Theory Simulation_SCLFOL_ORDRES

theory Simulation_SCLFOL_ORDRES
  imports
    Background
    ORD_RES
    ORD_RES_1
    ORD_RES_2
    ORD_RES_3
    ORD_RES_4
    ORD_RES_5
    ORD_RES_6
    ORD_RES_7
    ORD_RES_8
    ORD_RES_9
    ORD_RES_10
    ORD_RES_11
    Clause_Could_Propagate
begin

section ‹ORD-RES-1 (deterministic)›

type_synonym 'f ord_res_1_state = "'f gclause fset"

context simulation_SCLFOL_ground_ordered_resolution begin

sublocale backward_simulation_with_measuring_function where
  step1 = ord_res and
  step2 = ord_res_1 and
  final1 = ord_res_final and
  final2 = ord_res_1_final and
  order = "λ_ _. False" and
  match = "(=)" and
  measure = "λ_. ()"
proof unfold_locales
  show "wfP (λ_ _. False)"
    by simp
next
  show "N1 N2. N1 = N2  ord_res_1_final N2  ord_res_final N1"
    unfolding ord_res_1_final_def by metis
next
  fix N1 N2 N2' :: "'f ord_res_1_state"
  assume match: "N1 = N2" and step2: "ord_res_1 N2 N2'"
  show "(N1'. ord_res++ N1 N1'  N1' = N2')  N1 = N2'  False"
  proof (intro disjI1 exI conjI)

    have mempty_no_in: "{#} |∉| N2"
      if C_least: "linorder_cls.is_least_in_fset {|C |∈| N2.
        ¬ ord_res.interp (fset N2) C  ord_res.production (fset N2) C  C|} C" and
        L_max: "linorder_lit.is_maximal_in_mset C L"
      for C L
    proof (rule notI)
      assume "{#} |∈| N2"
      moreover have "¬ ord_res.interp (fset N2) {#}  ord_res.production (fset N2) {#}  {#}"
        by simp
      moreover have "C. {#} c C"
        using mempty_lesseq_cls by metis
      ultimately have "C = {#}"
        using C_least
        by (metis (no_types, lifting) ffmember_filter linorder_cls.is_least_in_fset_iff
            linorder_cls.less_le_not_le)
      moreover have "L ∈# C"
        using L_max by (simp add: linorder_lit.is_maximal_in_mset_iff)
      ultimately show "False"
        by simp
    qed

    have "ord_res N2 N2'"
      using step2
    proof (cases N2 N2' rule: ord_res_1.cases)
      case hyps: (factoring C L C')
      show ?thesis
      proof (rule ord_res.factoring)
        show "{#} |∉| N2"
          using hyps mempty_no_in is_least_false_clause_def by simp
      next
        show "ex_false_clause (fset N2)"
          unfolding ex_false_clause_def
          using hyps is_least_false_clause_def
          by (metis (no_types, lifting) linorder_cls.is_least_in_fset_ffilterD)
      next
        show "C |∈| N2"
          using hyps is_least_false_clause_def linorder_cls.is_least_in_fset_ffilterD(1) by blast
      next
        show "ord_res.ground_factoring C C'"
          using hyps by argo
      next
        show "N2' = finsert C' N2"
          using hyps by argo
      qed
    next
      case hyps: (resolution C L D CD)
      show ?thesis
      proof (rule ord_res.resolution)
        show "{#} |∉| N2"
          using hyps mempty_no_in is_least_false_clause_def by simp
      next
        show "ex_false_clause (fset N2)"
          unfolding ex_false_clause_def
          using hyps is_least_false_clause_def
          by (metis (no_types, lifting) linorder_cls.is_least_in_fset_ffilterD)
      next
        show "C |∈| N2"
          using hyps is_least_false_clause_def linorder_cls.is_least_in_fset_ffilterD(1) by blast
      next
        show "D |∈| N2"
          using hyps by argo
      next
        show "ord_res.ground_resolution C D CD"
          using hyps by argo
      next
        show "N2' = finsert CD N2"
          using hyps by argo
      qed
    qed
    thus "ord_res++ N1 N2'"
      unfolding match by simp
  next
    show "N2' = N2'" ..
  qed
qed

end


section ‹ORD-RES-2 (full factorization)›

type_synonym 'f ord_res_2_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"

context simulation_SCLFOL_ground_ordered_resolution begin

fun ord_res_1_matches_ord_res_2
  :: "'f ord_res_1_state  _  bool" where
  "ord_res_1_matches_ord_res_2 S1 (N, (Ur, Uef))  (Uf.
      S1 = N |∪| Ur |∪| Uef |∪| Uf 
      (Cf |∈| Uf. C |∈| N |∪| Ur |∪| Uef. ord_res.ground_factoring++ C Cf  Cf  efac Cf 
        (efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C)))"

lemma ord_res_1_matches_ord_res_2_simps':
  "ord_res_1_matches_ord_res_2 S1 (N, (Ur, Uef)) 
    (Uf. S1 = N |∪| Ur |∪| Uef |∪| Uf 
      (Cf |∈| Uf. Cf  efac Cf  (C |∈| N |∪| Ur |∪| Uef. ord_res.ground_factoring++ C Cf 
        (efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C))))"
  unfolding ord_res_1_matches_ord_res_2.simps by metis

lemma ord_res_1_matches_ord_res_2_simps'':
  "ord_res_1_matches_ord_res_2 S1 (N, (Ur, Uef)) 
    (Uf. S1 = N |∪| Ur |∪| Uef |∪| Uf 
      (Cf |∈| Uf. Cf  efac Cf  (C |∈| N |∪| Ur |∪| Uef. ord_res.ground_factoring++ C Cf 
        (efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C))))"
  unfolding ord_res_1_matches_ord_res_2_simps'
  by (metis ground_factorings_preserves_efac tranclp_into_rtranclp)

lemma ord_res_1_final_iff_ord_res_2_final:
  assumes match: "ord_res_1_matches_ord_res_2 S1 S2"
  shows "ord_res_1_final S1  ord_res_2_final S2"
proof -
  obtain N Ur Uef where "S2 = (N, (Ur, Uef))"
    by (metis prod.exhaust)
  with match obtain Uf where
    S1_def: "S1 = N |∪| Ur |∪| Uef |∪| Uf" and
    Uf_spec: "Cf |∈| Uf. C |∈| N |∪| Ur |∪| Uef. ord_res.ground_factoring++ C Cf  Cf  efac Cf 
      (efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C)"
    by auto

  have Uf_unproductive: "Cf |∈| Uf. ord_res.production (fset (N |∪| Ur |∪| Uef |∪| Uf)) Cf = {}"
  proof (intro ballI)
    fix Cf
    assume "Cf |∈| Uf"
    hence "Cf  efac Cf"
      using Uf_spec by metis
    hence "L. is_pos L  ord_res.is_strictly_maximal_lit L Cf"
      using nex_strictly_maximal_pos_lit_if_neq_efac by metis
    thus "ord_res.production (fset (N |∪| Ur |∪| Uef |∪| Uf)) Cf = {}"
      using unproductive_if_nex_strictly_maximal_pos_lit by metis
  qed

  have Interp_eq: "C. ord_res_Interp (fset (N |∪| Ur |∪| Uef |∪| Uf)) C =
    ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C"
    using Interp_union_unproductive[of "fset (N |∪| Ur |∪| Uef)" "fset Uf", folded union_fset,
        OF finite_fset finite_fset Uf_unproductive] .

  have "{#} |∈| N |∪| Ur |∪| Uef |∪| Uf  {#} |∈| N |∪| Ur |∪| Uef"
  proof (rule iffI)
    assume "{#} |∈| N |∪| Ur |∪| Uef |∪| Uf"
    hence "{#} |∈| N |∪| Ur |∪| Uef  {#} |∈| Uf"
      by simp
    thus "{#} |∈| N |∪| Ur |∪| Uef"
    proof (elim disjE)
      assume "{#} |∈| N |∪| Ur |∪| Uef"
      thus "{#} |∈| N |∪| Ur |∪| Uef"
        by assumption
    next
      assume "{#} |∈| Uf"
      hence "{#}  efac {#}"
        using Uf_spec[rule_format, of "{#}"] by metis
      hence False
        by simp
      thus "{#} |∈| N |∪| Ur |∪| Uef" ..
    qed
  next
    assume "{#} |∈| N |∪| Ur |∪| Uef"
    thus "{#} |∈| N |∪| Ur |∪| Uef |∪| Uf"
      by simp
  qed

  moreover have "¬ ex_false_clause (fset (N |∪| Ur |∪| Uef |∪| Uf)) 
    ¬ ex_false_clause (fset (N |∪| Ur |∪| Uef))"
  proof (rule iffI; erule contrapos_nn)
    assume "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
    thus "ex_false_clause (fset (N |∪| Ur |∪| Uef |∪| Uf))"
      unfolding ex_false_clause_def Interp_eq by auto
  next
    assume "ex_false_clause (fset (N |∪| Ur |∪| Uef |∪| Uf))"
    then obtain C where
      "C |∈| N |∪| Ur |∪| Uef |∪| Uf" and
      C_false: "¬ ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C  C"
      unfolding ex_false_clause_def Interp_eq by metis
    hence "C |∈| N |∪| Ur |∪| Uef  C |∈| Uf"
      by simp
    thus "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
    proof (elim disjE)
      assume "C |∈| N |∪| Ur |∪| Uef"
      thus "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
        unfolding ex_false_clause_def using C_false by metis
    next
      assume "C |∈| Uf"
      then obtain C' where "C' |∈| N |∪| Ur |∪| Uef" and
       "ord_res.ground_factoring++ C' C" and
       "C  efac C" and
       "efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C'"
        using Uf_spec[rule_format, of C] by metis
      thus "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
      proof (elim disjE exE conjE)
        assume "efac C |∈| Uef"

        show "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
        proof (cases "ord_res_Interp (fset (N |∪| Ur |∪| Uef)) (efac C)  efac C")
          case efac_C_true: True
          have "efac C ⊆# C"
            using efac_subset[of C] .
          hence "efac C c C"
            using subset_implies_reflclp_multp by metis
          hence "ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C  efac C"
            using efac_C_true ord_res.entailed_clause_stays_entailed by fastforce
          hence "ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C  C"
            using efac_C_true by (simp add: true_cls_def)
          with C_false have False
            by contradiction
          thus ?thesis ..
        next
          case False

          moreover have "efac C |∈| N |∪| Ur |∪| Uef"
            using efac C |∈| Uef by simp

          ultimately show "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
            unfolding ex_false_clause_def by metis
        qed
      next
        assume "is_least_false_clause (N |∪| Ur |∪| Uef) C'"
        hence "C' |∈| N |∪| Ur |∪| Uef" and "¬ ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C'  C'"
          using linorder_cls.is_least_in_ffilter_iff is_least_false_clause_def by simp_all
        thus "ex_false_clause (fset (N |∪| Ur |∪| Uef))"
          unfolding ex_false_clause_def by metis
      qed
    qed
  qed

  ultimately show ?thesis
    by (simp add: S1_def S2 = (N, Ur, Uef) ord_res_1_final_def ord_res_2_final.simps
        ord_res_final_def)
qed

lemma safe_states_if_ord_res_1_matches_ord_res_2:
  assumes match: "ord_res_1_matches_ord_res_2 S1 S2"
  shows "safe_state ord_res_1 ord_res_1_final S1  safe_state ord_res_2_step ord_res_2_final S2"
proof -
  have "safe_state ord_res_1 ord_res_1_final S1"
    using safe_state_if_all_states_safe ord_res_1_safe by metis

  moreover have "safe_state ord_res_2_step ord_res_2_final S2"
    using safe_state_if_all_states_safe ord_res_2_step_safe by metis

  ultimately show ?thesis
    by argo
qed

definition ord_res_1_measure where
  "ord_res_1_measure s1 =
    (if C. is_least_false_clause s1 C then
      The (is_least_false_clause s1)
    else
      {#})"

lemma forward_simulation:
  assumes match: "ord_res_1_matches_ord_res_2 s1 s2" and
    step1: "ord_res_1 s1 s1'"
  shows "(s2'. ord_res_2_step++ s2 s2'  ord_res_1_matches_ord_res_2 s1' s2') 
    ord_res_1_matches_ord_res_2 s1' s2  ord_res_1_measure s1' ⊂# ord_res_1_measure s1"
proof -
  let
    ?match = "ord_res_1_matches_ord_res_2" and
    ?measure = "ord_res_1_measure" and
    ?order = "(⊂#)"

  obtain N Ur Uef :: "'f gterm clause fset" where
    s2_def: "s2 = (N, (Ur, Uef))"
    by (metis prod.exhaust)

  from match obtain Uf where
    s1_def: "s1 = N |∪| Ur |∪| Uef |∪| Uf" and
    Uf_spec: "Cf |∈| Uf. C |∈| N |∪| Ur |∪| Uef. ord_res.ground_factoring++ C Cf  Cf  efac Cf 
      (efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C)"
    unfolding s2_def ord_res_1_matches_ord_res_2.simps by metis

  have Uf_unproductive: "Cf |∈| Uf. ord_res.production (fset (N |∪| Ur |∪| Uef |∪| Uf)) Cf = {}"
  proof (intro ballI)
    fix Cf
    assume "Cf |∈| Uf"
    hence "Cf  efac Cf"
      using Uf_spec by metis
    hence "L. is_pos L  ord_res.is_strictly_maximal_lit L Cf"
      using nex_strictly_maximal_pos_lit_if_neq_efac by metis
    thus "ord_res.production (fset (N |∪| Ur |∪| Uef |∪| Uf)) Cf = {}"
      using unproductive_if_nex_strictly_maximal_pos_lit by metis
  qed

  have Interp_eq: "C. ord_res_Interp (fset (N |∪| Ur |∪| Uef |∪| Uf)) C =
    ord_res_Interp (fset (N |∪| Ur |∪| Uef)) C"
    using Interp_union_unproductive[of "fset (N |∪| Ur |∪| Uef)" "fset Uf", folded union_fset,
        OF finite_fset finite_fset Uf_unproductive] .

  show "(s2'. ord_res_2_step++ s2 s2'  ?match s1' s2') 
    ?match s1' s2  ?order (?measure s1') (?measure s1)"
    using step1
  proof (cases s1 s1' rule: ord_res_1.cases)
    case (factoring C L C')

    have C_least_false: "is_least_false_clause (N |∪| Ur |∪| Uef |∪| Uf) C"
      using factoring
      unfolding is_least_false_clause_def s1_def by argo

    hence C_in: "C |∈| N |∪| Ur |∪| Uef |∪| Uf"
      unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff s1_def by argo
    hence C_in_disj: "C |∈| N |∪| Ur |∪| Uef  C |∈| Uf"
      by simp

    show ?thesis
    proof (cases "C' = efac C'")
      case True
      let ?s2' = "(N, (Ur, finsert C' Uef))"

      have "ord_res_2_step++ s2 ?s2'"
      proof (rule tranclp.r_into_trancl)
        show "ord_res_2_step s2 (N, Ur, finsert C' Uef)"
          using C_in_disj
        proof (elim disjE)
          assume "C |∈| N |∪| Ur |∪| Uef"
          show ?thesis
            unfolding s2_def
          proof (intro ord_res_2_step.intros ord_res_2.factoring)
            show "is_least_false_clause (N |∪| Ur |∪| Uef) C"
              using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
                  OF Uf_unproductive C |∈| N |∪| Ur |∪| Uef C_least_false]
              unfolding is_least_false_clause_def .
          next
            show "ord_res.is_maximal_lit L C"
              using ord_res.is_maximal_lit L C .
          next
            show "is_pos L"
              using is_pos L .
          next
            show "finsert C' Uef = finsert (efac C) Uef"
              using True factoring ground_factoring_preserves_efac by metis
          qed
        next
          assume "C |∈| Uf"
          then obtain x where
            "x |∈| N |∪| Ur |∪| Uef" and
            "ord_res.ground_factoring++ x C" and
            "C  efac C" and
            "efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x"
            using Uf_spec by metis

          show ?thesis
            unfolding s2_def
          proof (intro ord_res_2_step.intros ord_res_2.factoring)
            have efac C |∉| Uef
            proof (rule notI)
              have "efac C c C"
                using efac_subset[of C] subset_implies_reflclp_multp by metis
              hence "efac C c C"
                using C  efac C by order

              moreover assume "efac C |∈| Uef"

              ultimately show False
                using C_least_false[unfolded is_least_false_clause_def
                    linorder_cls.is_least_in_ffilter_iff]
                by (metis C  efac C funionCI linorder_cls.not_less_iff_gr_or_eq
                    ord_res.entailed_clause_stays_entailed set_mset_efac true_cls_def)
            qed
            thus "is_least_false_clause (N |∪| Ur |∪| Uef) x"
              using efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x by argo
          next
            show "ord_res.is_maximal_lit L x"
              using ord_res.ground_factoring++ x C ord_res.is_maximal_lit L C
              using ord_res.ground_factorings_preserves_maximal_literal
              by (metis tranclp_into_rtranclp)
          next
            show "is_pos L"
              using is_pos L .
          next
            show "finsert C' Uef = finsert (efac x) Uef"
              using ord_res.ground_factoring++ x C ord_res.ground_factoring C C'
              using True ground_factorings_preserves_efac ground_factoring_preserves_efac
              by (metis tranclp_into_rtranclp)
          qed
        qed
      qed

      moreover have "?match s1' ?s2'"
      proof -
        have "s1' = N |∪| Ur |∪| finsert C' Uef |∪| Uf"
          unfolding s1' = finsert C' s1 s1_def by simp

        moreover have "C |∈| N |∪| Ur |∪| finsert C' Uef.
          ord_res.ground_factoring++ C Cf  Cf  efac Cf 
          (efac Cf |∈| finsert C' Uef  is_least_false_clause (N |∪| Ur |∪| finsert C' Uef) C)"
          if "Cf |∈| Uf" for Cf
        proof -
          obtain x where
            "x |∈| N |∪| Ur |∪| Uef" and
            "ord_res.ground_factoring++ x Cf" and
            "Cf  efac Cf" and
            "efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x"
            using Cf |∈| Uf Uf_spec by metis

          show ?thesis
          proof (intro bexI conjI)
            show "x |∈| N |∪| Ur |∪| finsert C' Uef"
              using x |∈| N |∪| Ur |∪| Uef by simp
          next
            show "ord_res.ground_factoring++ x Cf"
              using ord_res.ground_factoring++ x Cf .
          next
            show "Cf  efac Cf"
              using Cf  efac Cf .
          next
            show "efac Cf |∈| finsert C' Uef  is_least_false_clause (N |∪| Ur |∪| finsert C' Uef) x"
              using efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x
            proof (elim disjE)
              assume "efac Cf |∈| Uef"
              thus ?thesis
                by simp
            next
              assume "is_least_false_clause (N |∪| Ur |∪| Uef) x"
              show ?thesis
              proof (cases "C' = efac x")
                case True
                moreover have "efac x = efac Cf"
                  using ord_res.ground_factoring++ x Cf ground_factorings_preserves_efac
                  by (metis tranclp_into_rtranclp)
                ultimately show ?thesis
                  by simp
              next
                case False
                show ?thesis
                  using C_in_disj
                proof (elim disjE)
                  assume "C |∈| N |∪| Ur |∪| Uef"
                  then show ?thesis
                    by (smt (verit) C_least_false True Uf_unproductive
                        is_least_false_clause (N |∪| Ur |∪| Uef) x ord_res.ground_factoring++ x Cf
                        finsert_iff ground_factoring_preserves_efac ground_factorings_preserves_efac
                        linorder_cls.Uniq_is_least_in_fset local.factoring(4)
                        is_least_false_clause_def
                        is_least_false_clause_if_is_least_false_clause_in_union_unproductive
                        the1_equality' tranclp_into_rtranclp)
                next
                  assume "C |∈| Uf"
                  then show ?thesis
                    using C_least_false
                    using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
                        OF Uf_unproductive]
                    by (smt (z3) True Uf_spec is_least_false_clause (N |∪| Ur |∪| Uef) x
                        ord_res.ground_factoring++ x Cf finsert_absorb finsert_iff
                        ground_factoring_preserves_efac ground_factorings_preserves_efac
                        linorder_cls.Uniq_is_least_in_fset local.factoring(4)
                        is_least_false_clause_def the1_equality' tranclp_into_rtranclp)
                qed
              qed
            qed
          qed
        qed

        ultimately show ?thesis
          by auto
      qed

      ultimately show ?thesis
        by metis
    next
      case False
      let ?Uf' = "finsert C' Uf"

      have "?match s1' s2"
      proof -
        have "finsert C' s1 = N |∪| Ur |∪| Uef |∪| ?Uf'"
          unfolding s1_def by simp

        moreover have "C |∈| N |∪| Ur |∪| Uef.
          ord_res.ground_factoring++ C Cf  Cf  efac Cf 
          (efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C)"
          if "Cf |∈| ?Uf'" for Cf
        proof -
          from Cf |∈| ?Uf' have "Cf = C'  Cf |∈| Uf"
            by simp
          thus ?thesis
          proof (elim disjE)
            assume "Cf = C'"
            thus ?thesis
              using C_in_disj
            proof (elim disjE)
              assume "C |∈| N |∪| Ur |∪| Uef"
              show ?thesis
              proof (intro bexI conjI)
                show "C |∈| N |∪| Ur |∪| Uef"
                  using C |∈| N |∪| Ur |∪| Uef .
              next
                show "ord_res.ground_factoring++ C Cf"
                  using ord_res.ground_factoring C C' Cf = C' by simp
              next
                show "Cf  efac Cf"
                  using False Cf = C' by argo
              next
                have "is_least_false_clause (N |∪| Ur |∪| Uef) C"
                  using factoring
                  using Interp_eq C |∈| N |∪| Ur |∪| Uef linorder_cls.is_least_in_ffilter_iff
                  by (simp add: s1_def is_least_false_clause_def)
                thus "efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) C" ..
              qed
            next
              assume "C |∈| Uf"
              then obtain x where
                "x |∈| N |∪| Ur |∪| Uef" and
                "ord_res.ground_factoring++ x C" and
                "C  efac C" and
                "efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x"
                using Uf_spec by metis

              show ?thesis
              proof (intro bexI conjI)
                show "x |∈| N |∪| Ur |∪| Uef"
                  using x |∈| N |∪| Ur |∪| Uef .
              next
                show "ord_res.ground_factoring++ x Cf"
                  using ord_res.ground_factoring++ x C ord_res.ground_factoring C C' Cf = C'
                  by simp
              next
                show "Cf  efac Cf"
                  using False Cf = C' by argo
              next
                show "efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x"
                  using efac C |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x
                proof (elim disjE)
                  assume "efac C |∈| Uef"

                  moreover have "efac C = efac Cf"
                    unfolding Cf = C'
                    using ord_res.ground_factoring C C' ground_factoring_preserves_efac by metis

                  ultimately show ?thesis
                    by argo
                next
                  assume "is_least_false_clause (N |∪| Ur |∪| Uef) x"
                  thus ?thesis
                    by argo
                qed
              qed
            qed
          next
            assume "Cf |∈| Uf"
            thus ?thesis
              using Uf_spec by metis
          qed
        qed

        ultimately have "ord_res_1_matches_ord_res_2 (finsert C' s1) (N, (Ur, Uef))"
          unfolding ord_res_1_matches_ord_res_2.simps by metis
        thus ?thesis
          unfolding s2_def s1' = finsert C' s1 by simp
      qed

      moreover have "?order (?measure s1') (?measure s1)"
      proof -
        have "?measure s1 = C"
          unfolding ord_res_1_measure_def
          using C_least_false[folded s1_def]
          by (metis (mono_tags, lifting) linorder_cls.Uniq_is_least_in_fset
              is_least_false_clause_def the1_equality' the_equality)

        moreover have "?measure s1' = C'"
        proof -
          have "C' c C"
            using factoring ord_res.ground_factoring_smaller_conclusion by metis

          have unproductive: "x{C'}. ord_res.production (fset s1  {C'}) x = {}"
            using C'  efac C'
            by (simp add: nex_strictly_maximal_pos_lit_if_neq_efac
                unproductive_if_nex_strictly_maximal_pos_lit)

          have Interp_eq: "D. ord_res_Interp (fset s1) D = ord_res_Interp (fset (finsert C' s1)) D"
            using Interp_union_unproductive[of "fset s1" "{C'}", folded union_fset,
                OF finite_fset _ unproductive]
            by simp

          have "is_least_false_clause (finsert C' s1) C'"
            unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
          proof (intro conjI ballI impI)
            have "¬ ord_res_Interp (fset s1) C  C"
              using C_least_false s1_def is_least_false_clause_def
                linorder_cls.is_least_in_ffilter_iff by simp
            thus "¬ ord_res_Interp (fset (finsert C' s1)) C'  C'"
              by (metis Interp_eq C' c C local.factoring(4)
                  ord_res.entailed_clause_stays_entailed
                  ord_res.set_mset_eq_set_mset_if_ground_factoring subset_refl true_cls_mono)
          next
            fix y
            assume "y |∈| finsert C' s1" and "y  C'" and
              y_false: "¬ ord_res_Interp (fset (finsert C' s1)) y  y"
            hence "y |∈| s1"
              by simp

            moreover have "¬ ord_res_Interp (fset s1) y  y"
              using y_false
              unfolding Interp_eq .

            ultimately have "C c y"
              using C_least_false[folded s1_def, unfolded is_least_false_clause_def]
              unfolding linorder_cls.is_least_in_ffilter_iff
              by force
            thus "C' c y"
              using C' c C by order
          qed simp
          thus ?thesis
            unfolding ord_res_1_measure_def s1' = finsert C' s1
            by (metis (mono_tags, lifting) linorder_cls.Uniq_is_least_in_fset
                is_least_false_clause_def the1_equality' the_equality)
        qed

        moreover have "C' ⊂# C"
          using factoring ord_res.strict_subset_mset_if_ground_factoring by metis

        ultimately show ?thesis
          unfolding s1_def  by simp
      qed

      ultimately show ?thesis
        by argo
    qed
  next
    case (resolution C L D CD)

    have "is_least_false_clause s1 C"
      using resolution unfolding is_least_false_clause_def by argo
    hence
      "C |∈| s1" and
      "¬ ord_res_Interp (fset s1) C  C" and
      "x |∈| s1. ¬ ord_res_Interp (fset s1) x  x  x  C  C c x"
      unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff by simp_all

    have "C |∉| Uf"
    proof (rule notI)
      assume "C |∈| Uf"
      then show False
        by (metis Uf_spec Uniq_D is_pos_def linorder_lit.Uniq_is_maximal_in_mset local.resolution(2)
            local.resolution(3) efac_spec)
    qed
    hence "C |∈| N |∪| Ur |∪| Uef"
      using C |∈| s1 by (simp add: s1_def)

    have C_least_false: "is_least_false_clause (N |∪| Ur |∪| Uef |∪| Uf) C"
      using resolution s1_def by metis
    hence C_least_false': "is_least_false_clause (N |∪| Ur |∪| Uef) C"
      using is_least_false_clause_if_is_least_false_clause_in_union_unproductive[
          OF Uf_unproductive C |∈| N |∪| Ur |∪| Uef] by argo

    define s2' where
      "s2' = (N, (finsert CD Ur, Uef))"

    have "ord_res_2_step++ s2 s2'"
    proof -
      have "D |∉| Uf"
      proof (rule notI)
        assume "D |∈| Uf"
        thus False
          using ord_res.production (fset s1) D = {atm_of L}
          using Uf_unproductive s1_def by simp
      qed
      hence D_in: "D |∈| N |∪| Ur |∪| Uef"
        using D |∈| s1[unfolded s1_def] by simp

      have "ord_res_2 N (Ur, Uef) (finsert CD Ur, Uef)"
      proof (rule ord_res_2.resolution)
        show "is_least_false_clause (N |∪| Ur |∪| Uef) C"
          using C_least_false' .
      next
        show "ord_res.is_maximal_lit L C"
          using resolution by argo
      next
        show "is_neg L"
          using resolution by argo
      next
        show "D |∈| N |∪| Ur |∪| Uef"
          using D_in .
      next
        show "D c C"
          using resolution by argo
      next
        show "ord_res.production (fset (N |∪| Ur |∪| Uef)) D = {atm_of L}"
          using resolution
          unfolding s1_def
          using production_union_unproductive[OF finite_fset finite_fset _ D_in] Uf_unproductive
          by (metis (no_types, lifting) union_fset)
      next
        show "ord_res.ground_resolution C D CD"
          using resolution by argo
      qed simp_all
      thus ?thesis
        by (auto simp: s2_def s2'_def ord_res_2_step.simps)
    qed

    moreover have "?match s1' s2'"
    proof -
      have "finsert CD (N |∪| Ur |∪| Uef |∪| Uf) = N |∪| finsert CD Ur |∪| Uef |∪| Uf"
        by simp

      moreover have "C |∈| N |∪| finsert CD Ur |∪| Uef.
        ord_res.ground_factoring++ C Cf  Cf  efac Cf 
        (efac Cf |∈| Uef  is_least_false_clause (N |∪| finsert CD Ur |∪| Uef) C)"
        if "Cf |∈| Uf" for Cf
      proof -
        obtain x where
          "x |∈| N |∪| Ur |∪| Uef" and
          "ord_res.ground_factoring++ x Cf" and
          "Cf  efac Cf" and
          "efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x"
          using Cf |∈| Uf Uf_spec by metis
        show ?thesis
        proof (intro bexI conjI)
          show "x |∈| N |∪| finsert CD Ur |∪| Uef"
            using x |∈| N |∪| Ur |∪| Uef by simp
        next
          show "ord_res.ground_factoring++ x Cf"
            using ord_res.ground_factoring++ x Cf .
        next
          show "Cf  efac Cf"
            using Cf  efac Cf .
        next
          show efac Cf |∈| Uef  is_least_false_clause (N |∪| finsert CD Ur |∪| Uef) x
            using efac Cf |∈| Uef  is_least_false_clause (N |∪| Ur |∪| Uef) x x |∈| N |∪| Ur |∪| Uef
            by (metis (no_types, lifting) C_least_false' Uniq_D ord_res.ground_factoring++ x Cf
                is_least_false_clause_def is_pos_def linorder_cls.Uniq_is_least_in_fset
                linorder_lit.Uniq_is_maximal_in_mset local.resolution(2) local.resolution(3)
                ord_res.ground_factoring.cases tranclpD)
        qed
      qed

      ultimately show ?thesis
        unfolding s1_def resolution s2'_def by auto
    qed

    ultimately show ?thesis
      by metis
  qed
qed

theorem bisimulation_ord_res_1_ord_res_2:
  defines "match  λi s1 s2. i = ord_res_1_measure s1  ord_res_1_matches_ord_res_2 s1 s2"
  shows "(MATCH :: nat × nat  'f ord_res_1_state  'f ord_res_2_state  bool) .
    bisimulation ord_res_1 ord_res_2_step ord_res_1_final ord_res_2_final  MATCH"  
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_1_state ⇒ 'f ord_res_2_state ⇒ bool) ℛfb.
    bisimulation ord_res_1 ord_res_1_final ord_res_2_step ord_res_2_final MATCH ℛfb" *)
proof (rule ex_bisimulation_from_forward_simulation)
  show "right_unique ord_res_1"
    using right_unique_ord_res_1 .
next
  show "right_unique ord_res_2_step"
    using right_unique_ord_res_2_step .
next
  show "s1. ord_res_1_final s1  (s1'. ord_res_1 s1 s1')"
    using ord_res_1_semantics.final_finished
    by (simp add: finished_def)
next
  show "s2. ord_res_2_final s2  (s2'. ord_res_2_step s2 s2')"
    using ord_res_2_semantics.final_finished
    by (simp add: finished_def)
next
  show "i s1 s2. match i s1 s2  ord_res_1_final s1 = ord_res_2_final s2"
    using ord_res_1_final_iff_ord_res_2_final
    by (simp add: match_def)
next
  show "i s1 s2. match i s1 s2 
    safe_state ord_res_1 ord_res_1_final s1 
    safe_state ord_res_2_step ord_res_2_final s2"
  proof (intro allI impI)
    fix i s1 S2
    assume "match i s1 S2"

    then obtain N s2 where
      S2_def: "S2 = (N, s2)" and
      "i = ord_res_1_measure s1" and
      match: "ord_res_1_matches_ord_res_2 s1 S2"
      unfolding match_def
      by (metis prod.exhaust)

    show "safe_state ord_res_1 ord_res_1_final s1  safe_state ord_res_2_step ord_res_2_final S2"
      using safe_states_if_ord_res_1_matches_ord_res_2[OF match] .
  qed
next
  show "wfP (⊂#)"
    using wfp_subset_mset .
next
  show "i s1 s2 s1'. match i s1 s2  ord_res_1 s1 s1' 
    (i' s2'. ord_res_2_step++ s2 s2'  match i' s1' s2')  (i'. match i' s1' s2  i' ⊂# i)"
  proof (intro allI impI)
    fix i s1 S2 s1'
    assume "match i s1 S2"
    then obtain N s2 where
      S2_def: "S2 = (N, s2)" and "i = ord_res_1_measure s1" and "ord_res_1_matches_ord_res_2 s1 S2"
      unfolding match_def
      by (metis prod.exhaust)

    moreover assume "ord_res_1 s1 s1'"

    ultimately have "(S2'. ord_res_2_step++ S2 S2'  ord_res_1_matches_ord_res_2 s1' S2') 
    ord_res_1_matches_ord_res_2 s1' S2  ord_res_1_measure s1' ⊂# ord_res_1_measure s1"
      using forward_simulation by metis

    thus "(i' S2'. ord_res_2_step++ S2 S2'  match i' s1' S2')  (i'. match i' s1' S2  i' ⊂# i)"
      unfolding S2_def prod.case
      using lift_tranclp_to_pairs_with_constant_fst[of ord_res_2 N s2]
      by (metis (mono_tags, lifting) i = ord_res_1_measure s1 match_def)
  qed
qed

end


section ‹ORD-RES-3 (full resolve)›

type_synonym 'f ord_res_3_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_2_matches_ord_res_3 :: "_  'f ord_res_3_state  bool" where
  "(C |∈| Upr. D1 |∈| N |∪| Uer |∪| Uef. D2 |∈| N |∪| Uer |∪| Uef.
      (ground_resolution D1)++ D2 C  C  eres D1 D2  eres D1 D2 |∈| Uer) 
  ord_res_2_matches_ord_res_3 (N, (Upr |∪| Uer, Uef)) (N, (Uer, Uef))"

lemma ord_res_2_final_iff_ord_res_3_final:
  assumes match: "ord_res_2_matches_ord_res_3 S2 S3"
  shows "ord_res_2_final S2  ord_res_3_final S3"
  using match
proof (cases S2 S3 rule: ord_res_2_matches_ord_res_3.cases)
  case match_hyps: (1 Upr N Uer Uef)

  note invars = match_hyps(3-)

  have Upr_spec: "C|∈|Upr. D1|∈|N |∪| Uer |∪| Uef. D2|∈|N |∪| Uer |∪| Uef.
    (ground_resolution D1)++ D2 C  C  eres D1 D2  eres D1 D2 |∈| Uer"
    using invars by argo

  have least_false_spec: "is_least_false_clause (N |∪| Upr |∪| Uer |∪| Uef) =
    is_least_false_clause (N |∪| Uer |∪| Uef)"
    using invars is_least_false_clause_conv_if_partial_resolution_invariant by metis

  have Upr_unproductive: "C |∈| Upr. ord_res.production (fset (N |∪| Uer |∪| Uef |∪| Upr)) C = {}"
  proof (intro ballI)
    fix C
    assume "C |∈| Upr"
    hence "L. is_pos L  ord_res.is_strictly_maximal_lit L C"
      using Upr_spec
      by (metis eres_eq_after_tranclp_ground_resolution nex_strictly_maximal_pos_lit_if_neq_eres)
    thus "ord_res.production (fset (N |∪| Uer |∪| Uef |∪| Upr)) C = {}"
      using unproductive_if_nex_strictly_maximal_pos_lit by metis
  qed

  hence Interp_N_Ur_Uef_eq_Interp_N_Uer_Uef: "C.
    ord_res_Interp (fset (N |∪| Upr |∪| Uer |∪| Uef)) C =
    ord_res_Interp (fset (N |∪| Uer |∪| Uef)) C"
    using Interp_union_unproductive[OF finite_fset finite_fset, folded union_fset,
        of Upr "N |∪| Uer |∪| Uef"]
    by (simp add: funion_left_commute sup_commute)

  have "ex_false_clause (fset (N |∪| Upr |∪| Uer |∪| Uef)) 
    ex_false_clause (fset (N |∪| Uer |∪| Uef))"
  proof (rule iffI)
    assume "ex_false_clause (fset (N |∪| Upr |∪| Uer |∪| Uef))"
    then obtain C where "is_least_false_clause (N |∪| Upr |∪| Uer |∪| Uef) C"
      using obtains_least_false_clause_if_ex_false_clause by metis
    thus "ex_false_clause (fset (N |∪| Uer |∪| Uef))"
      using least_false_spec ex_false_clause_iff by metis
  next
    assume "ex_false_clause (fset (N |∪| Uer |∪| Uef))"
    thus "ex_false_clause (fset (N |∪| Upr |∪| Uer |∪| Uef))"
      unfolding ex_false_clause_def
      unfolding Interp_N_Ur_Uef_eq_Interp_N_Uer_Uef
      by auto
  qed

  moreover have "{#} |∈| N |∪| Upr |∪| Uer |∪| Uef  {#} |∈| N |∪| Uer |∪| Uef"
  proof (rule iffI)
    assume "{#} |∈| N |∪| Upr |∪| Uer |∪| Uef"
    hence "{#} |∈| N |∪| Uef  {#} |∈| Upr |∪| Uer"
      by auto
    thus "{#} |∈| N |∪| Uer |∪| Uef"
    proof (elim disjE)
      assume "{#} |∈| N |∪| Uef"
      thus ?thesis
        by auto
    next
      have "{#} |∉| Upr"
        using Upr_spec[rule_format, of "{#}"]
        by (metis eres_eq_after_tranclp_ground_resolution eres_mempty_right)
      moreover assume "{#} |∈| Upr |∪| Uer"
      ultimately show ?thesis
        by simp
    qed
  next
    assume "{#} |∈| N |∪| Uer |∪| Uef"
    then show "{#} |∈| N |∪| Upr |∪| Uer |∪| Uef"
      by auto
  qed

  ultimately have "ord_res_final (N |∪| Upr |∪| Uer |∪| Uef) = ord_res_final (N |∪| Uer |∪| Uef)"
    unfolding ord_res_final_def by argo

  thus "ord_res_2_final S2  ord_res_3_final S3"
    unfolding match_hyps(1,2)
    by (simp add: ord_res_2_final.simps ord_res_3_final.simps sup_assoc)
qed

definition ord_res_2_measure where
  "ord_res_2_measure S1 =
    (let (N, (Ur, Uef)) = S1 in
    (if C. is_least_false_clause (N |∪| Ur |∪| Uef) C then
      The (is_least_false_clause (N |∪| Ur |∪| Uef))
    else
      {#}))"

definition resolvent_at where
  "resolvent_at C D i = (THE CD. (ground_resolution C ^^ i) D CD)"

lemma resolvent_at_0[simp]: "resolvent_at C D 0 = D"
  by (simp add: resolvent_at_def)

lemma resolvent_at_less_cls_resolvent_at:
  assumes reso_at: "(ground_resolution C ^^ n) D CD"
  assumes "i < j" and "j  n"
  shows "resolvent_at C D j c resolvent_at C D i"
proof -
  obtain j' where
    "j = i + Suc j'"
    using i < j by (metis less_iff_Suc_add nat_arith.suc1)

  obtain n' where
    "n = j + n'"
    using j  n by (metis le_add_diff_inverse)

  obtain CDi CDj CDn where
    "(ground_resolution C ^^ i) D CDi" and
    "(ground_resolution C ^^ Suc j') CDi CDj"
    "(ground_resolution C ^^ n') CDj CDn"
    using reso_at n = j + n' j = i + Suc j' by (metis relpowp_plusD)

  have *: "resolvent_at C D i = CDi"
    unfolding resolvent_at_def
    using (ground_resolution C ^^ i) D CDi
    by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')

  have "(ground_resolution C ^^ j) D CDj"
    unfolding j = i + Suc j'
    using (ground_resolution C ^^ i) D CDi (ground_resolution C ^^ Suc j') CDi CDj
    by (metis relpowp_trans)
  hence **: "resolvent_at C D j = CDj"
    unfolding resolvent_at_def
    by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')

  have "(ground_resolution C)++ CDi CDj"
    using (ground_resolution C ^^ Suc j') CDi CDj
    by (metis Zero_not_Suc tranclp_if_relpowp)
  hence "CDj c CDi"
    using resolvent_lt_right_premise_if_tranclp_ground_resolution by metis
  thus ?thesis
    unfolding * ** .
qed

lemma
  assumes reso_at: "(ground_resolution C ^^ n) D CD" and "i < n"
  shows
    left_premisse_lt_resolvent_at: "C c resolvent_at C D i" and
    max_lit_resolvent_at:
      "ord_res.is_maximal_lit L D  ord_res.is_maximal_lit L (resolvent_at C D i)" and
    nex_pos_strictly_max_lit_in_resolvent_at:
      "L. is_pos L  ord_res.is_strictly_maximal_lit L (resolvent_at C D i)" and
    ground_resolution_resolvent_at_resolvent_at_Suc:
      "ground_resolution C (resolvent_at C D i) (resolvent_at C D (Suc i))" and
    relpowp_to_resolvent_at: "(ground_resolution C ^^ i) D (resolvent_at C D i)"
proof -
  obtain j where n_def: "n = i + Suc j"
    using i < n less_natE by auto

  obtain CD' where "(ground_resolution C ^^ i) D CD'" and "(ground_resolution C ^^ Suc j) CD' CD"
    using reso_at n_def by (metis relpowp_plusD)

  have "resolvent_at C D i = CD'"
    unfolding resolvent_at_def
    using (ground_resolution C ^^ i) D CD'
    by (simp add: Uniq_ground_resolution Uniq_relpowp the1_equality')

  have "C c CD'"
  proof (rule left_premise_lt_right_premise_if_tranclp_ground_resolution)
    show "(ground_resolution C)++ CD' CD"
      using (ground_resolution C ^^ Suc j) CD' CD
      by (metis Zero_not_Suc tranclp_if_relpowp)
  qed
  thus "C c resolvent_at C D i"
    unfolding resolvent_at C D i = CD' by argo

  show "ord_res.is_maximal_lit L (resolvent_at C D i)" if "ord_res.is_maximal_lit L D"
    unfolding resolvent_at C D i = CD'
    using that
    using (ground_resolution C ^^ i) D CD'
    by (smt (verit, ccfv_SIG) Uniq_ground_resolution Uniq_relpowp Zero_not_Suc
        thesis. (CD'. (ground_resolution C ^^ i) D CD'; (ground_resolution C ^^ Suc j) CD' CD  thesis)  thesis
        linorder_lit.Uniq_is_greatest_in_mset linorder_lit.Uniq_is_maximal_in_mset literal.sel(1)
        n_def relpowp_ground_resolutionD reso_at the1_equality' zero_eq_add_iff_both_eq_0)

  show "L. is_pos L  ord_res.is_strictly_maximal_lit L (resolvent_at C D i)"
    unfolding resolvent_at C D i = CD'
    by (metis Zero_not_Suc (ground_resolution C ^^ Suc j) CD' CD
        nex_strictly_maximal_pos_lit_if_resolvable tranclpD tranclp_if_relpowp)

  show "ground_resolution C (resolvent_at C D i) (resolvent_at C D (Suc i))"
  proof -
    obtain CD'' where "ground_resolution C CD' CD''" and "(ground_resolution C ^^ j) CD'' CD"
      using (ground_resolution C ^^ Suc j) CD' CD by (metis relpowp_Suc_D2)
    hence "(ground_resolution C ^^ Suc i) D CD''"
      using (ground_resolution C ^^ i) D CD' by auto
    hence "resolvent_at C D (Suc i) = CD''"
      unfolding resolvent_at_def
      by (meson Uniq_ground_resolution Uniq_relpowp the1_equality')

    show ?thesis
      unfolding resolvent_at C D i = CD' resolvent_at C D (Suc i) = CD''
      using ground_resolution C CD' CD'' .
  qed

  show "(ground_resolution C ^^ i) D (resolvent_at C D i)"
    using (ground_resolution C ^^ i) D CD' resolvent_at C D i = CD' by argo
qed

definition resolvents_upto where
  "resolvents_upto C D n = resolvent_at C D |`| fset_upto (Suc 0) n"

lemma resolvents_upto_0[simp]:
  "resolvents_upto C D 0 = {||}"
  by (simp add: resolvents_upto_def)

lemma resolvents_upto_Suc[simp]:
  "resolvents_upto C D (Suc n) = finsert (resolvent_at C D (Suc n)) (resolvents_upto C D n)"
  by (simp add: resolvents_upto_def)

lemma resolvent_at_fmember_resolvents_upto:
  assumes "k  0"
  shows "resolvent_at C D k |∈| resolvents_upto C D k"
  unfolding resolvents_upto_def
proof (rule fimageI)
  show "k |∈| fset_upto (Suc 0) k"
    using assms by simp
qed

lemma backward_simulation_2_to_3:
  fixes match measure less
  defines "match  ord_res_2_matches_ord_res_3"
  assumes
    match: "match S2 S3" and
    step2: "ord_res_3_step S3 S3'"
  shows "(S2'. ord_res_2_step++ S2 S2'  match S2' S3')"
  using match[unfolded match_def]
proof (cases S2 S3 rule: ord_res_2_matches_ord_res_3.cases)
  case match_hyps: (1 Upr N Uer Uef)
  note invars = match_hyps(3-)

  have Upr_spec: "C|∈|Upr. D1|∈|N |∪| Uer |∪| Uef. D2|∈|N |∪| Uer |∪| Uef.
    (ground_resolution D1)++ D2 C  C  eres D1 D2  eres D1 D2 |∈| Uer"
    using invars by argo

  hence C_not_least_with_partial: "¬ is_least_false_clause (N |∪| Upr |∪| Uer |∪| Uef) C"
    if C_in: "C |∈| Upr" for C
  proof -
    obtain D1 D2 where
      "D1 |∈| N |∪| Uer |∪| Uef" and
      "D2 |∈| N |∪| Uer |∪| Uef" and
      "(ground_resolution D1)++ D2 C" and
      "C  eres D1 D2" and
      "eres D1 D2 |∈| Uer"
      using Upr_spec C_in by metis

    have "eres D1 C = eres D1 D2"
      using (ground_resolution D1)++ D2 C eres_eq_after_tranclp_ground_resolution by metis
    hence "eres D1 C c C"
      using eres_le[of D1 C] C  eres D1 D2 by order

    show ?thesis
    proof (cases "ord_res_Interp (fset (N |∪| Upr |∪| Uer |∪| Uef)) (eres D1 D2)  eres D1 D2")
      case True
      then show ?thesis
        by (metis (no_types, lifting) (ground_resolution D1)++ D2 C eres D1 C = eres D1 D2
            clause_true_if_eres_true is_least_false_clause_def
            linorder_cls.is_least_in_fset_ffilterD(2))
    next
      case False
      then show ?thesis
        by (metis (mono_tags, lifting) Un_iff eres D1 C = eres D1 D2 eres D1 C c C
            eres D1 D2 |∈| Uer is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
            linorder_cls.not_less_iff_gr_or_eq sup_fset.rep_eq)
    qed
  qed

  have least_false_conv: "is_least_false_clause (N |∪| Upr |∪| Uer |∪| Uef) =
    is_least_false_clause (N |∪| Uer |∪| Uef)"
    using invars is_least_false_clause_conv_if_partial_resolution_invariant by metis

  have Upr_unproductive: "N. C |∈| Upr. ord_res.production N C = {}"
  proof (intro ballI)
    fix C
    assume "C |∈| Upr"
    hence "D |∈| N |∪| Uer |∪| Uef. (C'. ground_resolution D C C')"
      using Upr_spec by (metis eres_eq_after_tranclp_ground_resolution resolvable_if_neq_eres)
    hence "L. is_pos L  ord_res.is_strictly_maximal_lit L C"
      using nex_strictly_maximal_pos_lit_if_resolvable by metis
    thus "N. ord_res.production N C = {}"
      using unproductive_if_nex_strictly_maximal_pos_lit by metis
  qed

  hence Interp_N_Ur_Uef_eq_Interp_N_Uer_Uef:
    "C. ord_res_Interp (fset (N |∪| Upr |∪| Uer |∪| Uef)) C = ord_res_Interp (fset (N |∪| Uer |∪| Uef)) C"
    using Interp_union_unproductive[OF finite_fset finite_fset, folded union_fset,
        of Upr "N |∪| Uer |∪| Uef"]
    by (simp add: funion_left_commute sup_commute)

  have Upr_have_generalization: "Ca |∈| Upr. D |∈| Uer. D c Ca  {D} ⊫e {Ca}"
  proof (intro ballI)
    fix Ca
    assume "Ca |∈| Upr"
    then obtain D1 D2 where
      "D1|∈|N |∪| Uer |∪| Uef" and
      "D2|∈|N |∪| Uer |∪| Uef" and
      "(ground_resolution D1)++ D2 Ca" and
      "Ca  eres D1 D2" and
      "eres D1 D2 |∈| Uer"
      using Upr_spec by metis

    have "eres D1 D2 = eres D1 Ca"
      using (ground_resolution D1)++ D2 Ca eres_eq_after_tranclp_ground_resolution by metis

    show "D |∈| Uer. D c Ca  {D} ⊫e {Ca}"
    proof (intro bexI conjI)
      have "eres D1 Ca c Ca"
        using eres_le .
      thus "eres D1 D2 c Ca"
        using Ca  eres D1 D2 eres D1 D2 = eres D1 Ca by order
    next
      show "{eres D1 D2} ⊫e {Ca}"
        using (ground_resolution D1)++ D2 Ca eres_entails_resolvent by metis
    next
      show "eres D1 D2 |∈| Uer"
        using eres D1 D2 |∈| Uer by simp
    qed
  qed

  from step2 obtain s3' where S3'_def: "S3' = (N, s3')" and "ord_res_3 N (Uer, Uef) s3'"
    by (auto simp: match_hyps(1,2) elim: ord_res_3_step.cases)

  show ?thesis
    using ord_res_3 N (Uer, Uef) s3'
  proof (cases N "(Uer, Uef)" s3' rule: ord_res_3.cases)
    case (factoring C L Uef')

    define S2' where
      "S2' = (N, (Upr |∪| Uer, finsert (efac C) Uef))"

    have "ord_res_2_step++ S2 S2'"
      unfolding match_hyps(1,2) S2'_def
    proof (intro tranclp.r_into_trancl ord_res_2_step.intros ord_res_2.factoring)
      show "is_least_false_clause (N |∪| (Upr |∪| Uer) |∪| Uef) C"
        using is_least_false_clause (N |∪| Uer |∪| Uef) C
        using least_false_conv
        by (metis sup_assoc)
    next
      show "ord_res.is_maximal_lit L C"
        using factoring by metis
    next
      show "is_pos L"
        using factoring by metis
    next
      show "finsert (efac C) Uef = finsert (efac C) Uef"
        by argo
    qed

    moreover have "match S2' S3'"
      unfolding S2'_def S3'_def
      unfolding factoring
      unfolding match_def
    proof (rule ord_res_2_matches_ord_res_3.intros)
      show "Ca|∈|Upr.
        D1|∈|N |∪| Uer |∪| finsert (efac C) Uef. D2|∈|N |∪| Uer |∪| finsert (efac C) Uef.
        (ground_resolution D1)++ D2 Ca  Ca  eres D1 D2  eres D1 D2 |∈| Uer"
        using Upr_spec by auto
    qed

    ultimately show ?thesis
      by metis
  next
    case (resolution C L D Urr')

    have "(ground_resolution D)** C (eres D C)" "x. ground_resolution D (eres D C) x"
      unfolding atomize_conj
      by (metis ex1_eres_eq_full_run_ground_resolution full_run_def)

    moreover have "x. ground_resolution D C x"
      unfolding ground_resolution_def
      using resolution
      by (metis Neg_atm_of_iff ex_ground_resolutionI ord_res.mem_productionE singletonI)

    ultimately have "(ground_resolution D)++ C (eres D C)"
      by (metis rtranclpD)

    then obtain n where "(ground_resolution D ^^ Suc n) C (eres D C)"
      by (metis not0_implies_Suc not_gr_zero tranclp_power)

    hence "resolvent_at D C (Suc n) = eres D C"
      by (metis Uniq_ground_resolution Uniq_relpowp resolvent_at_def the1_equality')

    have steps: "k  Suc n  (ord_res_2_step ^^ k)
      (N, Upr |∪| Uer, Uef) (N, Upr |∪| Uer |∪| resolvents_upto D C k, Uef)" for k
    proof (induction k)
      case 0
      show ?case
        by simp
    next
      case (Suc k)
      have "k < Suc n"
        using Suc k  Suc n by presburger
      hence "k  Suc n"
        by presburger
      hence "(ord_res_2_step ^^ k) (N, Upr |∪| Uer, Uef)
        (N, Upr |∪| Uer |∪| resolvents_upto D C k, Uef)"
        using Suc.IH by metis

      moreover have "ord_res_2_step
        (N, Upr |∪| Uer |∪| resolvents_upto D C k, Uef)
        (N, Upr |∪| Uer |∪| resolvents_upto D C (Suc k), Uef)"
        unfolding resolvents_upto_Suc
      proof (intro ord_res_2_step.intros ord_res_2.resolution)
        show "is_least_false_clause (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C k) |∪| Uef)
          (resolvent_at D C k)"
          using k < Suc n
        proof (induction k)
          case 0
          have "is_least_false_clause (N |∪| Upr |∪| Uer |∪| Uef) C"
            using is_least_false_clause (N |∪| Uer |∪| Uef) C
            unfolding least_false_conv .
          thus ?case
            unfolding funion_fempty_right funion_assoc[symmetric]
            by simp
        next
          case (Suc k')

          have "x. ord_res_Interp (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef)) x =
              ord_res_Interp (fset (N |∪| Uer |∪| Uef)  fset (Upr |∪| resolvents_upto D C (Suc k'))) x"
            by (simp add: funion_left_commute sup_assoc sup_commute)
          also have "x. ord_res_Interp (fset (N |∪| Uer |∪| Uef)  fset (Upr |∪| resolvents_upto D C (Suc k'))) x =
            ord_res_Interp (fset (N |∪| Uer |∪| Uef)) x"
          proof (intro Interp_union_unproductive ballI)
            fix x y assume "y |∈| Upr |∪| resolvents_upto D C (Suc k')"
            hence "y |∈| Upr  y |∈| resolvents_upto D C (Suc k')"
              by blast
            thus "ord_res.production (fset (N |∪| Uer |∪| Uef)  fset (Upr |∪| resolvents_upto D C (Suc k'))) y = {}"
            proof (elim disjE)
              assume "y |∈| Upr"
              thus ?thesis
                using Upr_unproductive by metis
            next
              assume "y |∈| resolvents_upto D C (Suc k')"
              then obtain i where "i |∈| fset_upto (Suc 0) (Suc k')" and "y = resolvent_at D C i"
                unfolding resolvents_upto_def by blast

              have "L. is_pos L  ord_res.is_strictly_maximal_lit L (resolvent_at D C i)"
              proof (rule nex_pos_strictly_max_lit_in_resolvent_at)
                show "(ground_resolution D ^^ Suc n) C (eres D C)"
                  using (ground_resolution D ^^ Suc n) C (eres D C) .
              next
                have "i  Suc k'"
                  using i |∈| fset_upto (Suc 0) (Suc k') by auto
                thus "i < Suc n"
                  using Suc k' < Suc n by presburger
              qed

              then show ?thesis
                using y = resolvent_at D C i unproductive_if_nex_strictly_maximal_pos_lit
                by metis
            qed
          qed simp_all
          finally have Interp_simp: "x.
            ord_res_Interp (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef)) x =
            ord_res_Interp (fset (N |∪| Uer |∪| Uef)) x" .

          show ?case
            unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
          proof (intro conjI ballI impI)
            have "resolvent_at D C (Suc k') |∈| resolvents_upto D C (Suc k')"
              using resolvent_at_fmember_resolvents_upto by simp
            thus "resolvent_at D C (Suc k') |∈| N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef"
              by simp
          next

            show "¬ ord_res_Interp (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef))
              (resolvent_at D C (Suc k'))  resolvent_at D C (Suc k')"
              unfolding Interp_simp
              by (metis (no_types, lifting) Suc.prems Zero_not_Suc
                  (ground_resolution D ^^ Suc n) C (eres D C) clause_true_if_resolved_true
                  insert_not_empty is_least_false_clause_def
                  linorder_cls.is_least_in_fset_ffilterD(2) local.resolution(2) local.resolution(7)
                  relpowp_to_resolvent_at tranclp_if_relpowp)
          next
            fix y
            assume "y  resolvent_at D C (Suc k')"
            assume "¬ ord_res_Interp (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef)) y  y"
            hence "¬ ord_res_Interp (fset (N |∪| Uer |∪| Uef)) y  y"
              unfolding Interp_simp .
            hence "¬ ord_res_Interp (fset (N |∪| Upr |∪| Uer |∪| Uef)) y  y"
              using Interp_N_Ur_Uef_eq_Interp_N_Uer_Uef by metis

            assume "y |∈| N |∪| (Upr |∪| Uer |∪| resolvents_upto D C (Suc k')) |∪| Uef"
            hence "y |∈| N |∪| Upr |∪| Uer |∪| Uef  y |∈| resolvents_upto D C (Suc k')"
              by auto
            thus "resolvent_at D C (Suc k') c y"
            proof (elim disjE)
              assume "y |∈| N |∪| Upr |∪| Uer |∪| Uef"
              have "C c y"
              proof (cases "y = C")
                case True
                thus ?thesis
                  by order
              next
                case False
                thus ?thesis
                  using y |∈| N |∪| Upr |∪| Uer |∪| Uef
                  using ¬ ord_res_Interp (fset (N |∪| Upr |∪| Uer |∪| Uef)) y  y
                  using is_least_false_clause (N |∪| Uer |∪| Uef) C
                  unfolding least_false_conv[symmetric]
                  unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
                  by simp
              qed

              moreover have "resolvent_at D C (Suc k') c C"
                by (metis Suc.prems (ground_resolution D ^^ Suc n) C (eres D C) less_or_eq_imp_le
                    resolvent_at_less_cls_resolvent_at resolvent_at_0 zero_less_Suc)

              ultimately show "resolvent_at D C (Suc k') c y"
                by order
            next
              assume "y |∈| resolvents_upto D C (Suc k')"
              then obtain i where
                i_in: "i |∈| fset_upto (Suc 0) (Suc k')" and y_def: "y = resolvent_at D C i"
                unfolding resolvents_upto_def by blast

              hence "i < Suc k'"
                using y  resolvent_at D C (Suc k')
                by auto

              show "resolvent_at D C (Suc k') c y"
                unfolding y_def
              proof (rule resolvent_at_less_cls_resolvent_at)
                show "(ground_resolution D ^^ Suc n) C (eres D C)"
                  using (ground_resolution D ^^ Suc n) C (eres D C) .
              next
                show "i < Suc k'"
                  using y  resolvent_at D C (Suc k') i_in y_def by auto
              next
                show "Suc k'  Suc n"
                  using Suc k' < Suc n by presburger
              qed
            qed
          qed
        qed
      next
        show "ord_res.is_maximal_lit L (resolvent_at D C k)"
        proof (rule max_lit_resolvent_at)
          show "(ground_resolution D ^^ Suc n) C (eres D C)"
            using (ground_resolution D ^^ Suc n) C (eres D C) .
        next
          show "k < Suc n"
            using k < Suc n .
        next
          show "ord_res.is_maximal_lit L C"
          using ord_res.is_maximal_lit L C .
        qed
      next
        show "is_neg L"
          using is_neg L .
      next
        show "D |∈| N |∪| (Upr |∪| Uer |∪| resolvents_upto D C k) |∪| Uef"
          using D |∈| N |∪| Uer |∪| Uef by auto
      next
        show "D c resolvent_at D C k"
        proof (rule left_premisse_lt_resolvent_at)
          show "(ground_resolution D ^^ Suc n) C (eres D C)"
            using (ground_resolution D ^^ Suc n) C (eres D C) .
        next
          show "k < Suc n"
            using k < Suc n .
        qed
      next
        have "ord_res.production (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C k) |∪| Uef)) D =
          ord_res.production (fset (N |∪| Uer |∪| Uef)  fset (Upr |∪| resolvents_upto D C k)) D"
          by (simp add: funion_left_commute sup_assoc sup_commute)
        also have " = ord_res.production (fset (N |∪| Uer |∪| Uef)) D"
        proof (intro production_union_unproductive ballI)
          fix x
          assume "x |∈| Upr |∪| resolvents_upto D C k"
          hence "L. is_pos L  ord_res.is_strictly_maximal_lit L x"
            unfolding funion_iff
          proof (elim disjE)
            assume "x |∈| Upr"
            thus ?thesis
              using Upr_spec
              by (metis eres_eq_after_tranclp_ground_resolution nex_strictly_maximal_pos_lit_if_neq_eres)
          next
            assume "x |∈| resolvents_upto D C k"
            then obtain i where "i |∈| fset_upto (Suc 0) k" and x_def: "x = resolvent_at D C i"
              unfolding resolvents_upto_def by auto

            have "0 < i" and "i  k"
              using i |∈| fset_upto (Suc 0) k by simp_all

            show ?thesis
              unfolding x_def
            proof (rule nex_pos_strictly_max_lit_in_resolvent_at)
              show "(ground_resolution D ^^ Suc n) C (eres D C)"
                using (ground_resolution D ^^ Suc n) C (eres D C) .
            next
              show "i < Suc n"
                using i  k k < Suc n by presburger
            qed
          qed
          thus "ord_res.production (fset (N |∪| Uer |∪| Uef) 
            fset (Upr |∪| resolvents_upto D C k)) x = {}"
            using unproductive_if_nex_strictly_maximal_pos_lit by metis
        next
          show "D |∈| N |∪| Uer |∪| Uef"
            using D |∈| N |∪| Uer |∪| Uef .
        qed simp_all
        finally show "ord_res.production (fset (N |∪| (Upr |∪| Uer |∪| resolvents_upto D C k) |∪| Uef)) D =
          {atm_of L}"
          using ord_res.production (fset (N |∪| Uer |∪| Uef)) D = {atm_of L} by argo
      next
        show "ord_res.ground_resolution (resolvent_at D C k) D (resolvent_at D C (Suc k))"
          unfolding ground_resolution_def[symmetric]
        proof (rule ground_resolution_resolvent_at_resolvent_at_Suc)
          show "(ground_resolution D ^^ Suc n) C (eres D C)"
            using (ground_resolution D ^^ Suc n) C (eres D C) .
        next
          show "k < Suc n"
            using k < Suc n .
        qed
      next
        show "Upr |∪| Uer |∪| finsert (resolvent_at D C (Suc k)) (resolvents_upto D C k) =
          finsert (resolvent_at D C (Suc k)) (Upr |∪| Uer |∪| resolvents_upto D C k)"
          by simp
      qed

      ultimately show ?case
        by (meson relpowp_Suc_I)
    qed

    hence "(ord_res_2_step ^^ Suc n) S2 (N, Upr |∪| Uer |∪| resolvents_upto D C (Suc n), Uef)"
      unfolding match_hyps(1,2) by blast

    moreover have "match (N, Upr |∪| Uer |∪| resolvents_upto D C (Suc n), Uef) S3'"
    proof -
      have 1: "S3' = (N, finsert (eres D C) Uer, Uef)"
        unfolding S3'_def s3' = (Urr', Uef) Urr' = finsert (eres D C) Uer ..

      have 2: "Upr |∪| Uer |∪| resolvents_upto D C (Suc n) =
        Upr |∪| resolvents_upto D C n |∪| finsert (eres D C) Uer"
        by (auto simp: resolvent_at D C (Suc n) = eres D C)

      show ?thesis
        unfolding match_def 1 2
      proof (rule ord_res_2_matches_ord_res_3.intros)
        show "E|∈|Upr |∪| resolvents_upto D C n.
          D1|∈|N |∪| finsert (eres D C) Uer |∪| Uef. D2|∈|N |∪| finsert (eres D C) Uer |∪| Uef.
          (ground_resolution D1)++ D2 E  E  eres D1 D2  eres D1 D2 |∈| finsert (eres D C) Uer"
        proof (intro ballI)
          fix Ca
          assume "Ca |∈| Upr |∪| resolvents_upto D C n"
          hence "Ca |∈| Upr  Ca |∈| resolvents_upto D C n"
            by simp
          thus "D1|∈|N |∪| finsert (eres D C) Uer |∪| Uef. D2|∈|N |∪| finsert (eres D C) Uer |∪| Uef.
            (ground_resolution D1)++ D2 Ca  Ca  eres D1 D2  eres D1 D2 |∈| finsert (eres D C) Uer"
          proof (elim disjE)
            show "Ca |∈| Upr  ?thesis"
              using Upr_spec by auto
          next
            assume "Ca |∈| resolvents_upto D C n"
            then obtain i where i_in: "i |∈| fset_upto (Suc 0) n" and Ca_def:"Ca = resolvent_at D C i"
              unfolding resolvents_upto_def by auto

            from i_in have "0 < i" "i  n"
              by simp_all

            show "?thesis"
            proof (intro bexI conjI)
              have "(ground_resolution D ^^ i) C Ca"
                unfolding Ca = resolvent_at D C i
              proof (rule relpowp_to_resolvent_at)
                show "(ground_resolution D ^^ Suc n) C (eres D C)"
                  using (ground_resolution D ^^ Suc n) C (eres D C) .
              next
                show "i < Suc n"
                  using i  n by presburger
              qed
              thus "(ground_resolution D)++ C Ca"
                using 0 < i by (simp add: tranclp_if_relpowp)
            next
              show "Ca  eres D C"
                by (metis Ca_def (ground_resolution D ^^ Suc n) C (eres D C)
                  x. ground_resolution D (eres D C) x i  n
                  ground_resolution_resolvent_at_resolvent_at_Suc less_Suc_eq_le)
            next
              show "eres D C |∈| finsert (eres D C) Uer"
                by simp
            next
              show "D |∈| N |∪| finsert (eres D C) Uer |∪| Uef"
                using resolution by simp
            next
              have "C |∈| N |∪| Uer |∪| Uef"
                using resolution
                by (simp add: is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff)
              thus "C |∈| N |∪| finsert (eres D C) Uer |∪| Uef"
                by simp
            qed
          qed
        qed
      qed
    qed

    ultimately have "S2'. (ord_res_2_step ^^ Suc n) S2 S2'  match S2' S3'"
      by metis

    thus "S2'. ord_res_2_step++ S2 S2'  match S2' S3'"
      by (metis Zero_neq_Suc tranclp_if_relpowp)
  qed
qed

lemma safe_states_if_ord_res_2_matches_ord_res_3:
  assumes match: "ord_res_2_matches_ord_res_3 S2 S3"
  shows
    "safe_state ord_res_2_step ord_res_2_final S2"
    "safe_state ord_res_3_step ord_res_3_final S3"
proof -
  show "safe_state ord_res_2_step ord_res_2_final S2"
    using safe_state_if_all_states_safe ord_res_2_step_safe by metis

  show "safe_state ord_res_3_step ord_res_3_final S3"
    using safe_state_if_all_states_safe ord_res_3_step_safe by metis
qed

theorem bisimulation_ord_res_2_ord_res_3:
  defines "match  λ_ S2 S3. ord_res_2_matches_ord_res_3 S2 S3"
  shows "(MATCH :: nat × nat  'f ord_res_2_state  'f ord_res_3_state  bool) .
    bisimulation ord_res_2_step ord_res_3_step ord_res_2_final ord_res_3_final  MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_2_state ⇒ 'f ord_res_3_state ⇒ bool) ℛfb.
    bisimulation ord_res_2_step ord_res_2_final ord_res_3_step ord_res_3_final MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique ord_res_2_step"
    using right_unique_ord_res_2_step .
next
  show "right_unique ord_res_3_step"
    using right_unique_ord_res_3_step .
next
  show "s1. ord_res_2_final s1  (s1'. ord_res_2_step s1 s1')"
    by (metis finished_def ord_res_2_semantics.final_finished)
next
  show "s2. ord_res_3_final s2  (s2'. ord_res_3_step s2 s2')"
    by (metis finished_def ord_res_3_semantics.final_finished)
next
  show "i s1 s2. match i s1 s2  ord_res_2_final s1 = ord_res_3_final s2"
    unfolding match_def
    using ord_res_2_final_iff_ord_res_3_final by metis
next
  show "i s1 s2. match i s1 s2 
    safe_state ord_res_2_step ord_res_2_final s1  safe_state ord_res_3_step ord_res_3_final s2"
    unfolding match_def
    using safe_states_if_ord_res_2_matches_ord_res_3 by metis
next
  show "wfP (λ_ _. False)"
    by simp
next
  show "i s1 s2 s2'.
       match i s1 s2 
       ord_res_3_step s2 s2' 
       (i' s1'. ord_res_2_step++ s1 s1'  match i' s1' s2')  (i'. match i' s1 s2'  False)"
    unfolding match_def
    using backward_simulation_2_to_3 by metis
qed

end


section ‹ORD-RES-4 (implicit factorization)›

type_synonym 'f ord_res_4_state = "'f gclause fset × 'f gclause fset × 'f gclause fset"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_3_matches_ord_res_4 :: "'f ord_res_3_state  'f ord_res_4_state  bool" where
  " |⊆| N |∪| Uer  Uef = iefac  |`| {|C |∈| N |∪| Uer. iefac  C  C|} 
  ord_res_3_matches_ord_res_4 (N, (Uer, Uef)) (N, Uer, )"

lemma ord_res_3_final_iff_ord_res_4_final:
  assumes match: "ord_res_3_matches_ord_res_4 S3 S4"
  shows "ord_res_3_final S3  ord_res_4_final S4"
  using match
proof (cases S3 S4 rule: ord_res_3_matches_ord_res_4.cases)
  case match_hyps: (1  N Uer Uef)
  note invars = match_hyps(3-)

  have "{#} |∈| N |∪| Uer |∪| Uef  {#} |∈| iefac  |`| (N |∪| Uer)"
    using invars by (auto simp: iefac_def)

  moreover have "ex_false_clause (fset (N |∪| Uer |∪| Uef)) 
    ex_false_clause (fset (iefac  |`| (N |∪| Uer)))"
    unfolding ex_false_clause_iff
    unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF invars(2)]
    unfolding is_least_false_clause_with_iefac_conv ..

  ultimately have "ord_res_final (N |∪| Uer |∪| Uef)  ord_res_final (iefac  |`| (N |∪| Uer))"
    unfolding ord_res_final_def by argo

  thus ?thesis
    unfolding match_hyps(1,2)
    by (simp add: ord_res_3_final.simps ord_res_4_final.simps)
qed

lemma forward_simulation_between_3_and_4:
  assumes
    match: "ord_res_3_matches_ord_res_4 S3 S4" and
    step: "ord_res_3_step S3 S3'"
  shows "(S4'. ord_res_4_step++ S4 S4'  ord_res_3_matches_ord_res_4 S3' S4')"
  using match
proof (cases S3 S4 rule: ord_res_3_matches_ord_res_4.cases)
  case match_hyps: (1  N Uer Uef)
  note match_invars = match_hyps(3-)

  from step obtain s3' where step': "ord_res_3 N (Uer, Uef) s3'" and "S3' = (N, s3')"
    unfolding match_hyps(1,2)
    by (auto elim: ord_res_3_step.cases)

  from step' show ?thesis
  proof (cases N "(Uer, Uef)" s3' rule: ord_res_3.cases)
    case (factoring C L Uef')

    have "¬ ord_res.is_strictly_maximal_lit L C"
      using is_least_false_clause (N |∪| Uer |∪| Uef) C ord_res.is_maximal_lit L C is_pos L
      by (metis (no_types, lifting) is_least_false_clause_def is_pos_def
        pos_lit_not_greatest_if_maximal_in_least_false_clause)

    have "C |∈| N |∪| Uer"
    proof -
      have "C |∈| N |∪| Uer |∪| Uef"
        using is_least_false_clause (N |∪| Uer |∪| Uef) C
        by (simp add: is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff)
      moreover have "C |∉| Uef"
      proof (rule notI)
        assume "C |∈| Uef"
        then obtain C0 where "C = iefac  C0" and "C0 |∈| N |∪| Uer" and "iefac  C0  C0"
          using match_invars(2) by force
        then show False
          by (metis Uniq_D ¬ ord_res.is_strictly_maximal_lit L C iefac_def
            linorder_lit.Uniq_is_maximal_in_mset
            linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset local.factoring(3)
            obtains_positive_greatest_lit_if_efac_not_ident)
      qed
      ultimately show ?thesis
        by simp
    qed

    show ?thesis
    proof (intro exI conjI)
      show "ord_res_4_step++ S4 (N, Uer, finsert C )"
        unfolding match_hyps(1,2)
      proof (intro tranclp.r_into_trancl ord_res_4_step.intros ord_res_4.factoring)
        have "is_least_false_clause (N |∪| Uer |∪| Uef) C"
          using factoring by argo
        hence "is_least_false_clause (N |∪| Uer |∪| iefac  |`| (N |∪| Uer)) C"
          unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
        thus "is_least_false_clause (iefac  |`| (N |∪| Uer)) C"
          unfolding is_least_false_clause_with_iefac_conv .
      next
        show "ord_res.is_maximal_lit L C"
          using ord_res.is_maximal_lit L C .
      next
        show "is_pos L"
          using is_pos L .
      qed (rule refl)+
    next
      show "ord_res_3_matches_ord_res_4 S3' (N, Uer, finsert C )"
        unfolding S3' = (N, s3') s3' = (Uer, Uef') Uef' = finsert (efac C) Uef
      proof (rule ord_res_3_matches_ord_res_4.intros)
        show "finsert C  |⊆| N |∪| Uer"
          using match_invars C |∈| N |∪| Uer by simp
      next
        have "C'. ord_res.ground_factoring C C'"
          using ord_res.is_maximal_lit L C is_pos L
          by (metis ¬ ord_res.is_strictly_maximal_lit L C ex_ground_factoringI is_pos_def)
        hence "efac C  C"
          by (metis ex1_efac_eq_factoring_chain)
        hence "iefac (finsert C ) C  C"
          by (simp add: iefac_def)

        have "{|Ca |∈| N |∪| Uer. iefac (finsert C ) Ca  Ca|} =
          finsert C {|Ca |∈| N |∪| Uer. iefac  Ca  Ca|}"
        proof (intro fsubset_antisym fsubsetI)
          fix x
          assume "x |∈| {|Ca |∈| N |∪| Uer. iefac (finsert C ) Ca  Ca|}"
          hence "x |∈| N |∪| Uer" and "iefac (finsert C ) x  x"
            by simp_all
          then show "x |∈| finsert C {|Ca |∈| N |∪| Uer. iefac  Ca  Ca|}"
            by (smt (verit, best) ffmember_filter finsert_iff iefac_def)
        next
          fix x
          assume "x |∈| finsert C {|Ca |∈| N |∪| Uer. iefac  Ca  Ca|}"
          hence "x = C  x |∈| N |∪| Uer  iefac  x  x"
            by auto
          thus "x |∈| {|Ca |∈| N |∪| Uer. iefac (finsert C ) Ca  Ca|}"
          proof (elim disjE conjE)
            assume "x = C"
            thus ?thesis
              using C |∈| N |∪| Uer iefac (finsert C ) C  C by auto
          next
            assume "x |∈| N |∪| Uer" and "iefac  x  x"
            thus ?thesis
              by (smt (verit, best) ffmember_filter finsertCI iefac_def)
          qed
        qed
        thus "finsert (efac C) Uef =
          iefac (finsert C ) |`| {|Ca |∈| N |∪| Uer. iefac (finsert C ) Ca  Ca|}"
          using iefac_def match_invars(2) by auto
      qed
    qed
  next
    case (resolution C L D Urr')

    have "D |∈| iefac  |`| (N |∪| Uer)"
    proof -
      have "D |∈| N |∪| Uer |∪| Uef"
        using resolution by argo
      hence "D |∈| N |∪| Uer |∪| iefac  |`| (N |∪| Uer)"
        unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
      moreover have "D |∉| N |∪| Uer - iefac  |`| (N |∪| Uer)"
        by (metis clauses_for_iefac_are_unproductive insert_not_empty local.resolution(7))
      ultimately show ?thesis
        by blast
    qed

    show ?thesis
    proof (intro exI conjI)
      show "ord_res_4_step++ S4 (N, finsert (eres D C) Uer, )"
        unfolding match_hyps(1,2)
        proof (intro tranclp.r_into_trancl ord_res_4_step.intros ord_res_4.resolution)
          have "is_least_false_clause (N |∪| Uer |∪| Uef) C"
            using resolution by argo
          hence "is_least_false_clause (N |∪| Uer |∪| iefac  |`| (N |∪| Uer)) C"
            unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] .
          thus "is_least_false_clause (iefac  |`| (N |∪| Uer)) C"
            unfolding is_least_false_clause_with_iefac_conv .
        next
          show "ord_res.is_maximal_lit L C"
            using resolution by argo
        next
          show "is_neg L"
            using resolution by argo
        next
          show "D |∈| iefac  |`| (N |∪| Uer)"
            using D |∈| iefac  |`| (N |∪| Uer) .
        next
          show "D c C"
            using resolution by argo
        next
          have "ord_res.production (fset (N |∪| Uer |∪| Uef)) D =
            ord_res.production (fset (N |∪| Uer |∪| iefac  |`| (N |∪| Uer))) D"
            unfolding funion_funion_eq_funion_funion_fimage_iefac_if[OF match_invars(2)] ..
          also have " = ord_res.production (fset (iefac  |`| (N |∪| Uer))  fset (N |∪| Uer)) D"
            by (simp add: sup.commute)
          also have " = ord_res.production (fset (iefac  |`| (N |∪| Uer))) D"
          proof (rule production_union_unproductive_strong)
            show "x  fset (N |∪| Uer) - fset (iefac  |`| (N |∪| Uer)).
              ord_res.production (fset (iefac  |`| (N |∪| Uer))  fset (N |∪| Uer)) x = {}"
              using clauses_for_iefac_are_unproductive[of "N |∪| Uer" ] by simp
          next
            show "D |∈| iefac  |`| (N |∪| Uer)"
              using D |∈| iefac  |`| (N |∪| Uer) .
          qed (rule finite_fset)+

          finally show "ord_res.production (fset (iefac  |`| (N |∪| Uer))) D = {atm_of L}"
            using resolution by argo
        qed (rule refl)+
    next
      show "ord_res_3_matches_ord_res_4 S3' (N, finsert (eres D C) Uer, )"
        unfolding S3' = (N, s3') s3' = (Urr', Uef) Urr' = finsert (eres D C) Uer
      proof (rule ord_res_3_matches_ord_res_4.intros)
        show " |⊆| N |∪| finsert (eres D C) Uer"
          using match_invars by auto
      next
        show "Uef = iefac  |`| {|C |∈| N |∪| finsert (eres D C) Uer. iefac  C  C|} "
        proof (cases "eres D C |∈| ")
          case True
          then show ?thesis
            using  |⊆| N |∪| Uer
            using match_invars by force
        next
          case False
          hence "iefac  (eres D C) = eres D C"
            by (simp add: iefac_def)
          hence "{|C |∈| N |∪| finsert (eres D C) Uer. iefac  C  C|} = {|C |∈| N |∪| Uer. iefac  C  C|}"
            using ffilter_eq_ffilter_minus_singleton by auto
          thus ?thesis
            using match_invars by argo
        qed
      qed
    qed
  qed
qed

theorem bisimulation_ord_res_3_ord_res_4:
  defines "match  λ_ S3 S4. ord_res_3_matches_ord_res_4 S3 S4"
  shows "(MATCH :: nat × nat  'f ord_res_3_state  'f ord_res_4_state  bool) .
    bisimulation ord_res_3_step ord_res_4_step ord_res_3_final ord_res_4_final  MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_3_state ⇒ 'f ord_res_4_state ⇒ bool) ℛfb.
    bisimulation ord_res_3_step ord_res_3_final ord_res_4_step ord_res_4_final MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_forward_simulation)
  show "right_unique ord_res_3_step"
    using right_unique_ord_res_3_step .
next
  show "right_unique ord_res_4_step"
    using right_unique_ord_res_4_step .
next
  show "s1. ord_res_3_final s1  (s1'. ord_res_3_step s1 s1')"
    by (metis finished_def ord_res_3_semantics.final_finished)
next
  show "s2. ord_res_4_final s2  (s2'. ord_res_4_step s2 s2')"
    by (metis finished_def ord_res_4_semantics.final_finished)
next
  show "i s1 s2. match i s1 s2  ord_res_3_final s1  ord_res_4_final s2"
    unfolding match_def
    using ord_res_3_final_iff_ord_res_4_final by metis
next
  show "i s1 s2. match i s1 s2 
    safe_state ord_res_3_step ord_res_3_final s1  safe_state ord_res_4_step ord_res_4_final s2"
    using ord_res_3_step_safe ord_res_4_step_safe
    by (simp add: safe_state_if_all_states_safe)
next
  show "wfP (λi' i. False)"
    by simp
next
  show "i s1 s2 s1'. match i s1 s2  ord_res_3_step s1 s1' 
    (i' s2'. ord_res_4_step++ s2 s2'  match i' s1' s2')  (i'. match i' s1' s2  False)"
    unfolding match_def
    using forward_simulation_between_3_and_4 by metis
qed

end


section ‹ORD-RES-5 (explicit model construction)›

type_synonym 'f ord_res_5_state = "'f gclause fset × 'f gclause fset × 'f gclause fset ×
  ('f gterm  'f gclause option) × 'f gclause option"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_4_matches_ord_res_5 :: "'f ord_res_4_state  'f ord_res_5_state  bool" where
  "ord_res_5_invars N (Uer, , , 𝒞) 
    (C. 𝒞 = Some C  is_least_false_clause (iefac  |`| (N |∪| Uer)) C) 
    ord_res_4_matches_ord_res_5 (N, Uer, ) (N, Uer, , , 𝒞)"

lemma ord_res_4_final_iff_ord_res_5_final:
  assumes match: "ord_res_4_matches_ord_res_5 S4 S5"
  shows "ord_res_4_final S4  ord_res_5_final S5"
  using match
proof (cases S4 S5 rule: ord_res_4_matches_ord_res_5.cases)
  case match_hyps: (1 N Uer   𝒞)

  show ?thesis
    unfolding match_hyps(1,2,3)
  proof (intro iffI ord_res_5_final.intros)
    assume "ord_res_4_final (N, Uer, )"
    hence "{#} |∈| iefac  |`| (N |∪| Uer)  ¬ ex_false_clause (fset (iefac  |`| (N |∪| Uer)))"
      by (simp add: ord_res_4_final.simps ord_res_final_def)
    thus "ord_res_5_final (N, Uer, , , 𝒞)"
    proof (elim disjE)
      assume "{#} |∈| iefac  |`| (N |∪| Uer)"
      hence "is_least_false_clause (iefac  |`| (N |∪| Uer)) {#}"
        using is_least_false_clause_mempty by metis
      hence "𝒞 = Some {#}"
        by (smt (verit) all_smaller_clauses_true_wrt_respective_Interp_def is_least_false_clause_def
            linorder_cls.is_least_in_ffilter_iff linorder_cls.le_imp_less_or_eq match_hyps(3)
            mempty_lesseq_cls ord_res_5_invars_def)
      thus ?thesis
        using ord_res_5_final.contradiction_found by metis
    next
      assume "¬ ex_false_clause (fset (iefac  |`| (N |∪| Uer)))"
      hence "𝒞 = None"
        using match_hyps(2-)
        by (metis ex_false_clause_if_least_false_clause option.exhaust)
      thus ?thesis
        using ord_res_5_final.model_found by metis
    qed
  next
    assume "ord_res_5_final (N, Uer, , , 𝒞)"
    thus "ord_res_4_final (N, Uer, )"
    proof (cases "(N, Uer, , , 𝒞)" rule: ord_res_5_final.cases)
      case model_found
      have "all_smaller_clauses_true_wrt_respective_Interp N (Uer, , , 𝒞)"
        using ord_res_5_invars N (Uer, , , 𝒞)
        unfolding ord_res_5_invars_def by metis
      hence "C |∈| iefac  |`| (N |∪| Uer). ord_res_Interp (iefac  ` (fset N  fset Uer)) C  C"
        by (simp add: model_found all_smaller_clauses_true_wrt_respective_Interp_def)
      hence "¬ ex_false_clause (fset (iefac  |`| (N |∪| Uer)))"
        by (simp add: ex_false_clause_def)
      then show ?thesis
        by (metis ord_res_4_final.intros ord_res_final_def)
    next
      case contradiction_found
      hence "{#} |∈| iefac  |`| (N |∪| Uer)"
        using ord_res_5_invars N (Uer, , , 𝒞)
        by (metis next_clause_in_factorized_clause_def ord_res_5_invars_def)
      then show ?thesis
        by (metis ord_res_4_final.intros ord_res_final_def)
    qed
  qed
qed

lemma forward_simulation_between_4_and_5:
  fixes S4 S4' S5
  assumes match: "ord_res_4_matches_ord_res_5 S4 S5" and step: "ord_res_4_step S4 S4'"
  shows "S5'. ord_res_5_step++ S5 S5'  ord_res_4_matches_ord_res_5 S4' S5'"
  using match
proof (cases S4 S5 rule: ord_res_4_matches_ord_res_5.cases)
  case match_hyps: (1 N Uer   𝒞)
  hence
    S4_def: "S4 = (N, Uer, )" and
    S5_def: "S5 = (N, Uer, , , 𝒞)"
    unfolding atomize_conj by metis

  have dom_ℳ_eq: "C. 𝒞 = Some C  dom  = ord_res.interp (fset (iefac  |`| (N |∪| Uer))) C"
    using match_hyps unfolding ord_res_5_invars_def model_eq_interp_upto_next_clause_def by simp

  obtain s4' where S4'_def: "S4' = (N, s4')" and step': "ord_res_4 N (Uer, ) s4'"
    using step unfolding S4_def by (auto simp: ord_res_4_step.simps)

  show ?thesis
    using step'
  proof (cases N "(Uer, )" s4' rule: ord_res_4.cases)
    case step_hyps: (factoring NN C L ℱ')
    have "𝒞 = Some C"
      using match_hyps(3-) step_hyps by metis

    define ℳ' :: "'f gterm  'f gterm literal multiset option" where
      "ℳ' = (λ_. None)"

    define 𝒞' :: "'f gclause option" where
      "𝒞' = The_optional (linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer)))"

    have ord_res_5_step: "ord_res_5 N (Uer, , , Some C) (Uer, ℱ', ℳ', 𝒞')"
    proof (rule ord_res_5.factoring)
      have "is_least_false_clause (iefac  |`| (N |∪| Uer)) C"
        using step_hyps by argo
      then show "¬ dom   C"
        using dom_ℳ_eq[OF 𝒞 = Some C]
        by (metis (mono_tags, lifting) is_least_false_clause_def
            linorder_cls.is_least_in_ffilter_iff ord_res_Interp_entails_if_greatest_lit_is_pos
            unproductive_if_nex_strictly_maximal_pos_lit sup_bot.right_neutral)
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by metis
    next
      show "is_pos L"
        using step_hyps by metis
    next
      show "¬ ord_res.is_strictly_maximal_lit L C"
        using step_hyps
        by (metis (no_types, lifting) is_least_false_clause_def literal.collapse(1)
            pos_lit_not_greatest_if_maximal_in_least_false_clause)
    next
      show "ℱ' = finsert C "
        using step_hyps by metis
    qed (simp_all add: ℳ'_def 𝒞'_def)

    moreover have "ℳ'' 𝒞''.
       (ord_res_5 N)** (Uer, ℱ', ℳ', 𝒞') (Uer, ℱ', ℳ'', 𝒞'') 
       (C. (𝒞'' = Some C)  is_least_false_clause (iefac ℱ' |`| (N |∪| Uer)) C)"
    proof (rule ord_res_5_construct_model_upto_least_false_clause)
      show "ord_res_5_invars N (Uer, ℱ', ℳ', 𝒞')"
        using ord_res_5_step ord_res_5_invars N (Uer, , , 𝒞) 𝒞 = Some C
        by (metis ord_res_5_preserves_invars)
    qed

    ultimately obtain ℳ'' 𝒞'' where
      s5_steps: "(ord_res_5 N)++ (Uer, , , Some C) (Uer, ℱ', ℳ'', 𝒞'')" and
      next_clause_least_false:
        "(C. (𝒞'' = Some C)  is_least_false_clause (iefac ℱ' |`| (N |∪| Uer)) C)"
      by (meson rtranclp_into_tranclp2)

    have "ord_res_5_step++ S5 (N, Uer, ℱ', ℳ'', 𝒞'')"
      unfolding S5_def 𝒞 = Some C
      using s5_steps by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)

    moreover have "ord_res_4_matches_ord_res_5 S4' (N, Uer, ℱ', ℳ'', 𝒞'')"
      unfolding S4'_def s4' = (Uer, ℱ')
    proof (intro ord_res_4_matches_ord_res_5.intros)
      show "ord_res_5_invars N (Uer, ℱ', ℳ'', 𝒞'')"
        using s5_steps 𝒞 = Some C ord_res_5_invars N (Uer, , , 𝒞)
        by (smt (verit, best) ord_res_5_preserves_invars tranclp_induct)
    next
      show "C. (𝒞'' = Some C) = is_least_false_clause (iefac ℱ' |`| (N |∪| Uer)) C"
        using next_clause_least_false .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (resolution NN C L D Uer')
    have "𝒞 = Some C"
      using match_hyps(3-) step_hyps by metis

    define ℳ' :: "'f gterm  'f gterm literal multiset option" where
      "ℳ' = (λ_. None)"

    define 𝒞' :: "'f gclause option" where
      "𝒞' = The_optional (linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')))"

    have ord_res_5_step: "ord_res_5 N (Uer, , , Some C) (Uer', , ℳ', 𝒞')"
    proof (rule ord_res_5.resolution)
      have "is_least_false_clause (iefac  |`| (N |∪| Uer)) C"
        using step_hyps by argo
      then show "¬ dom   C"
        using dom_ℳ_eq[OF 𝒞 = Some C]
        by (metis (mono_tags, lifting) is_least_false_clause_def
            linorder_cls.is_least_in_ffilter_iff ord_res_Interp_entails_if_greatest_lit_is_pos
            unproductive_if_nex_strictly_maximal_pos_lit sup_bot.right_neutral)
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by metis
    next
      show "is_neg L"
        using step_hyps by metis
    next
      show " (atm_of L) = Some D"
        using step_hyps
        by (smt (verit) 𝒞 = Some C all_produced_atoms_in_model_def insertI1 match_hyps(3)
            ord_res_5_invars_def)
    next
      show "Uer' = finsert (eres D C) Uer"
        using step_hyps by metis
    qed (simp_all add: ℳ'_def 𝒞'_def)

    moreover have "ℳ'' 𝒞''.
       (ord_res_5 N)** (Uer', , ℳ', 𝒞') (Uer', , ℳ'', 𝒞'') 
       (C. (𝒞'' = Some C)  is_least_false_clause (iefac  |`| (N |∪| Uer')) C)"
    proof (rule ord_res_5_construct_model_upto_least_false_clause)
      show "ord_res_5_invars N (Uer', , ℳ', 𝒞')"
        using ord_res_5_step ord_res_5_invars N (Uer, , , 𝒞) 𝒞 = Some C
        by (metis ord_res_5_preserves_invars)
    qed

    ultimately obtain ℳ'' 𝒞'' where
      s5_steps: "(ord_res_5 N)++ (Uer, , , Some C) (Uer', , ℳ'', 𝒞'')" and
      next_clause_least_false:
        "(C. (𝒞'' = Some C)  is_least_false_clause (iefac  |`| (N |∪| Uer')) C)"
      by (meson rtranclp_into_tranclp2)

    have "ord_res_5_step++ S5 (N, Uer', , ℳ'', 𝒞'')"
      unfolding S5_def 𝒞 = Some C
      using s5_steps by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)

    moreover have "ord_res_4_matches_ord_res_5 S4' (N, Uer', , ℳ'', 𝒞'')"
      unfolding S4'_def s4' = (Uer', )
    proof (intro ord_res_4_matches_ord_res_5.intros)
      show "ord_res_5_invars N (Uer', , ℳ'', 𝒞'')"
        using s5_steps 𝒞 = Some C ord_res_5_invars N (Uer, , , 𝒞)
        by (smt (verit, best) ord_res_5_preserves_invars tranclp_induct)
    next
      show "C. (𝒞'' = Some C) = is_least_false_clause (iefac  |`| (N |∪| Uer')) C"
        using next_clause_least_false .
    qed

    ultimately show ?thesis
      by metis
  qed
qed

theorem bisimulation_ord_res_4_ord_res_5:
  defines "match  λ_. ord_res_4_matches_ord_res_5"
  shows "(MATCH :: nat × nat  'f ord_res_4_state  'f ord_res_5_state  bool) .
    bisimulation ord_res_4_step ord_res_5_step ord_res_4_final ord_res_5_final  MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_4_state ⇒ 'f ord_res_5_state ⇒ bool) ℛfb.
    bisimulation ord_res_4_step ord_res_4_final ord_res_5_step ord_res_5_final MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_forward_simulation)
  show "right_unique ord_res_4_step"
    using right_unique_ord_res_4_step .
next
  show "right_unique ord_res_5_step"
    using right_unique_ord_res_5_step .
next
  show "s. ord_res_4_final s  (s'. ord_res_4_step s s')"
    by (metis finished_def ord_res_4_semantics.final_finished)
next
  show "s. ord_res_5_final s  (s'. ord_res_5_step s s')"
    by (metis finished_def ord_res_5_semantics.final_finished)
next
  show "i s4 s5. match i s4 s5  ord_res_4_final s4  ord_res_5_final s5"
    unfolding match_def
    using ord_res_4_final_iff_ord_res_5_final by metis
next
  show "i S4 S5. match i S4 S5 
    safe_state ord_res_4_step ord_res_4_final S4  safe_state ord_res_5_step ord_res_5_final S5"
  proof (intro allI impI conjI)
    fix i S4 S5
    show "safe_state ord_res_4_step ord_res_4_final S4"
      using ord_res_4_step_safe safe_state_if_all_states_safe by metis

    assume "match i S4 S5"
    thus "safe_state ord_res_5_step ord_res_5_final S5"
      using match i S4 S5
      using ord_res_5_safe_state_if_invars
      using match_def ord_res_4_matches_ord_res_5.cases by metis
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i s1 s2 s1'.
       match i s1 s2 
       ord_res_4_step s1 s1' 
       (i' s2'. ord_res_5_step++ s2 s2'  match i' s1' s2')  (i'. match i' s1' s2  False)"
    unfolding match_def
    using forward_simulation_between_4_and_5 by metis
qed

end


section ‹ORD-RES-6 (model backjump)›

type_synonym 'f ord_res_6_state = "'f gclause fset × 'f gclause fset × 'f gclause fset ×
  ('f gterm  'f gclause option) × 'f gclause option"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_5_matches_ord_res_6 :: "'f ord_res_5_state  'f ord_res_6_state  bool" where
  "ord_res_5_invars N (Uer, , , 𝒞) 
    ord_res_5_matches_ord_res_6 (N, Uer, , , 𝒞) (N, Uer, , , 𝒞)"

lemma ord_res_5_final_iff_ord_res_6_final:
  fixes i S5 S6
  assumes match: "ord_res_5_matches_ord_res_6 S5 S6"
  shows "ord_res_5_final S5  ord_res_6_final S6"
  using match
proof (cases S5 S6 rule: ord_res_5_matches_ord_res_6.cases)
  case (1 N Uer   𝒞)
  thus ?thesis
    by (metis (no_types, opaque_lifting) ord_res_5_final.simps ord_res_6_final.cases
        ord_res_6_final.contradiction_found ord_res_6_final.model_found)
qed

lemma backward_simulation_between_5_and_6:
  fixes S5 S6 S6'
  assumes match: "ord_res_5_matches_ord_res_6 S5 S6" and step: "ord_res_6_step S6 S6'"
  shows "S5'. ord_res_5_step++ S5 S5'  ord_res_5_matches_ord_res_6 S5' S6'"
  using match
proof (cases S5 S6 rule: ord_res_5_matches_ord_res_6.cases)
  case match_hyps: (1 N Uer   𝒞)
  hence S5_def: "S5 = (N, Uer, , , 𝒞)" and S6_def: "S6 = (N, Uer, , , 𝒞)"
    by metis+

  obtain s6' where S6'_def: "S6' = (N, s6')" and step': "ord_res_6 N (Uer, , , 𝒞) s6'"
    using step unfolding S6_def
    using ord_res_6_step.simps by auto

  show ?thesis
    using step'
  proof (cases N "(Uer, , , 𝒞)" s6' rule: ord_res_6.cases)
    case step_hyps: (skip C 𝒞')

    define S5' where
      "S5' = (N, Uer, , , 𝒞')"

    show ?thesis
    proof (intro exI conjI)
      have step5: "ord_res_5 N (Uer, , , Some C) (Uer, , , 𝒞')"
        using ord_res_5.skip step_hyps by metis
      hence "ord_res_5_step S5 S5'"
        unfolding S5_def S5'_def
        by (metis ord_res_5_step.simps step_hyps(1))
      thus "ord_res_5_step++ S5 S5'"
        by simp

      have "ord_res_5_invars N (Uer, , , 𝒞')"
        using step5 match_hyps(3) ord_res_5_preserves_invars step_hyps(1) by metis
      thus "ord_res_5_matches_ord_res_6 S5' S6'"
        unfolding S5'_def S6'_def s6' = (Uer, , , 𝒞')
        using ord_res_5_matches_ord_res_6.intros by metis
    qed
  next
    case step_hyps: (production C L ℳ' 𝒞')

    define S5' where
      "S5' = (N, Uer, , ℳ', 𝒞')"

    show ?thesis
    proof (intro exI conjI)
      have step5: "ord_res_5 N (Uer, , , Some C) (Uer, , ℳ', 𝒞')"
        using ord_res_5.production step_hyps by metis
      hence "ord_res_5_step S5 S5'"
        unfolding S5_def S5'_def
        by (metis ord_res_5_step.simps step_hyps(1))
      thus "ord_res_5_step++ S5 S5'"
        by simp

      have "ord_res_5_invars N (Uer, , ℳ', 𝒞')"
        using step5 match_hyps(3) ord_res_5_preserves_invars step_hyps(1) by metis
      thus "ord_res_5_matches_ord_res_6 S5' S6'"
        unfolding S5'_def S6'_def s6' = (Uer, , ℳ', 𝒞')
        using ord_res_5_matches_ord_res_6.intros by metis
    qed
  next
    case step_hyps: (factoring D K ℱ')

    define S5' where
      "S5' = (N, Uer, ℱ', , Some (efac D))"

    have "D |∈| iefac  |`| (N |∪| Uer)"
      by (metis match_hyps(3) next_clause_in_factorized_clause_def ord_res_5_invars_def step_hyps(1))
    hence "iefac ℱ' |`| (N |∪| Uer)  {||}"
      by blast
    then obtain C where C_least: "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer)) C"
      by (metis linorder_cls.ex1_least_in_fset)

    have "efac D  D"
      by (metis ex1_efac_eq_factoring_chain is_pos_def ex_ground_factoringI step_hyps(4,5,6))

    show ?thesis
    proof (intro exI conjI)
      have "The_optional (linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer))) = Some C"
      proof (rule The_optional_eq_SomeI)
        show "1 x. linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer)) x"
          by blast
      next
        show "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer)) C"
          using C_least .
      qed
      hence step5: "ord_res_5 N (Uer, , , Some D) (Uer, ℱ', Map.empty, Some C)"
        using ord_res_5.factoring step_hyps by metis
      moreover have "(ord_res_5 N)**  (Uer, ℱ', , Some (efac D))"
      proof (rule full_rtranclp_ord_res_5_run_upto)
        show "ord_res_6 N (Uer, , , Some D) (Uer, ℱ', , Some (efac D))"
          using step' S6_def S6'_def s6' = (Uer, ℱ', , Some (efac D)) 𝒞 = Some D by argo
      next
        show "ord_res_5_invars N (Uer, ℱ', , Some (efac D))"
          using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
      next
        have "iefac  D = D" and "D |∈| N |∪| Uer"
          unfolding atomize_conj
          using efac D  D D |∈| iefac  |`| (N |∪| Uer)[unfolded fimage_iff]
          unfolding iefac_def
          by (metis ex1_efac_eq_factoring_chain factorizable_if_neq_efac)

        have iefac_ℱ'_eq: "iefac ℱ' = (iefac )(D := efac D)"
          unfolding ℱ' = finsert D  iefac_def by auto

        have fimage_iefac_ℱ'_eq:
          "iefac ℱ' |`| (N |∪| Uer) = finsert (efac D) (iefac  |`| (N |∪| Uer - {|D|}))"
          unfolding iefac_ℱ'_eq
          unfolding fun_upd_fimage[of "iefac " D "efac D"] D |∈| N |∪| Uer
          using D |∈| N |∪| Uer by argo

        have "{|C |∈| iefac ℱ' |`| (N |∪| Uer). C c efac D|} =
          {|C |∈| finsert (efac D) (iefac  |`| (N |∪| Uer - {|D|})). C c efac D|}"
          unfolding fimage_iefac_ℱ'_eq ..

        also have " = {|C |∈| iefac  |`| (N |∪| Uer - {|D|}). C c efac D|}"
          by auto

        also have " = {|C |∈| iefac  |`| (N |∪| Uer). C c efac D|}"
          by (smt (verit, ccfv_SIG) iefac  D = D efac_properties_if_not_ident(1)
              ffilter_eq_ffilter_minus_singleton fimage_finsert finsertI1 finsert_fminus1
              finsert_fminus_single linorder_cls.less_imp_not_less)

        finally have "{|C |∈| iefac ℱ' |`| (N |∪| Uer). C c efac D|} =
          {|C |∈| iefac  |`| (N |∪| Uer). C c efac D|}" .
      next
        have dom_ℳ_eq: "dom  = ord_res.interp (fset (iefac  |`| (N |∪| Uer))) D"
          using ord_res_5_invars N (Uer, , , 𝒞) 𝒞 = Some D
          unfolding ord_res_5_invars_def model_eq_interp_upto_next_clause_def
          by metis

        have "atm_of K  dom "
          by (metis linorder_lit.is_maximal_in_mset_iff literal.collapse(1)
              pos_literal_in_imp_true_cls step_hyps(3) step_hyps(4) step_hyps(5))

        have "A t atm_of K" if "A  dom " for A
        proof -
          obtain C where
            "C |∈| iefac  |`| (N |∪| Uer)" and
            "C c D" and
            "A  ord_res.production (fset (iefac  |`| (N |∪| Uer))) C"
            using A  dom  unfolding dom_ℳ_eq
            unfolding ord_res.interp_def UN_iff
            by blast

          hence "ord_res.is_strictly_maximal_lit (Pos A) C"
            using ord_res.mem_productionE by metis

          hence "Pos A l K"
            using ord_res.is_maximal_lit K D C c D
            by (metis ord_res.asymp_less_lit ord_res.transp_less_lit linorder_cls.less_asym
                linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset linorder_lit.leI
                linorder_lit.multpHO_if_maximal_less_that_maximal multp_eq_multpHO)

          hence "A t atm_of K"
            by (metis literal.collapse(1) literal.sel(1) ord_res.less_lit_simps(1) reflclp_iff
                step_hyps(5))

          moreover have "A  atm_of K"
            using atm_of K  dom  A  dom  by metis

          ultimately show ?thesis
            by order
        qed
        hence "dom   {A. K. ord_res.is_maximal_lit K (efac D)  A t atm_of K}"
          using linorder_lit.is_maximal_in_mset_iff step_hyps(4) by auto
        thus " = restrict_map  {A. K. ord_res.is_maximal_lit K (efac D)  A t atm_of K}"
          using restrict_map_ident_if_dom_subset by fastforce
      next
        show "linorder_cls.is_least_in_fset (iefac ℱ' |`| (N |∪| Uer)) C"
          using C_least .
      qed
      ultimately have steps5: "(ord_res_5 N)++ (Uer, , , Some D) (Uer, ℱ', , Some (efac D))"
        by simp
      thus "ord_res_5_step++ S5 S5'"
        using S5'_def S5_def step_hyps(1) tranclp_ord_res_5_step_if_tranclp_ord_res_5 by metis

      have "ord_res_5_invars N (Uer, ℱ', , Some (efac D))"
        using steps5 match_hyps(3) tranclp_ord_res_5_preserves_invars step_hyps(1) by metis
      thus "ord_res_5_matches_ord_res_6 S5' S6'"
        unfolding S5'_def S6'_def s6' = (Uer, ℱ', , Some (efac D))
        using ord_res_5_matches_ord_res_6.intros by metis
    qed
  next
    case step_hyps: (resolution_bot C L D Uer' ℳ')

    define S5' :: "_ × _ × _ × ('f gterm  'f gclause option) × 'f gclause option" where
      "S5' = (N, Uer', , ℳ', Some {#})"

    show ?thesis
    proof (intro exI conjI)
      have "{#} |∈| iefac  |`| (N |∪| Uer')"
        using Uer' = finsert (eres D C) Uer eres D C = {#}
        using iefac_def by simp

      hence "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) {#}"
        by (metis linorder_cls.is_minimal_in_fset_eq_is_least_in_fset
            linorder_cls.is_minimal_in_fset_iff linorder_cls.leD mempty_lesseq_cls)

      hence "The_optional (linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer'))) = Some {#}"
        by (metis linorder_cls.Uniq_is_least_in_fset The_optional_eq_SomeI)

      hence step5: "ord_res_5 N (Uer, , , Some C) (Uer', , ℳ', Some {#})"
        using ord_res_5.resolution step_hyps by metis

      thus "ord_res_5_step++ S5 S5'"
        unfolding S5_def 𝒞 = Some C S5'_def
        by (simp only: ord_res_5_step.intros tranclp.r_into_trancl)

      show "ord_res_5_matches_ord_res_6 S5' S6'"
        using step5
        by (metis S5'_def S6'_def match_hyps(3) ord_res_5_matches_ord_res_6.intros
            ord_res_5_preserves_invars step_hyps(1) step_hyps(2))
    qed
  next
    case step_hyps: (resolution_pos E L D Uer' ℳ' K)

    define S5' :: "_ × _ × _ × ('f gterm  'f gclause option) × 'f gclause option" where
      "S5' = (N, Uer', , ℳ', Some (eres D E))"

    hence "iefac  |`| (N |∪| Uer')  {||}"
      using Uer' = finsert (eres D E) Uer by simp
    then obtain C where C_least: "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) C"
      by (metis linorder_cls.ex1_least_in_fset)

    show ?thesis
    proof (intro exI conjI)
      have "The_optional (linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer'))) = Some C"
      proof (rule The_optional_eq_SomeI)
        show "1 x. linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) x"
          by blast
      next
        show "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) C"
          using C_least .
      qed

      hence step5: "ord_res_5 N (Uer, , , Some E) (Uer', , Map.empty, Some C)"
        using ord_res_5.resolution step_hyps by metis

      moreover have "(ord_res_5 N)**  (Uer', , ℳ', Some (eres D E))"
      proof (rule full_rtranclp_ord_res_5_run_upto)
        show "ord_res_6 N (Uer, , , Some E) (Uer', , ℳ', Some (eres D E))"
          using step' 𝒞 = Some E s6' = (Uer', , ℳ', Some (eres D E)) by argo
      next
        show "ord_res_5_invars N (Uer', , ℳ', Some (eres D E))"
          using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
      next
        have "eres D E  E"
          using step_hyps by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')

        moreover have "eres D E c E"
          using eres_le .

        ultimately have "eres D E c E"
          by order

        have "F. is_least_false_clause (iefac  |`| (N |∪| Uer)) F  E c F"
          using ord_res_5_invars N (Uer, , , 𝒞)
          unfolding ord_res_5_invars_def 𝒞 = Some E
          using next_clause_lt_least_false_clause[of N "(Uer, , , Some E)"]
          by simp

        have E_least_false: "is_least_false_clause (iefac  |`| (N |∪| Uer)) E"
          unfolding is_least_false_clause_def linorder_cls.is_least_in_ffilter_iff
        proof (intro conjI ballI impI)
          show "E |∈| iefac  |`| (N |∪| Uer)"
            using ord_res_5_invars N (Uer, , , 𝒞)
            unfolding ord_res_5_invars_def 𝒞 = Some E
            by (metis next_clause_in_factorized_clause_def)
        next
          have "¬ ord_res.interp (fset (iefac  |`| (N |∪| Uer))) E  E"
            using ord_res_5_invars N (Uer, , , 𝒞)
            unfolding ord_res_5_invars_def 𝒞 = Some E
            using ¬ dom   E by (metis model_eq_interp_upto_next_clause_def)
          moreover have "ord_res.production (fset (iefac  |`| (N |∪| Uer))) E = {}"
          proof -
            have "L. is_pos L  ord_res.is_strictly_maximal_lit L E"
              using ord_res.is_maximal_lit L E is_neg L
              by (metis Uniq_D linorder_lit.Uniq_is_maximal_in_mset
                  linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset)
            thus ?thesis
              using unproductive_if_nex_strictly_maximal_pos_lit by metis
          qed
          ultimately show "¬ ord_res_Interp (fset (iefac  |`| (N |∪| Uer))) E  E"
            by simp
        next
          fix F
          assume F_in: "F |∈| iefac  |`| (N |∪| Uer)" and "F  E" and
            F_false: "¬ ord_res_Interp (fset (iefac  |`| (N |∪| Uer))) F  F"
          have "¬ F c E"
            using ord_res_5_invars N (Uer, , , 𝒞)
            unfolding ord_res_5_invars_def 𝒞 = Some E
            unfolding all_smaller_clauses_true_wrt_respective_Interp_def
            using F_in F_false
            by (metis option.inject)
          thus "E c F"
            using F  E by order
        qed

        have D_prod: "ord_res.production (fset (iefac  |`| (N |∪| Uer))) D  {}"
          using ord_res_5_invars N (Uer, , , 𝒞)
          unfolding ord_res_5_invars_def 𝒞 = Some E
          by (metis atoms_in_model_were_produced_def empty_iff step_hyps(6))

        have "iefac  (eres D E) = eres D E"
          using E_least_false D_prod
          by (smt (verit, ccfv_threshold)
              F. is_least_false_clause (iefac  |`| (N |∪| Uer)) F  (≺c)== E F
              eres D E c E clause_true_if_resolved_true ex1_eres_eq_full_run_ground_resolution
              fimage_finsert finsert_absorb finsert_iff full_run_def funion_finsert_right
              is_least_false_clause_def is_least_false_clause_finsert_smaller_false_clause
              linorder_cls.is_least_in_fset_ffilterD(2) linorder_cls.leD match_hyps(3)
              next_clause_in_factorized_clause_def ord_res_5_invars_def ord_res_6_preserves_invars
              rtranclpD step' step_hyps(2) step_hyps(7))

        hence "{|C |∈| iefac  |`| (N |∪| Uer'). C c eres D E|} =
          {|C |∈| iefac  |`| (N |∪| Uer). C c eres D E|}"
          unfolding Uer' = finsert (eres D E) Uer by auto
      next
        show "ℳ' = restrict_map  {A. K. ord_res.is_maximal_lit K (eres D E)  A t atm_of K}"
          using ℳ' = restrict_map  {A. A t atm_of K}
          by (smt (verit, ccfv_SIG) Collect_cong linorder_lit.Uniq_is_maximal_in_mset step_hyps(10)
              the1_equality')
      next
        show "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) C"
          using C_least .
      qed

      ultimately have steps5: "(ord_res_5 N)++ (Uer, , , Some E) (Uer', , ℳ', Some (eres D E))"
        by simp

      thus "ord_res_5_step++ S5 S5'"
        unfolding S5_def 𝒞 = Some E S5'_def
        by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)

      show "ord_res_5_matches_ord_res_6 S5' S6'"
        unfolding S5'_def S6'_def s6' = (Uer', , ℳ', Some (eres D E))
        using steps5
        using match_hyps(3) ord_res_5_matches_ord_res_6.intros ord_res_6_preserves_invars step'
          step_hyps(2) by metis
    qed
  next
    case step_hyps: (resolution_neg E L D Uer' ℳ' K C)

    define S5' :: "_ × _ × _ × ('f gterm  'f gclause option) × 'f gclause option" where
      "S5' = (N, Uer', , ℳ', Some C)"

    hence "iefac  |`| (N |∪| Uer')  {||}"
      using Uer' = finsert (eres D E) Uer by simp
    then obtain B where B_least: "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) B"
      by (metis linorder_cls.ex1_least_in_fset)

    show ?thesis
    proof (intro exI conjI)
      have "The_optional (linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer'))) = Some B"
      proof (rule The_optional_eq_SomeI)
        show "1 x. linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) x"
          by blast
      next
        show "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) B"
          using B_least .
      qed

      hence step5: "ord_res_5 N (Uer, , , Some E) (Uer', , Map.empty, Some B)"
        using ord_res_5.resolution step_hyps by metis

      moreover have "(ord_res_5 N)**  (Uer', , ℳ', Some C)"
      proof (rule full_rtranclp_ord_res_5_run_upto)
        show "ord_res_6 N (Uer, , , Some E) (Uer', , ℳ', Some C)"
          using step' 𝒞 = Some E s6' = (Uer', , ℳ', Some C) by argo
      next
        show "ord_res_5_invars N (Uer', , ℳ', Some C)"
          using match_hyps(3) ord_res_6_preserves_invars step' step_hyps(2) by blast
      next
        have "ord_res.is_strictly_maximal_lit (Pos (atm_of K)) C"
          using  (atm_of K) = Some C
            ord_res_5_invars N (Uer, , , 𝒞)[unfolded ord_res_5_invars_def
              atoms_in_model_were_produced_def, simplified]
          using ord_res.mem_productionE by blast

        moreover have "Pos (atm_of K) l K"
          using is_neg K by (cases K) simp_all

        ultimately have "C c eres D E"
          using ord_res.is_maximal_lit K (eres D E)
          by (metis ord_res.asymp_less_lit ord_res.transp_less_lit
              linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset
              linorder_lit.multpHO_if_maximal_less_that_maximal multp_eq_multpHO)

        hence "C c E"
          using eres_le[of D E] by order

        have "C c efac (eres D E)"
          by (metis Uniq_D C c eres D E efac_spec is_pos_def linorder_lit.Uniq_is_maximal_in_mset
              step_hyps(10) step_hyps(11))

        moreover have "efac (eres D E) c eres D E"
          by (metis efac_subset subset_implies_reflclp_multp)

        ultimately have "C c iefac  (eres D E)"
          unfolding iefac_def by auto

        hence "{|Ca |∈| iefac  |`| (N |∪| Uer'). Ca c C|} =
          {|Ca |∈| iefac  |`| (N |∪| Uer). Ca c C|}"
          unfolding Uer' = finsert (eres D E) Uer by auto

        have "(K. ord_res.is_maximal_lit K C  A t atm_of K)  A t atm_of K" for A
          using ord_res.is_strictly_maximal_lit (Pos (atm_of K)) C
          by (metis Uniq_def linorder_lit.Uniq_is_maximal_in_mset
              linorder_lit.is_maximal_in_mset_if_is_greatest_in_mset literal.sel(1))

        hence "{A. K. ord_res.is_maximal_lit K C  A t atm_of K} = {A. A t atm_of K}"
          by metis

        thus "ℳ' = restrict_map  {A. K. ord_res.is_maximal_lit K C  A t atm_of K}"
          using ℳ' = restrict_map  {A. A t atm_of K} by argo
      next
        show "linorder_cls.is_least_in_fset (iefac  |`| (N |∪| Uer')) B"
          using B_least .
      qed

      ultimately have steps5: "(ord_res_5 N)++ (Uer, , , Some E) (Uer', , ℳ', Some C)"
        by simp

      thus "ord_res_5_step++ S5 S5'"
        unfolding S5_def 𝒞 = Some E S5'_def
        by (metis tranclp_ord_res_5_step_if_tranclp_ord_res_5)

      show "ord_res_5_matches_ord_res_6 S5' S6'"
        unfolding S5'_def S6'_def s6' = (Uer', , ℳ', Some C)
        using steps5
        using match_hyps(3) ord_res_5_matches_ord_res_6.intros ord_res_6_preserves_invars step'
          step_hyps(2) by metis
    qed
  qed
qed

theorem bisimulation_ord_res_5_ord_res_6:
  defines "match  λ_. ord_res_5_matches_ord_res_6"
  shows "(MATCH :: nat × nat  'f ord_res_5_state  'f ord_res_6_state  bool) .
    bisimulation ord_res_5_step ord_res_6_step ord_res_5_final ord_res_6_final  MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_5_state ⇒ 'f ord_res_6_state ⇒ bool) ℛfb.
    bisimulation ord_res_5_step ord_res_5_final ord_res_6_step ord_res_6_final MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique ord_res_5_step"
    using right_unique_ord_res_5_step .
next
  show "right_unique ord_res_6_step"
    using right_unique_ord_res_6_step .
next
  show "s. ord_res_5_final s  (s'. ord_res_5_step s s')"
    by (metis finished_def ord_res_5_semantics.final_finished)
next
  show "s. ord_res_6_final s  (s'. ord_res_6_step s s')"
    by (metis finished_def ord_res_6_semantics.final_finished)
next
  show "i S5 S6. match i S5 S6  ord_res_5_final S5  ord_res_6_final S6"
    unfolding match_def
    using ord_res_5_final_iff_ord_res_6_final by metis
next
  show "i S5 S6.
       match i S5 S6 
       safe_state ord_res_5_step ord_res_5_final S5  safe_state ord_res_6_step ord_res_6_final S6"
  proof (intro allI impI conjI)
    fix i S5 S6
    assume "match i S5 S6"
    show "safe_state ord_res_5_step ord_res_5_final S5"
      using match i S5 S6
      using ord_res_5_safe_state_if_invars
      using match_def ord_res_5_matches_ord_res_6.cases by metis
    show "safe_state ord_res_6_step ord_res_6_final S6"
      using match i S5 S6
      using ord_res_6_safe_state_if_invars
      using match_def ord_res_5_matches_ord_res_6.cases by metis
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i S5 S6 S6'.
       match i S5 S6 
       ord_res_6_step S6 S6' 
       (i' S5'. ord_res_5_step++ S5 S5'  match i' S5' S6')  (i'. match i' S5 S6'  False)"
    unfolding match_def
    using backward_simulation_between_5_and_6 by metis
qed

end


section ‹ORD-RES-7 (clause-guided literal trail construction)›

type_synonym 'f ord_res_7_state =
  "'f gclause fset ×'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list ×
    'f gclause option"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_6_matches_ord_res_7 ::
  "'f gterm fset  'f ord_res_6_state  'f ord_res_7_state  bool" where
  "ord_res_5_invars N (Uer, , , 𝒞) 
    ord_res_7_invars N (Uer, , Γ, 𝒞) 
    (A C.  A = Some C  map_of Γ (Pos A) = Some (Some C)) 
    (A.  A = None  map_of Γ (Neg A)  None  A |∉| trail_atms Γ) 
    i = atms_of_clss (N |∪| Uer) - trail_atms Γ 
    ord_res_6_matches_ord_res_7 i (N, Uer, , , 𝒞) (N, Uer, , Γ, 𝒞)"

lemma ord_res_6_final_iff_ord_res_7_final:
  fixes i S6 S7
  assumes match: "ord_res_6_matches_ord_res_7 i S6 S7"
  shows "ord_res_6_final S6  ord_res_7_final S7"
  using match
proof (cases i S6 S7 rule: ord_res_6_matches_ord_res_7.cases)
  case match_hyps: (1 N Uer   𝒞 Γ)

  show "ord_res_6_final S6  ord_res_7_final S7"
  proof (rule iffI)
    assume "ord_res_6_final S6"
    thus "ord_res_7_final S7"
      unfolding S6 = (N, Uer, , , 𝒞)
    proof (cases "(N, Uer, , , 𝒞)" rule: ord_res_6_final.cases)
      case model_found
      thus "ord_res_7_final S7"
        unfolding S7 = (N, Uer, , Γ, 𝒞)
        using ord_res_7_final.model_found
        by metis
    next
      case contradiction_found
      thus "ord_res_7_final S7"
        unfolding S7 = (N, Uer, , Γ, 𝒞)
        using ord_res_7_final.contradiction_found
        by metis
    qed
  next
    assume "ord_res_7_final S7"
    thus "ord_res_6_final S6"
      unfolding S7 = (N, Uer, , Γ, 𝒞)
    proof (cases "(N, Uer, , Γ, 𝒞)" rule: ord_res_7_final.cases)
      case model_found
      thus "ord_res_6_final S6"
        unfolding S6 = (N, Uer, , , 𝒞)
        using ord_res_6_final.model_found
        by metis
    next
      case contradiction_found
      thus "ord_res_6_final S6"
        unfolding S6 = (N, Uer, , , 𝒞)
        using ord_res_6_final.contradiction_found
        by metis
    qed
  qed
qed

lemma backward_simulation_between_6_and_7:
  fixes i S6 S7 S7'
  assumes match: "ord_res_6_matches_ord_res_7 i S6 S7" and step: "constant_context ord_res_7 S7 S7'"
  shows "
    (i' S6'. ord_res_6_step++ S6 S6'  ord_res_6_matches_ord_res_7 i' S6' S7') 
    (i'. ord_res_6_matches_ord_res_7 i' S6 S7'  i' |⊂| i)"
  using match
proof (cases i S6 S7 rule: ord_res_6_matches_ord_res_7.cases)
  case match_hyps: (1 N Uer   𝒞 Γ)

  note S6_def = S6 = (N, Uer, , , 𝒞)
  note invars_6 = ord_res_5_invars N (Uer, , , 𝒞)
  note invars_7 = ord_res_7_invars N (Uer, , Γ, 𝒞)[
      unfolded ord_res_7_invars_def, rule_format, OF refl] 

  have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) Γ"
    using invars_7 by argo

  have Γ_consistent: "trail_consistent Γ"
    using invars_7 by (metis trail_consistent_if_sorted_wrt_atoms)

  hence Γ_distinct_atoms: "distinct (map fst Γ)"
    using distinct_lits_if_trail_consistent by iprover

  have clause_true_wrt_model_if_true_wrt_Γ: "dom   D"
    if D_true: "trail_true_cls Γ D" for D
  proof -
    obtain L where "L ∈# D" and L_true: "trail_true_lit Γ L"
      using D_true unfolding trail_true_cls_def by auto

    have "𝒞. (L, 𝒞)  set Γ"
      using L_true unfolding trail_true_lit_def by auto

    show ?thesis
    proof (cases L)
      case (Pos A)

      then obtain C where "(Pos A, Some C)  set Γ"
        using invars_7 𝒞. (L, 𝒞)  set Γ
        by (metis fst_conv literal.disc(1) not_None_eq snd_conv)

      hence "map_of Γ (Pos A) = Some (Some C)"
        using Γ_distinct_atoms by (metis map_of_is_SomeI)

      hence " A = Some C"
        using A C. ( A = Some C) = (map_of Γ (Pos A) = Some (Some C)) by metis

      hence "A  dom "
        by blast

      then show ?thesis
        using L ∈# D L = Pos A by blast
    next
      case (Neg A)

      hence "(Neg A, None)  set Γ"
        using invars_7 𝒞. (L, 𝒞)  set Γ
        by (metis fst_conv literal.disc(2) snd_conv)

      hence "map_of Γ (Neg A)  None"
        by (simp add: weak_map_of_SomeI)

      hence " A = None"
        using A. ( A = None) = (map_of Γ (Neg A)  None  A |∉| trail_atms Γ) by metis

      hence "A  dom "
        by blast

      then show ?thesis
        using L ∈# D L = Neg A by blast
    qed
  qed

  have clause_false_wrt_model_if_false_wrt_Γ: "¬ dom   D"
    if D_false: "trail_false_cls Γ D" for D
    unfolding true_cls_def
  proof (intro notI , elim bexE)
    fix L :: "'f gterm literal"
    assume "L ∈# D" and "dom  ⊫l L"

    have "trail_false_lit Γ L"
      using L ∈# D D_false unfolding trail_false_cls_def by metis

    hence "¬ trail_true_lit Γ L" and "trail_defined_lit Γ L"
      unfolding atomize_conj
      using Γ_consistent L ∈# D not_trail_true_cls_and_trail_false_cls that
        trail_defined_lit_iff_true_or_false trail_true_cls_def by blast

    show False
    proof (cases L)
      case (Pos A)

      hence " A  None"
        using dom  ⊫l L by blast

      hence "map_of Γ (Pos A)  None"
        using A C. ( A = Some C) = (map_of Γ (Pos A) = Some (Some C)) by blast

      hence "Pos A  fst ` set Γ"
        by (simp add: map_of_eq_None_iff)

      hence "trail_true_lit Γ (Pos A)"
        unfolding trail_true_lit_def .

      moreover have "¬ trail_true_lit Γ (Pos A)"
        using ¬ trail_true_lit Γ L L = Pos A by argo

      ultimately show False
        by contradiction
    next
      case (Neg A)

      hence " A = None"
        using dom  ⊫l L by blast

      hence "map_of Γ (Neg A)  None  A |∉| trail_atms Γ"
        using A. ( A = None) = (map_of Γ (Neg A)  None  A |∉| trail_atms Γ) by blast

      hence "trail_true_lit Γ (Neg A)  ¬ trail_defined_lit Γ (Neg A)"
        unfolding map_of_eq_None_iff not_not
        unfolding trail_true_lit_def trail_defined_lit_iff_trail_defined_atm literal.sel
        .

      then show ?thesis
        using ¬ trail_true_lit Γ L trail_defined_lit Γ L L = Neg A by argo
    qed
  qed

  obtain s7' where
    "S7' = (N, s7')" and
    step': "ord_res_7 N (Uer, , Γ, 𝒞) s7'"
    using step unfolding S7 = (N, Uer, , Γ, 𝒞)
    by (auto elim: constant_context.cases)

  have invars_s7': "ord_res_7_invars N s7'"
    using ord_res_7_preserves_invars[OF step' ord_res_7_invars N (Uer, , Γ, 𝒞)] .

  show ?thesis
    using step'
  proof (cases N "(Uer, , Γ, 𝒞)" s7' rule: ord_res_7.cases)
    case step_hyps: (decide_neg C L A Γ')

    define i' where
      "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"

    have "A |∈| atms_of_clss (N |∪| Uer)" and "A t atm_of L" and "A |∉| trail_atms Γ"
      using step_hyps unfolding atomize_conj linorder_trm.is_least_in_ffilter_iff by argo

    have "ord_res_6_matches_ord_res_7 i' S6 S7'"
      unfolding S6_def 𝒞 = Some C S7' = (N, s7') s7' = (Uer, , Γ', Some C)
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , , Some C)"
        using invars_6 unfolding 𝒞 = Some C .
    next
      show "ord_res_7_invars N (Uer, , Γ', Some C)"
        using invars_s7' unfolding s7' = (Uer, , Γ', Some C) .
    next
      show "A C. ( A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
        using match_hyps unfolding Γ' = (Neg A, None) # Γ by simp
    next
      show "A. ( A = None) = (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        unfolding Γ' = (Neg A, None) # Γ
        using match_hyps A |∉| trail_atms Γ by force
    next
      show "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"
        unfolding i'_def ..
    qed

    moreover have "i' |⊂| i"
    proof -
      have "i = finsert A i'"
        unfolding match_hyps i'_def
        using A |∈| atms_of_clss (N |∪| Uer) A |∉| trail_atms Γ step_hyps(6) by force

      moreover have "A |∉| i'"
        unfolding i'_def
        using step_hyps(6) by fastforce

      ultimately show ?thesis
        by auto
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (skip_defined C L 𝒞')

    define S6' where
      "S6' = (N, Uer, , , 𝒞')"

    have C_almost_defined: "trail_defined_cls Γ {#x ∈# C. x  L#}"
      using step_hyps by (metis clause_almost_definedI invars_7)

    hence C_defined: "trail_defined_cls Γ C"
      using step_hyps unfolding trail_defined_cls_def by auto

    hence C_true: "trail_true_cls Γ C"
      using step_hyps by (metis trail_true_or_false_cls_if_defined)

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, , , 𝒞')"
    proof (rule ord_res_6.skip)
      show "dom   C"
        using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
    next
      show "𝒞' = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        using step_hyps by argo
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, , Γ, 𝒞')
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , , 𝒞')"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, , Γ, 𝒞')"
        using invars_s7' unfolding s7' = (Uer, , Γ, 𝒞') .
    next
      show "A C. ( A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
        using match_hyps by argo
    next
      show "A. ( A = None) = (map_of Γ (Neg A)  None  A |∉| trail_atms Γ)"
        using match_hyps by argo
    next
      show "i = atms_of_clss (N |∪| Uer) |-| trail_atms Γ"
        using match_hyps by argo
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (skip_undefined_neg C L Γ' 𝒞')

    define S6' where
      "S6' = (N, Uer, , , 𝒞')"

    define i' where
      "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"

    have "trail_true_lit Γ' L"
      unfolding Γ' = (L, None) # Γ by (simp add: trail_true_lit_def)

    hence C_true: "trail_true_cls Γ' C"
      using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff trail_true_cls_def by metis

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, , , 𝒞')"
    proof (rule ord_res_6.skip)
      show "dom   C"
        using C_true
        by (metis domIff linorder_lit.is_maximal_in_mset_iff literal.collapse(2) match_hyps(6)
            step_hyps(4) step_hyps(6) step_hyps(7) trail_defined_lit_iff_trail_defined_atm
            true_cls_def true_lit_simps(2))
    next
      show "𝒞' = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        using step_hyps by argo
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, , Γ', 𝒞')
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , , 𝒞')"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, , Γ', 𝒞')"
        using invars_s7' unfolding s7' = (Uer, , Γ', 𝒞') .
    next
      show "A C. ( A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
        using match_hyps
        unfolding Γ' = (L, None) # Γ
        by (metis literal.disc(1) map_of_Cons_code(2) step_hyps(7))
    next
      show "A. ( A = None) = (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        using match_hyps
        unfolding Γ' = (L, None) # Γ
        by (metis finsert_iff literal.collapse(2) literal.sel(2) map_of_Cons_code(2) option.discI
            prod.sel(1) step_hyps(6) step_hyps(7) trail_atms.simps(2)
            trail_defined_lit_iff_trail_defined_atm)
    next
      show "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (skip_undefined_pos C L D)

    define S6' where
      "S6' = (N, Uer, , , Some D)"

    have "trail_defined_cls Γ {#x ∈# C. x  L  x  - L#}"
    proof (rule clause_almost_almost_definedI)
      show "C |∈| iefac  |`| (N |∪| Uer)"
        using invars_7 step_hyps by metis
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by argo
    next
      show "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)"
        using step_hyps by argo
    qed

    moreover have "- L ∉# C"
      by (metis atm_of_uminus is_pos_def linorder_lit.is_maximal_in_mset_iff linorder_lit.neqE
          linorder_trm.less_irrefl literal.collapse(2) literal.sel(1) ord_res.less_lit_simps(4)
          step_hyps(4) step_hyps(7) uminus_not_id')

    ultimately have "trail_defined_cls Γ {#x ∈# C. x  L#}"
      unfolding trail_defined_cls_def by auto

    hence "trail_true_cls Γ {#x ∈# C. x  L#}"
      using ¬ trail_false_cls Γ {#K ∈# C. K  L#} by (metis trail_true_or_false_cls_if_defined)

    hence C_true: "trail_true_cls Γ C"
      by (auto simp: trail_true_cls_def)

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, , , Some D)"
    proof (rule ord_res_6.skip)
      show "dom   C"
        using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
    next
      show "Some D = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        using linorder_cls.Uniq_is_least_in_fset step_hyps(9) The_optional_eq_SomeI
        by fastforce
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, , Γ, Some D)
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , , Some D)"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, , Γ, Some D)"
        using invars_s7' unfolding s7' = (Uer, , Γ, Some D) .
    next
      show "A C. ( A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
        using match_hyps by argo
    next
      show "A. ( A = None) = (map_of Γ (Neg A)  None  A |∉| trail_atms Γ)"
        using match_hyps by argo
    next
      show "i = atms_of_clss (N |∪| Uer) |-| trail_atms Γ"
        using match_hyps by argo
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (skip_undefined_pos_ultimate C L Γ')
    
    define S6' where
      "S6' = (N, Uer, , , None :: 'f gclause option)"

    define i' where
      "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"

    have "trail_defined_cls Γ {#x ∈# C. x  L  x  - L#}"
    proof (rule clause_almost_almost_definedI)
      show "C |∈| iefac  |`| (N |∪| Uer)"
        using invars_7 step_hyps by metis
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by argo
    next
      show "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)"
        using step_hyps by argo
    qed

    moreover have "- L ∉# C"
      by (metis atm_of_uminus is_pos_def linorder_lit.is_maximal_in_mset_iff linorder_lit.neqE
          linorder_trm.less_irrefl literal.collapse(2) literal.sel(1) ord_res.less_lit_simps(4)
          step_hyps(4) step_hyps(7) uminus_not_id')

    ultimately have "trail_defined_cls Γ {#x ∈# C. x  L#}"
      unfolding trail_defined_cls_def by auto

    hence "trail_true_cls Γ {#x ∈# C. x  L#}"
      using ¬ trail_false_cls Γ {#K ∈# C. K  L#} by (metis trail_true_or_false_cls_if_defined)

    hence C_true: "trail_true_cls Γ C"
      by (auto simp: trail_true_cls_def)

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, , , None)"
    proof (rule ord_res_6.skip)
      show "dom   C"
        using clause_true_wrt_model_if_true_wrt_Γ[OF C_true] .
    next
      have "¬ (D. linorder_cls.is_least_in_fset (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))) D)"
        using ¬ fBex (iefac  |`| (N |∪| Uer)) ((≺c) C)
        by (meson linorder_cls.is_least_in_ffilter_iff)

      thus "None = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        unfolding The_optional_def by metis
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, , Γ', None)
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , , None)"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, , Γ', None)"
        using invars_s7' unfolding s7' = (Uer, , Γ', None) .
    next
      show "A C. ( A = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
        using match_hyps(3-)
        unfolding Γ' = (- L, None) # Γ
        by (metis is_pos_neg_not_is_pos literal.disc(1) map_of_Cons_code(2) step_hyps(7))
    next
      show "A. ( A = None) = (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        using match_hyps(3-)
        unfolding Γ' = (- L, None) # Γ
        by (metis (no_types, opaque_lifting) atm_of_eq_atm_of eq_fst_iff fset_simps(2) insertCI
            insertE literal.discI(2) literal.sel(2) map_of_Cons_code(2) option.distinct(1)
            trail_defined_lit_iff_trail_defined_atm step_hyps(6) step_hyps(7) trail_atms.simps(2))
    next
      show "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (production C L Γ' 𝒞')
    
    define S6' where
      "S6' = (N, Uer, , (atm_of L  C), 𝒞')"

    define i' where
      "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"

    have "L ∈# C"
      using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff by argo

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, , (atm_of L  C), 𝒞')"
    proof (rule ord_res_6.production)
      have "atm_of L |∉| trail_atms Γ"
        using ¬ trail_defined_lit Γ L
        unfolding trail_defined_lit_iff_trail_defined_atm .

      hence " (atm_of L) = None"
        using match_hyps(3-) by metis

      hence "atm_of L  dom "
        unfolding dom_def by simp

      hence "¬ dom  ⊫l L"
        using is_pos L unfolding true_lit_def by metis

      moreover have "¬ dom   {#K ∈# C. K  L#}"
        using clause_false_wrt_model_if_false_wrt_Γ[OF trail_false_cls Γ {#K ∈# C. K  L#}] .

      ultimately show "¬ dom   C"
        using L ∈# C
        unfolding true_cls_def by auto
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by argo
    next
      show "is_pos L"
        using step_hyps by argo
    next
      show "ord_res.is_strictly_maximal_lit L C"
        using step_hyps by argo
    next
      show "(atm_of L  C) = (atm_of L  C)" ..
    next
      show "𝒞' = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        using step_hyps by argo
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, , Γ', 𝒞')
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, , (atm_of L  C), 𝒞')"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, , Γ', 𝒞')"
        using invars_s7' unfolding s7' = (Uer, , Γ', 𝒞') .
    next
      show "A D. (((atm_of L  C)) A = Some D) = (map_of Γ' (Pos A) = Some (Some D))"
        using match_hyps(3-)
        unfolding Γ' = (L, Some C) # Γ
        using step_hyps(7) by auto
    next
      show "A. (((atm_of L  C)) A = None) = (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        using match_hyps(3-)
        unfolding Γ' = (L, Some C) # Γ
        by (metis (no_types, opaque_lifting) domI domIff finsert_iff fun_upd_apply
            literal.collapse(1) literal.discI(2) map_of_Cons_code(2) map_of_eq_None_iff prod.sel(1)
            step_hyps(6) step_hyps(7) trail_atms.simps(2) trail_defined_lit_def uminus_Pos)
    next
      show "i' = atms_of_clss (N |∪| Uer) |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (factoring C L ℱ')
    
    define S6' where
      "S6' = (N, Uer, ℱ', , Some (efac C))"

    have "L ∈# C"
      using step_hyps unfolding linorder_lit.is_maximal_in_mset_iff by argo

    have step6: "ord_res_6 N (Uer, , , Some C) (Uer, ℱ', , Some (efac C))"
    proof (rule ord_res_6.factoring)
      have "atm_of L |∉| trail_atms Γ"
        using ¬ trail_defined_lit Γ L
        unfolding trail_defined_lit_iff_trail_defined_atm .

      hence " (atm_of L) = None"
        using match_hyps(3-) by metis

      hence "atm_of L  dom "
        unfolding dom_def by simp

      hence "¬ dom  ⊫l L"
        using is_pos L unfolding true_lit_def by metis

      moreover have "¬ dom   {#K ∈# C. K  L#}"
        using clause_false_wrt_model_if_false_wrt_Γ[OF trail_false_cls Γ {#K ∈# C. K  L#}] .

      ultimately show "¬ dom   C"
        using L ∈# C
        unfolding true_cls_def by auto
    next
      show "ord_res.is_maximal_lit L C"
        using step_hyps by argo
    next
      show "is_pos L"
        using step_hyps by argo
    next
      show "¬ ord_res.is_strictly_maximal_lit L C"
        using step_hyps by argo
    next
      show "ℱ' = finsert C "
        using step_hyps by argo
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some C S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer, ℱ', Γ, Some (efac C))
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer, ℱ', , Some (efac C))"
        using invars_6 unfolding 𝒞 = Some C
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer, ℱ', Γ, Some (efac C))"
        using invars_s7' unfolding s7' = (Uer, ℱ', Γ, Some (efac C)) .
    next
      show "A C. ( A = Some C) = (map_of Γ (Pos A) = Some (Some C))"
        using match_hyps(3-) by argo
    next
      show "A. ( A = None) = (map_of Γ (Neg A)  None  A |∉| trail_atms Γ)"
        using match_hyps(3-) by argo
    next
      show "i = atms_of_clss (N |∪| Uer) |-| trail_atms Γ"
        using match_hyps(3-) by argo
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (resolution_bot E L D Uer' Γ')
   
    define S6' where
      "S6' = (N, Uer', , (λ_. None) :: 'f gterm  'f gclause option, Some ({#} :: 'f gclause))"

    define i' where
      "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"

    have step6: "ord_res_6 N (Uer, , , Some E) (Uer', , λ_. None, Some {#})"
    proof (rule ord_res_6.resolution_bot)
      show "¬ dom   E"
        using clause_false_wrt_model_if_false_wrt_Γ[OF trail_false_cls Γ E] .
    next
      show "ord_res.is_maximal_lit L E"
        using step_hyps by argo
    next
      show "is_neg L"
        using step_hyps by argo
    next
      show " (atm_of L) = Some D"
        by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
    next
      show "Uer' = finsert (eres D E) Uer"
        using step_hyps by argo
    next
      show "eres D E = {#}"
        using step_hyps by argo
    next
      show "((λ_. None)) = (λ_. None)" ..
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some E S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer', , Γ', Some {#})
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer', , λ_. None, Some {#})"
        using invars_6 unfolding 𝒞 = Some E
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer', , Γ', Some {#})"
        using invars_s7' unfolding s7' = (Uer', , Γ', Some {#}) .
    next
      show "A C. (None = Some C) = (map_of Γ' (Pos A) = Some (Some C))"
        unfolding Γ' = [] by simp
    next
      show "A. (None = None) = (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        unfolding Γ' = [] by simp
    next
      show "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (resolution_pos E L D Uer' Γ' K)

    define S6' where
      "S6' = (N, Uer', , restrict_map  {A. A t atm_of K}, Some (eres D E))"

    define i' where
      "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"

    have mem_set_Γ'_iff: "x. (x  set Γ') = (atm_of (fst x) t atm_of K  x  set Γ)"
      unfolding Γ' = dropWhile (λLn. atm_of K t atm_of (fst Ln)) Γ
      unfolding mem_set_dropWhile_conv_if_list_sorted_and_pred_monotone[OF Γ_sorted mono_atms_lt]
      by auto

    have step6: "ord_res_6 N (Uer, , , Some E) (Uer', , restrict_map  {A. A t atm_of K}, Some (eres D E))"
    proof (rule ord_res_6.resolution_pos)
      show "¬ dom   E"
        using clause_false_wrt_model_if_false_wrt_Γ[OF trail_false_cls Γ E] .
    next
      show "ord_res.is_maximal_lit L E"
        using step_hyps by argo
    next
      show "is_neg L"
        using step_hyps by argo
    next
      show " (atm_of L) = Some D"
        by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
    next
      show "Uer' = finsert (eres D E) Uer"
        using step_hyps by argo
    next
      show "eres D E  {#}"
        using step_hyps by argo
    next
      show "restrict_map  {A. A t atm_of K} = restrict_map  {A. A t atm_of K}"
        ..
    next
      show "ord_res.is_maximal_lit K (eres D E)"
        using step_hyps by argo
    next
      show "is_pos K"
        using step_hyps by argo
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some E S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer', , Γ', Some (eres D E))
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer', , restrict_map  {A. A t atm_of K}, Some (eres D E))"
        using invars_6 unfolding 𝒞 = Some E
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer', , Γ', Some (eres D E))"
        using invars_s7' unfolding s7' = (Uer', , Γ', Some (eres D E)) .
    next
      show "A C. (restrict_map  {A. A t atm_of K} A = Some C) =
        (map_of Γ' (Pos A) = Some (Some C))"
      proof (intro allI)
        fix A :: "'f gterm" and C :: "'f gclause"
        show "restrict_map  {A. A t atm_of K} A = Some C  map_of Γ' (Pos A) = Some (Some C)"
        proof (cases "A  dom   A t atm_of K")
          case True
          have "restrict_map  {A. A t atm_of K} A = Some C   A = Some C"
            using True by simp

          also have "  map_of Γ (Pos A) = Some (Some C)"
            using match_hyps(3-) by metis

          also have "  map_of Γ' (Pos A) = Some (Some C)"
          proof -
            have "Pos A  fst ` set Γ"
              using True 
              by (metis domIff map_of_eq_None_iff match_hyps(5) not_None_eq)

            hence "𝒞. (Pos A, 𝒞)  set Γ"
              by fastforce

            hence "𝒞. (Pos A, 𝒞)  set Γ  (Pos A, 𝒞)  set Γ'"
              using True unfolding mem_set_Γ'_iff prod.sel literal.sel by metis

            moreover have "distinct (map fst Γ')"
              using Γ_distinct_atoms
            proof (rule distinct_suffix)
              show "suffix (map fst Γ') (map fst Γ)"
                using map_mono_suffix step_hyps(9) suffix_dropWhile by blast
            qed

            ultimately have "map_of Γ (Pos A) = map_of Γ' (Pos A)"
              using Γ_distinct_atoms by (auto dest: map_of_is_SomeI)

            thus ?thesis
              by argo
          qed

          finally show ?thesis .
        next
          case False
          have "restrict_map  {A. A t atm_of K} A = None"
            using False unfolding restrict_map_def by auto

          moreover have "map_of Γ' (Pos A)  Some (Some C)"
            using False unfolding de_Morgan_conj
          proof (elim disjE)
            assume "A  dom "

            hence "𝒞. (Pos A, 𝒞)  set Γ"
              using match_hyps(5)
              by (metis (no_types, opaque_lifting) domIff fst_eqD invars_7 is_pos_def map_of_SomeD
                  not_None_eq snd_conv weak_map_of_SomeI)

            hence "𝒞. (Pos A, 𝒞)  set Γ'"
              unfolding mem_set_Γ'_iff by simp

            then show "map_of Γ' (Pos A)  Some (Some C)"
              by (meson map_of_SomeD)
          next
            assume "¬ A t atm_of K"

            hence "𝒞. (Pos A, 𝒞)  set Γ'"
              unfolding mem_set_Γ'_iff by simp

            then show "map_of Γ' (Pos A)  Some (Some C)"
              by (meson map_of_SomeD)
          qed

          ultimately show ?thesis
            by simp
        qed
      qed
    next
      show "A. (restrict_map  {A. A t atm_of K} A = None) =
        (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
      proof (intro allI)
        fix A :: "'f gterm"
        show "(restrict_map  {A. A t atm_of K} A = None) =
          (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        proof (cases "A t atm_of K")
          case True

          have "restrict_map  {A. A t atm_of K} A = None   A = None"
            using True by simp

          also have "  map_of Γ (Neg A)  None  A |∉| trail_atms Γ"
            using match_hyps(6) by metis

          also have "  map_of Γ' (Neg A)  None  A |∉| trail_atms Γ"
            using True mem_set_Γ'_iff
            by (metis eq_fst_iff literal.sel(2) map_of_SomeD not_None_eq weak_map_of_SomeI)

          also have "  map_of Γ' (Neg A)  None  A |∉| trail_atms Γ'"
            using True mem_set_Γ'_iff
            by (smt (verit, best) fset_trail_atms image_iff)

          finally show ?thesis .
        next
          case False

          have "restrict_map  {A. A t atm_of K} A = None"
            using False by simp

          moreover have "A |∉| trail_atms Γ'"
            using False mem_set_Γ'_iff
            by (smt (verit, ccfv_threshold) fset_trail_atms image_iff)

          ultimately show ?thesis
            by metis
        qed
      qed
    next
      show "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  next
    case step_hyps: (resolution_neg E L D Uer' Γ' K C)
    
    define S6' where
      "S6' = (N, Uer', , restrict_map  {A. A t atm_of K}, Some C)"

    define i' where
      "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"

    have mem_set_Γ'_iff: "x. (x  set Γ') = (atm_of (fst x) t atm_of K  x  set Γ)"
      unfolding Γ' = dropWhile (λLn. atm_of K t atm_of (fst Ln)) Γ
      unfolding mem_set_dropWhile_conv_if_list_sorted_and_pred_monotone[OF Γ_sorted mono_atms_lt]
      by auto

    have step6: "ord_res_6 N (Uer, , , Some E) (Uer', , restrict_map  {A. A t atm_of K}, Some C)"
    proof (rule ord_res_6.resolution_neg)
      show "¬ dom   E"
        using clause_false_wrt_model_if_false_wrt_Γ[OF trail_false_cls Γ E] .
    next
      show "ord_res.is_maximal_lit L E"
        using step_hyps by argo
    next
      show "is_neg L"
        using step_hyps by argo
    next
      show " (atm_of L) = Some D"
        by (metis literal.collapse(2) match_hyps(5) step_hyps(5) step_hyps(6) uminus_Neg)
    next
      show "Uer' = finsert (eres D E) Uer"
        using step_hyps by argo
    next
      show "eres D E  {#}"
        using step_hyps by argo
    next
      show "restrict_map  {A. A t atm_of K} = restrict_map  {A. A t atm_of K}"
        ..
    next
      show "ord_res.is_maximal_lit K (eres D E)"
        using step_hyps by argo
    next
      show "is_neg K"
        using step_hyps by argo
    next
      show " (atm_of K) = Some C"
        using match_hyps(3-)
        by (metis (mono_tags, lifting) step_hyps(11) step_hyps(12) uminus_literal_def)
    qed

    hence "ord_res_6_step++ S6 S6'"
      using S6_def 𝒞 = Some E S6'_def ord_res_6_step.intros by blast

    moreover have "ord_res_6_matches_ord_res_7 i' S6' S7'"
      unfolding S6'_def S7' = (N, s7') s7' = (Uer', , Γ', Some C)
    proof (rule ord_res_6_matches_ord_res_7.intros)
      show "ord_res_5_invars N (Uer', , restrict_map  {A. A t atm_of K}, Some C)"
        using invars_6 unfolding 𝒞 = Some E
        using ord_res_6_preserves_invars[OF step6] by argo
    next
      show "ord_res_7_invars N (Uer', , Γ', Some C)"
        using invars_s7' unfolding s7' = (Uer', , Γ', Some C) .
    next
      show "A C. (restrict_map  {A. A t atm_of K} A = Some C) =
        (map_of Γ' (Pos A) = Some (Some C))"
      proof (intro allI)
        fix A :: "'f gterm" and C :: "'f gclause"
        show "restrict_map  {A. A t atm_of K} A = Some C  map_of Γ' (Pos A) = Some (Some C)"
        proof (cases "A  dom   A t atm_of K")
          case True
          have "restrict_map  {A. A t atm_of K} A = Some C   A = Some C"
            using True by simp

          also have "  map_of Γ (Pos A) = Some (Some C)"
            using match_hyps(3-) by metis

          also have "  map_of Γ' (Pos A) = Some (Some C)"
          proof -
            have "Pos A  fst ` set Γ"
              using True 
              by (metis domIff map_of_eq_None_iff match_hyps(5) not_None_eq)

            hence "𝒞. (Pos A, 𝒞)  set Γ"
              by fastforce

            hence "𝒞. (Pos A, 𝒞)  set Γ  (Pos A, 𝒞)  set Γ'"
              using True unfolding mem_set_Γ'_iff prod.sel literal.sel by metis

            moreover have "distinct (map fst Γ')"
              using Γ_distinct_atoms
            proof (rule distinct_suffix)
              show "suffix (map fst Γ') (map fst Γ)"
                using map_mono_suffix step_hyps(9) suffix_dropWhile by blast
            qed

            ultimately have "map_of Γ (Pos A) = map_of Γ' (Pos A)"
              using Γ_distinct_atoms by (auto dest: map_of_is_SomeI)

            thus ?thesis
              by argo
          qed

          finally show ?thesis .
        next
          case False
          have "restrict_map  {A. A t atm_of K} A = None"
            using False unfolding restrict_map_def by auto

          moreover have "map_of Γ' (Pos A)  Some (Some C)"
            using False unfolding de_Morgan_conj
          proof (elim disjE)
            assume "A  dom "

            hence "𝒞. (Pos A, 𝒞)  set Γ"
              using match_hyps(5)
              by (metis (no_types, opaque_lifting) domIff fst_eqD invars_7 is_pos_def map_of_SomeD
                  not_None_eq snd_conv weak_map_of_SomeI)

            hence "𝒞. (Pos A, 𝒞)  set Γ'"
              unfolding mem_set_Γ'_iff by simp

            then show "map_of Γ' (Pos A)  Some (Some C)"
              by (meson map_of_SomeD)
          next
            assume "¬ A t atm_of K"

            hence "𝒞. (Pos A, 𝒞)  set Γ'"
              unfolding mem_set_Γ'_iff by simp

            then show "map_of Γ' (Pos A)  Some (Some C)"
              by (meson map_of_SomeD)
          qed

          ultimately show ?thesis
            by simp
        qed
      qed
    next
      show "A. (restrict_map  {A. A t atm_of K} A = None) =
        (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
      proof (intro allI)
        fix A :: "'f gterm"
        show "(restrict_map  {A. A t atm_of K} A = None) =
          (map_of Γ' (Neg A)  None  A |∉| trail_atms Γ')"
        proof (cases "A t atm_of K")
          case True

          have "restrict_map  {A. A t atm_of K} A = None   A = None"
            using True by simp

          also have "  map_of Γ (Neg A)  None  A |∉| trail_atms Γ"
            using match_hyps(6) by metis

          also have "  map_of Γ' (Neg A)  None  A |∉| trail_atms Γ"
            using True mem_set_Γ'_iff
            by (metis eq_fst_iff literal.sel(2) map_of_SomeD not_None_eq weak_map_of_SomeI)

          also have "  map_of Γ' (Neg A)  None  A |∉| trail_atms Γ'"
            using True mem_set_Γ'_iff
            by (smt (verit, best) fset_trail_atms image_iff)

          finally show ?thesis .
        next
          case False

          have "restrict_map  {A. A t atm_of K} A = None"
            using False by simp

          moreover have "A |∉| trail_atms Γ'"
            using False mem_set_Γ'_iff
            by (smt (verit, ccfv_threshold) fset_trail_atms image_iff)

          ultimately show ?thesis
            by metis
        qed
      qed
    next
      show "i' = atms_of_clss (N |∪| Uer') |-| trail_atms Γ'"
        using i'_def .
    qed

    ultimately show ?thesis
      by metis
  qed
qed

theorem bisimulation_ord_res_6_ord_res_7:
  defines "match  ord_res_6_matches_ord_res_7"
  shows "(MATCH :: nat × nat  'f ord_res_6_state  'f ord_res_7_state  bool) .
    bisimulation ord_res_6_step (constant_context ord_res_7) ord_res_6_final ord_res_7_final
       MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_6_state ⇒ 'f ord_res_7_state ⇒ bool) ℛfb.
    bisimulation ord_res_6_step ord_res_6_final (constant_context ord_res_7) ord_res_7_final
      MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique ord_res_6_step"
    using right_unique_ord_res_6_step .
next
  show "right_unique (constant_context ord_res_7)"
    using right_unique_constant_context right_unique_ord_res_7 by metis
next
  show "S. ord_res_6_final S  (S'. ord_res_6_step S S')"
    by (metis finished_def ord_res_6_semantics.final_finished)
next
  show "S. ord_res_7_final S  (S'. constant_context ord_res_7 S S')"
    by (metis finished_def ord_res_7_semantics.final_finished)
next
  show "i S6 S7. match i S6 S7  ord_res_6_final S6  ord_res_7_final S7"
    unfolding match_def
    using ord_res_6_final_iff_ord_res_7_final by metis
next
  show "i S6 S7. match i S6 S7 
       safe_state ord_res_6_step ord_res_6_final S6 
       safe_state (constant_context ord_res_7) ord_res_7_final S7"
  proof (intro allI impI conjI)
    fix i S6 S7
    assume "match i S6 S7"
    show "safe_state ord_res_6_step ord_res_6_final S6"
      using match i S6 S7[unfolded match_def]
      using ord_res_6_safe_state_if_invars
      using ord_res_6_matches_ord_res_7.simps by auto

    show "safe_state (constant_context ord_res_7) ord_res_7_final S7"
      using match i S6 S7[unfolded match_def]
      using ord_res_7_safe_state_if_invars
      using ord_res_6_matches_ord_res_7.simps by auto
  qed
next
  show "wfp (|⊂|)"
    using wfP_pfsubset .
next
  show "i S6 S7 S7'. match i S6 S7  constant_context ord_res_7 S7 S7' 
    (i' S6'. ord_res_6_step++ S6 S6'  match i' S6' S7') 
    (i'. match i' S6 S7'  i' |⊂| i)"
    unfolding match_def
    using backward_simulation_between_6_and_7 by metis
qed

end


section ‹ORD-RES-8 (atom-guided literal trail construction)›

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_8_can_decide_neg where
  "¬ trail_false_cls Γ C 
    linorder_lit.is_maximal_in_mset C L 
    linorder_trm.is_least_in_fset {|A |∈| atms_of_clss (N |∪| Uer).
      A t atm_of L  A |∉| trail_atms Γ|} A 
    ord_res_8_can_decide_neg N Uer  Γ C"

inductive ord_res_8_can_skip_undefined_neg where
  "¬ trail_false_cls Γ C 
    linorder_lit.is_maximal_in_mset C L 
    ¬(A |∈| atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ) 
    ¬ trail_defined_lit Γ L 
    is_neg L 
    ord_res_8_can_skip_undefined_neg N Uer  Γ C"

inductive ord_res_8_can_skip_undefined_pos_ultimate where
  "¬ trail_false_cls Γ C 
    linorder_lit.is_maximal_in_mset C L 
    ¬(A |∈| atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ) 
    ¬ trail_defined_lit Γ L 
    is_pos L 
    ¬ trail_false_cls Γ {#K ∈# C. K  L#} 
    ¬(D |∈| iefac  |`| (N |∪| Uer). C c D) 
    ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C"

inductive ord_res_8_can_produce where
  "¬ trail_false_cls Γ C 
    linorder_lit.is_maximal_in_mset C L 
    ¬ (A |∈| atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ) 
    ¬ trail_defined_lit Γ L 
    is_pos L 
    trail_false_cls Γ {#K ∈# C. K  L#} 
    linorder_lit.is_greatest_in_mset C L 
    ord_res_8_can_produce N Uer  Γ C"

inductive ord_res_8_can_factorize where
  "¬ trail_false_cls Γ C 
    linorder_lit.is_maximal_in_mset C L 
    ¬ (A |∈| atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ) 
    ¬ trail_defined_lit Γ L 
    is_pos L 
    trail_false_cls Γ {#K ∈# C. K  L#} 
    ¬ linorder_lit.is_greatest_in_mset C L 
    ord_res_8_can_factorize N Uer  Γ C"

definition is_least_nonskipped_clause where
  "is_least_nonskipped_clause N Uer  Γ C 
    linorder_cls.is_least_in_fset {|C |∈| iefac  |`| (N |∪| Uer).
      trail_false_cls Γ C 
      ord_res_8_can_decide_neg N Uer  Γ C 
      ord_res_8_can_skip_undefined_neg N Uer  Γ C 
      ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C 
      ord_res_8_can_produce N Uer  Γ C 
      ord_res_8_can_factorize N Uer  Γ C|} C"

lemma is_least_nonskipped_clause_mempty:
  assumes bot_in: "{#} |∈| iefac  |`| (N |∪| Uer)"
  shows "is_least_nonskipped_clause N Uer  Γ {#}"
  unfolding is_least_nonskipped_clause_def linorder_cls.is_least_in_ffilter_iff
proof (intro conjI ballI impI)
  show "{#} |∈| iefac  |`| (N |∪| Uer)"
    using bot_in .
next
  show "trail_false_cls Γ {#} 
    ord_res_8_can_decide_neg N Uer  Γ {#} 
    ord_res_8_can_skip_undefined_neg N Uer  Γ {#} 
    ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ {#} 
    ord_res_8_can_produce N Uer  Γ {#}  ord_res_8_can_factorize N Uer  Γ {#}"
    by simp
next
  fix C :: "'f gclause"
  assume "C  {#}"
  thus "{#} c C"
    using mempty_lesseq_cls by blast
qed

lemma nex_is_least_nonskipped_clause_if:
  assumes
    no_undef_atom: "¬ (A|∈|atms_of_clss (N |∪| Uer). A |∉| trail_atms Γ)" and
    no_false_clause: "¬ fBex (iefac  |`| (N |∪| Uer)) (trail_false_cls Γ)"
  shows "C. is_least_nonskipped_clause N Uer  Γ C"
  unfolding not_ex
proof (intro allI notI)
  fix C :: "'f gclause"
  assume "is_least_nonskipped_clause N Uer  Γ C"

  hence C_in: "C |∈| iefac  |`| (N |∪| Uer)" and
    "ord_res_8_can_decide_neg N Uer  Γ C 
     ord_res_8_can_skip_undefined_neg N Uer  Γ C 
     ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C 
     ord_res_8_can_produce N Uer  Γ C 
     ord_res_8_can_factorize N Uer  Γ C"
    unfolding atomize_conj
    unfolding is_least_nonskipped_clause_def
    unfolding linorder_cls.is_least_in_ffilter_iff
    using no_false_clause by metis

  hence "L. linorder_lit.is_maximal_in_mset C L  ¬ trail_defined_lit Γ L"
  proof (elim disjE)
    assume "ord_res_8_can_decide_neg N Uer  Γ C"
    then show ?thesis
      by (metis (mono_tags, lifting) assms(1) linorder_trm.is_least_in_ffilter_iff
          ord_res_8_can_decide_neg.cases)
  qed (auto elim:
      ord_res_8_can_skip_undefined_neg.cases
      ord_res_8_can_skip_undefined_pos_ultimate.cases
      ord_res_8_can_produce.cases
      ord_res_8_can_factorize.cases)

  hence "L. atm_of L |∈| atms_of_clss (N |∪| Uer)  atm_of L |∉| trail_atms Γ"
    using C_in
    unfolding linorder_lit.is_maximal_in_mset_iff trail_defined_lit_iff_trail_defined_atm
    by (metis atm_of_in_atms_of_clssI)

  thus False
    using no_undef_atom by metis
qed

lemma MAGIC5:
  assumes invars: "ord_res_7_invars N (Uer, , Γ, 𝒞)" and
    no_more_steps: "𝒞'. ord_res_7 N (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞')"
  shows "(C. 𝒞 = Some C  is_least_nonskipped_clause N Uer  Γ C)"
proof (cases 𝒞)
  case None

  have "trail_atms Γ = atms_of_clss (N |∪| Uer)"
    using None invars[unfolded ord_res_7_invars_def] by simp

  have "C |∈| iefac  |`| (N |∪| Uer). trail_true_cls Γ C"
    using None invars[unfolded ord_res_7_invars_def] by simp

  hence no_false: "C |∈| iefac  |`| (N |∪| Uer). ¬ trail_false_cls Γ C"
    using invars[unfolded ord_res_7_invars_def]
    by (meson invars not_trail_true_cls_and_trail_false_cls
        ord_res_7_invars_implies_trail_consistent)

  have "C. is_least_nonskipped_clause N Uer  Γ C"
  proof (rule notI, elim exE)
    fix C
    assume "is_least_nonskipped_clause N Uer  Γ C"

    hence
      C_in: "C |∈| iefac  |`| (N |∪| Uer)" and
      C_spec_disj: "
        ord_res_8_can_decide_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C 
        ord_res_8_can_produce N Uer  Γ C 
        ord_res_8_can_factorize N Uer  Γ C"
      unfolding atomize_conj
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
      using no_false by metis

    thus False
    proof (elim disjE)
      assume "ord_res_8_can_decide_neg N Uer  Γ C"
      thus ?thesis
        using trail_atms Γ = atms_of_clss (N |∪| Uer)
        using ord_res_8_can_decide_neg.cases
        by (metis (no_types, lifting) linorder_trm.is_least_in_ffilter_iff)
    next
      assume "ord_res_8_can_skip_undefined_neg N Uer  Γ C"
      thus ?thesis
        using trail_atms Γ = atms_of_clss (N |∪| Uer)
        using ord_res_8_can_skip_undefined_neg.cases
        by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
            trail_defined_lit_iff_trail_defined_atm)
    next
      assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C"
      thus ?thesis
        using trail_atms Γ = atms_of_clss (N |∪| Uer)
        using ord_res_8_can_skip_undefined_pos_ultimate.cases
        by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
            trail_defined_lit_iff_trail_defined_atm)
    next
      assume "ord_res_8_can_produce N Uer  Γ C"
      thus False
        using trail_atms Γ = atms_of_clss (N |∪| Uer)
        using ord_res_8_can_produce.cases
        by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
            trail_defined_lit_iff_trail_defined_atm)
    next
      assume "ord_res_8_can_factorize N Uer  Γ C"
      thus False
        using trail_atms Γ = atms_of_clss (N |∪| Uer)
        using ord_res_8_can_factorize.cases
        by (metis C_in atm_of_in_atms_of_clssI linorder_lit.is_maximal_in_mset_iff
            trail_defined_lit_iff_trail_defined_atm)
    qed
  qed

  thus ?thesis
    using None by simp
next
  case (Some D)

  have D_in: "D |∈| iefac  |`| (N |∪| Uer)"
    using Some invars[unfolded ord_res_7_invars_def] by simp

  have "is_least_nonskipped_clause N Uer  Γ D"
  proof (cases "D = {#}")
    case True
    thus ?thesis
      using D_in is_least_nonskipped_clause_mempty by metis
  next
    case False

    then obtain LD where D_max_lit: "linorder_lit.is_maximal_in_mset D LD"
      using linorder_lit.ex_maximal_in_mset by presburger

    show ?thesis
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
    proof (intro conjI ballI impI)
      show "D |∈| iefac  |`| (N |∪| Uer)"
        using D_in .
    next
      show "trail_false_cls Γ D 
      ord_res_8_can_decide_neg N Uer  Γ D 
      ord_res_8_can_skip_undefined_neg N Uer  Γ D 
      ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D 
      ord_res_8_can_produce N Uer  Γ D  ord_res_8_can_factorize N Uer  Γ D"
      proof (cases "trail_false_cls Γ D")
        case True
        then show ?thesis
          by metis
      next
        case D_not_false: False

        then obtain L where D_max_lit: "linorder_lit.is_maximal_in_mset D L"
          by (metis linorder_lit.ex_maximal_in_mset trail_false_cls_mempty)

        show ?thesis
        proof (cases "A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ")
          case True

          hence "ord_res_8_can_decide_neg N Uer  Γ D"
            using ord_res_8_can_decide_neg.intros[OF D_not_false D_max_lit]
            by (metis (no_types, lifting) equalsffemptyD ffmember_filter linorder_trm.ex_least_in_fset)

          thus ?thesis
            by argo
        next
          case no_undef_atm: False
          show ?thesis
          proof (cases "trail_defined_lit Γ L")
            case L_defined: True

            hence "𝒞'. ord_res_7 N (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞')"
              unfolding 𝒞 = Some D
              using ord_res_7.skip_defined[OF D_not_false D_max_lit no_undef_atm]
              by metis

            hence False
              using no_more_steps by contradiction

            thus ?thesis ..
          next
            case L_undef: False
            show ?thesis
            proof (cases L)
              case (Pos A)

              hence L_pos: "is_pos L"
                by simp

              show ?thesis
              proof (cases "trail_false_cls Γ {#K ∈# D. K  L#}")
                case D_almost_false: True
                thus ?thesis
                  using ord_res_8_can_factorize.intros[
                      OF D_not_false D_max_lit no_undef_atm L_undef L_pos]
                  using ord_res_8_can_produce.intros[
                      OF D_not_false D_max_lit no_undef_atm L_undef L_pos]
                  by metis
              next
                case D_not_flagrantly_false: False
                show ?thesis
                proof (cases "E |∈| iefac  |`| (N |∪| Uer). D c E")
                  case True

                  hence "𝒞'. ord_res_7 N (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞')"
                    unfolding 𝒞 = Some D
                    using ord_res_7.skip_undefined_pos[
                        OF D_not_false D_max_lit no_undef_atm L_undef L_pos D_not_flagrantly_false]
                    by (metis femptyE ffmember_filter linorder_cls.ex_least_in_fset)

                  hence False
                    using no_more_steps by contradiction

                  thus ?thesis ..
                next
                  case False

                  hence "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D"
                    using ord_res_8_can_skip_undefined_pos_ultimate.intros[
                        OF D_not_false D_max_lit no_undef_atm L_undef L_pos D_not_flagrantly_false]
                    by metis

                  thus ?thesis
                    by argo
                qed
              qed
            next
              case (Neg A)
              hence L_neg: "is_neg L"
                by simp

              hence "ord_res_8_can_skip_undefined_neg N Uer  Γ D"
                unfolding 𝒞 = Some D
                using ord_res_8_can_skip_undefined_neg.intros[
                    OF D_not_false D_max_lit no_undef_atm L_undef]
                by metis

              thus ?thesis
                by argo
            qed
          qed
        qed
      qed

      fix E :: "'f gterm literal multiset"
      assume
        E_in: "E |∈| iefac  |`| (N |∪| Uer)" and
        E_neq: "E  D" and
        E_spec: "trail_false_cls Γ E 
        ord_res_8_can_decide_neg N Uer  Γ E 
        ord_res_8_can_skip_undefined_neg N Uer  Γ E 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ E 
        ord_res_8_can_produce N Uer  Γ E 
        ord_res_8_can_factorize N Uer  Γ E"

      have true_cls_if_lt_D:
        "C |∈| iefac  |`| (N |∪| Uer). C c D  trail_true_cls Γ C"
        using invars[unfolded ord_res_7_invars_def] Some by simp

      have Γ_lower_set: "linorder_trm.is_lower_fset (trail_atms Γ) (atms_of_clss (N |∪| Uer))"
        using invars[unfolded ord_res_7_invars_def] by simp

      have FOO:
        "C|∈|iefac  |`| (N |∪| Uer). C c D 
          (K. ord_res.is_maximal_lit K C 
          ¬ (A |∈| atms_of_clss (N |∪| Uer). A t atm_of K  A |∉| trail_atms Γ))"
        using invars[unfolded ord_res_7_invars_def] Some E_in by simp

      hence BAR:
        "C |∈| iefac  |`| (N |∪| Uer). C c D 
          (K. linorder_lit.is_maximal_in_mset C K 
            ¬ trail_defined_lit Γ K  (trail_true_cls Γ {#x ∈# C. x  K#}  is_pos K))"
        using invars[unfolded ord_res_7_invars_def] Some by simp

      show "D c E"
        using E_spec
      proof (elim disjE)
        assume "trail_false_cls Γ E"
        hence "¬ trail_true_cls Γ E"
          using invars not_trail_true_cls_and_trail_false_cls
            ord_res_7_invars_implies_trail_consistent by blast
        thus "D c E"
          using E_in E_neq true_cls_if_lt_D by force
      next
        assume "ord_res_8_can_decide_neg N Uer  Γ E"
        thus "D c E"
        proof (cases N Uer  Γ E rule: ord_res_8_can_decide_neg.cases)
          case (1 LE A)
          hence "A|∈|atms_of_clss (N |∪| Uer). A t atm_of LE  A |∉| trail_atms Γ"
            unfolding linorder_trm.is_least_in_ffilter_iff by metis
          thus ?thesis
            using FOO[rule_format, OF E_in _ ord_res.is_maximal_lit LE E] E_in E_neq
            by force
        qed
      next
        assume "ord_res_8_can_skip_undefined_neg N Uer  Γ E"
        thus "D c E"
        proof (cases N Uer  Γ E rule: ord_res_8_can_skip_undefined_neg.cases)
          case hyps: (1 LE)
          thus ?thesis
            using BAR[rule_format, OF E_in _ ord_res.is_maximal_lit LE E]
            using invars[unfolded ord_res_7_invars_def Some, rule_format, OF refl] Some E_in
            using E_neq by fastforce
        qed
      next
        assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ E"
        thus "D c E"
        proof (cases N Uer  Γ E rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
          case (1 L)
          then show ?thesis using E_neq D_in by force
        qed
      next
        assume "ord_res_8_can_produce N Uer  Γ E"
        hence "¬ trail_true_cls Γ E"
        proof (cases N Uer  Γ E rule: ord_res_8_can_produce.cases)
          case (1 L)
          then show ?thesis
            using invars[THEN ord_res_7_invars_implies_trail_consistent]
            by (smt (verit, ccfv_SIG) mem_Collect_eq not_trail_true_cls_and_trail_false_cls
                set_mset_filter trail_defined_lit_iff_true_or_false trail_true_cls_def)
        qed
        thus "D c E"
          using E_in E_neq true_cls_if_lt_D by force
      next
        assume "ord_res_8_can_factorize N Uer  Γ E"
        hence "¬ trail_true_cls Γ E"
        proof (cases N Uer  Γ E rule: ord_res_8_can_factorize.cases)
          case (1 L)
          then show ?thesis
            using invars[THEN ord_res_7_invars_implies_trail_consistent]
            by (smt (verit, ccfv_SIG) mem_Collect_eq not_trail_true_cls_and_trail_false_cls
                set_mset_filter trail_defined_lit_iff_true_or_false trail_true_cls_def)
        qed
        thus "D c E"
          using E_in E_neq true_cls_if_lt_D by force
      qed
    qed
  qed

  moreover have "Uniq (is_least_nonskipped_clause N Uer  Γ)"
    unfolding is_least_nonskipped_clause_def
    using linorder_cls.Uniq_is_least_in_fset
    by simp

  ultimately show ?thesis
    using Some
    by (metis (no_types) The_optional_eq_SomeD The_optional_eq_SomeI)
qed

lemma MAGIC6:
  assumes invars: "ord_res_7_invars N (Uer, , Γ, 𝒞)"
  shows "𝒞'. (ord_res_7 N)** (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞') 
    (𝒞''. ord_res_7 N (Uer, , Γ, 𝒞') (Uer, , Γ, 𝒞''))"
proof -
  define R where
    "R = (λ𝒞 𝒞'. ord_res_7 N (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞'))"

  define f :: "'f gclause option  'f gclause fset" where
    "f = (λ𝒞. case 𝒞 of None  {||} | Some C  {|D |∈| iefac  |`| (N |∪| Uer). C c D|})"

  let ?less_f = "(|⊂|)"

  have "Wellfounded.wfp_on {x'. R** 𝒞 x'} R¯¯"
  proof (rule wfp_on_if_convertible_to_wfp_on)
    have "wfp (|⊂|)"
      by auto
    thus "Wellfounded.wfp_on (f ` {x'. R** 𝒞 x'}) ?less_f"
      using Wellfounded.wfp_on_subset subset_UNIV by metis
  next
    fix 𝒞x 𝒞y :: "'f gclause option"

    have rtranclp_with_constsD: "(λy y'. R (x, y) (x, y'))** y y' 
      R** (x, y) (x, y')" for R x y y'
    proof (induction y arbitrary: rule: converse_rtranclp_induct)
      case base
      show ?case
        by simp
    next
      case (step w)
      thus ?case
        by force
    qed 

    assume "𝒞x  {x'. R** 𝒞 x'}" and  "𝒞y  {x'. R** 𝒞 x'}"
    hence
      "(ord_res_7 N)** (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞x)" and
      "(ord_res_7 N)** (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞y)"
      unfolding atomize_conj mem_Collect_eq R_def
      by (auto intro: rtranclp_with_constsD)
    hence
      x_invars: "ord_res_7_invars N (Uer, , Γ, 𝒞x)" and
      y_invars: "ord_res_7_invars N (Uer, , Γ, 𝒞y)"
      using ord_res_7_preserves_invars
      using invars by (metis rtranclp_ord_res_7_preserves_ord_res_7_invars)+

    have Γ_consistent: "trail_consistent Γ"
      using x_invars by (metis ord_res_7_invars_implies_trail_consistent)

    have less_f_if: "?less_f (f 𝒞y) (f 𝒞x)"
      if "𝒞x = Some C" and
        𝒞y_disj: "𝒞y = None  𝒞y = Some D  C c D" and
        C_in: "C |∈| iefac  |`| (N |∪| Uer)"
      for C D
    proof -
      have f_x: "f 𝒞x = {|D |∈| iefac  |`| (N |∪| Uer). C c D|}"
        by (auto simp add: 𝒞x = Some C f_def)

      moreover have "C |∈| f 𝒞x"
        using C_in f_x by simp

      moreover have "C |∉| f 𝒞y  f 𝒞y |⊆| f 𝒞x"
        using 𝒞y_disj
      proof (elim disjE conjE; intro conjI)
        assume "𝒞y = None"
        thus "C |∉| f 𝒞y" and "f 𝒞y |⊆| f 𝒞x"
          unfolding f_x
          by (simp_all add: f_def)
      next
        assume "𝒞y = Some D" and "C c D"

        have "x. D c x  C c x"
          using C c D by auto

        moreover have fst_f_y: "f 𝒞y = {|E |∈| iefac  |`| (N |∪| Uer). D c E|}"
          by (auto simp add: 𝒞y = Some D f_def)

        ultimately show "f 𝒞y |⊆| f 𝒞x"
          using f_x by auto

        show "C |∉| f 𝒞y"
          using C c D fst_f_y by auto
      qed

      ultimately have "f 𝒞y |⊂| f 𝒞x"
        by blast

      thus ?thesis
        by (simp add: lex_prodp_def)
    qed

    have eres_not_in_if: "eres D E |∉| Uer"
      if "𝒞x = Some E" and E_false: "trail_false_cls Γ E" and
        E_max_lit: "ord_res.is_maximal_lit L E" and L_neg: "is_neg L"
        "map_of Γ (- L) = Some (Some D)"
      for D E L
    proof -
      have
        clauses_lt_E_true:
        "C |∈| iefac  |`| (N |∪| Uer). C c E  trail_true_cls Γ C" and
        Γ_prop_greatest:
        "Ln  set Γ. C. snd Ln = Some C  linorder_lit.is_greatest_in_mset C (fst Ln)"
        using x_invars unfolding 𝒞x = Some E ord_res_7_invars_def by simp_all

      have "(- L, Some D)  set Γ"
        using map_of Γ (- L) = Some (Some D) by (metis map_of_SomeD)

      hence D_greatest_lit: "linorder_lit.is_greatest_in_mset D (- L)"
        using Γ_prop_greatest by fastforce

      have "eres D E c E"
        using eres_lt_if
        using E_max_lit L_neg D_greatest_lit
        by metis

      hence "eres D E  E"
        by order

      have "L ∈# E"
        using E_max_lit unfolding linorder_lit.is_maximal_in_mset_iff by metis

      hence "- L ∉# E"
        using not_both_lit_and_comp_lit_in_false_clause_if_trail_consistent
        using Γ_consistent E_false by metis

      hence "K ∈# eres D E. atm_of K t atm_of L"
        using lit_in_eres_lt_greatest_lit_in_grestest_resolvant[OF eres D E  E E_max_lit]
        by metis

      hence "K ∈# eres D E. K  L  K  - L"
        by fastforce

      moreover have "L ∈# eres D E. L ∈# D  L ∈# E"
        using lit_in_one_of_resolvents_if_in_eres by metis

      moreover have D_almost_false: "trail_false_cls Γ {#K ∈# D. K  - L#}"
        using ord_res_7_invars_implies_propagated_clause_almost_false
        using (- L, Some D)  set Γ x_invars
        by metis

      ultimately have "trail_false_cls Γ (eres D E)"
        using E_false unfolding trail_false_cls_def by fastforce

      have "eres D E |∉| N |∪| Uer"
        using eres_not_in_known_clauses_if_trail_false_cls
        using Γ_consistent clauses_lt_E_true eres D E c E trail_false_cls Γ (eres D E)
        by metis

      thus "eres D E |∉| Uer"
        by blast
    qed

    assume "R¯¯ 𝒞y 𝒞x"
    hence "ord_res_7 N (Uer, , Γ, 𝒞x) (Uer, , Γ, 𝒞y)"
      unfolding conversep_iff R_def .
    thus "?less_f (f 𝒞y) (f 𝒞x)"
    proof (cases N "(Uer, , Γ, 𝒞x)" "(Uer, , Γ, 𝒞y)" rule: ord_res_7.cases)
      case step_hyps: (decide_neg C L A)
      hence False by simp
      thus ?thesis ..
    next
      case step_hyps: (skip_defined C L)

      have "C |∈| iefac  |`| (N |∪| Uer)"
        using ord_res_7_invars_def step_hyps(1) x_invars by presburger

      moreover have "D. 𝒞y = None  𝒞y = Some D  C c D"
      proof (cases 𝒞y)
        case None
        thus ?thesis
          by simp
      next
        case (Some D)
        thus ?thesis
          using step_hyps
          by (metis linorder_cls.is_least_in_ffilter_iff Some_eq_The_optionalD)
      qed

      ultimately show ?thesis
        using less_f_if step_hyps by metis
    next
      case step_hyps: (skip_undefined_neg C L)
      hence False by simp
      thus ?thesis ..
    next
      case step_hyps: (skip_undefined_pos C L D)

      have "C |∈| iefac  |`| (N |∪| Uer)"
        using ord_res_7_invars_def step_hyps(1) x_invars by presburger

      moreover have "D. 𝒞y = None  𝒞y = Some D  C c D"
        using step_hyps by (metis linorder_cls.is_least_in_ffilter_iff)

      ultimately show ?thesis
        using less_f_if step_hyps by metis
    next
      case step_hyps: (skip_undefined_pos_ultimate C L)
      hence False by simp
      thus ?thesis ..
    next
      case step_hyps: (production C L)
      hence False by simp
      thus ?thesis ..
    next
      case step_hyps: (factoring C L)

      have "C |∈| iefac  |`| (N |∪| Uer)"
        using ord_res_7_invars_def step_hyps(1) x_invars by presburger

      moreover have "efac C  C"
        using step_hyps by (metis greatest_literal_in_efacI)
      
      ultimately have "C |∉| "
        by (smt (verit, ccfv_threshold) fimage_iff iefac_def ex1_efac_eq_factoring_chain
            factorizable_if_neq_efac)

      hence False
        using  = finsert C  by blast

      thus ?thesis ..
    next
      case step_hyps: (resolution_bot E L D)
      hence "eres D E |∉| Uer"
        using eres_not_in_if by metis
      hence False
        using Uer = finsert (eres D E) Uer by blast
      thus ?thesis ..
    next
      case (resolution_pos E L D K)
      hence "eres D E |∉| Uer"
        using eres_not_in_if by metis
      hence False
        using Uer = finsert (eres D E) Uer by blast
      thus ?thesis ..
    next
      case (resolution_neg E L D K C)
      hence "eres D E |∉| Uer"
        using eres_not_in_if by metis
      hence False
        using Uer = finsert (eres D E) Uer by blast
      thus ?thesis ..
    qed
  qed

  then obtain 𝒞' where "R** 𝒞 𝒞'" and "z. R 𝒞' z"
    using ex_terminating_rtranclp_strong by metis

  show ?thesis
  proof (intro exI conjI)
    show "(ord_res_7 N)** (Uer, , Γ, 𝒞) (Uer, , Γ, 𝒞')"
      using R** 𝒞 𝒞'
      by (induction 𝒞 rule: converse_rtranclp_induct) (auto simp: R_def)
  next
    show "𝒞''. ord_res_7 N (Uer, , Γ, 𝒞') (Uer, , Γ, 𝒞'')"
      using z. R 𝒞' z
      by (simp add: R_def)
  qed
qed

inductive ord_res_7_matches_ord_res_8 :: "'f ord_res_7_state  'f ord_res_8_state  bool" where
  "ord_res_7_invars N (Uer, , Γ, 𝒞) 
    ord_res_8_invars N (Uer, , Γ) 
    (C. 𝒞 = Some C  is_least_nonskipped_clause N Uer  Γ C) 
    ord_res_7_matches_ord_res_8 (N, Uer, , Γ, 𝒞) (N, Uer, , Γ)"

lemma ord_res_7_final_iff_ord_res_8_final:
  fixes S7 S8
  assumes match: "ord_res_7_matches_ord_res_8 S7 S8"
  shows "ord_res_7_final S7  ord_res_8_final S8"
  using match
proof (cases S7 S8 rule: ord_res_7_matches_ord_res_8.cases)
  case match_hyps: (1 N Uer  Γ 𝒞)

  note invars7 = ord_res_7_invars N (Uer, , Γ, 𝒞)[unfolded ord_res_7_invars_def,
      rule_format, OF refl]

  have Γ_consistent: "trail_consistent Γ"
    using invars7 by (metis trail_consistent_if_sorted_wrt_atoms)

  show "ord_res_7_final S7  ord_res_8_final S8"
  proof (rule iffI)
    assume "ord_res_7_final S7"
    thus "ord_res_8_final S8"
      unfolding S7 = (N, Uer, , Γ, 𝒞)
    proof (cases "(N, Uer, , Γ, 𝒞)" rule: ord_res_7_final.cases)
      case model_found
      show "ord_res_8_final S8"
        unfolding S8 = (N, Uer, , Γ)
      proof (rule ord_res_8_final.model_found)
        have "𝒞 = None  trail_atms Γ = atms_of_clss (N |∪| Uer)"
          using invars7 by argo

        hence "trail_atms Γ = atms_of_clss (N |∪| Uer)"
          using model_found by argo

        thus "¬ (A|∈|atms_of_clss (N |∪| Uer). A |∉| trail_atms Γ)"
          by metis
      next
        have "C |∈| iefac  |`| (N |∪| Uer). trail_true_cls Γ C"
          using invars7 model_found by simp

        moreover have "¬ (trail_true_cls Γ C  trail_false_cls Γ C)" for C
          using not_trail_true_cls_and_trail_false_cls[OF Γ_consistent] .

        ultimately show "¬ (C|∈|iefac  |`| (N |∪| Uer). trail_false_cls Γ C)"
          by metis
      qed
    next
      case contradiction_found
      show "ord_res_8_final S8"
        unfolding S8 = (N, Uer, , Γ)
      proof (rule ord_res_8_final.contradiction_found)
        show "{#} |∈| iefac  |`| (N |∪| Uer)"
          using invars7 𝒞 = Some {#} by metis
      qed
    qed
  next
    assume "ord_res_8_final S8"
    thus "ord_res_7_final S7"
      unfolding S8 = (N, Uer, , Γ)
    proof (cases "(N, Uer, , Γ)" rule: ord_res_8_final.cases)
      case model_found

      hence "C. is_least_nonskipped_clause N Uer  Γ C"
        using nex_is_least_nonskipped_clause_if by metis

      hence "𝒞 = None"
        using match_hyps by simp

      thus "ord_res_7_final S7"
        unfolding S7 = (N, Uer, , Γ, 𝒞)
        using ord_res_7_final.model_found by metis
    next
      case contradiction_found

      hence "is_least_nonskipped_clause N Uer  Γ {#}"
        using is_least_nonskipped_clause_mempty by metis

      hence "𝒞 = Some {#}"
        using match_hyps by presburger

      thus "ord_res_7_final S7"
        unfolding S7 = (N, Uer, , Γ, 𝒞)
        using ord_res_7_final.contradiction_found by metis
    qed
  qed
qed

lemma backward_simulation_between_7_and_8:
  fixes i S7 S8 S8'
  assumes match: "ord_res_7_matches_ord_res_8 S7 S8" and step: "constant_context ord_res_8 S8 S8'"
  shows "S7'. (constant_context ord_res_7)++ S7 S7'  ord_res_7_matches_ord_res_8 S7' S8'"
  using match
proof (cases S7 S8 rule: ord_res_7_matches_ord_res_8.cases)
  case match_hyps: (1 N Uer  Γ 𝒞)

  note S7_def = S7 = (N, Uer, , Γ, 𝒞)

  note invars7 = ord_res_7_invars N (Uer, , Γ, 𝒞)[unfolded ord_res_7_invars_def,
      rule_format, OF refl]

  have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) Γ"
    using invars7 by argo

  have Γ_consistent: "trail_consistent Γ"
    using trail_consistent_if_sorted_wrt_atoms[OF Γ_sorted] .

  have Γ_lower_set: "linorder_trm.is_lower_fset (trail_atms Γ) (atms_of_clss (N |∪| Uer))"
    using invars7 by argo

  have 𝒞_eq_None_implies_all_atoms_defined: "𝒞 = None  trail_atms Γ = atms_of_clss (N |∪| Uer)"
    using invars7 by argo

  obtain s8' where
    S8'_def: "S8' = (N, s8')" and
    step': "ord_res_8 N (Uer, , Γ) s8'"
    using step unfolding S8 = (N, Uer, , Γ)
    by (auto elim: constant_context.cases)

  have invars_s8': "ord_res_8_invars N s8'"
    using ord_res_8_preserves_invars[OF step' ord_res_8_invars N (Uer, , Γ)] .

  show ?thesis
    using step'
  proof (cases N "(Uer, , Γ)" s8' rule: ord_res_8.cases)
    case step_hyps: (decide_neg A Γ')

    have
      A_in: "A |∈| atms_of_clss (N |∪| Uer)" and
      A_undef: "A |∉| trail_atms Γ" and
      A_least: "y|∈|atms_of_clss (N |∪| Uer). y  A  (A1|∈|trail_atms Γ. A1 t y)  A t y"
      using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto

    have "𝒞  None"
      using 𝒞_eq_None_implies_all_atoms_defined A_in A_undef by metis

    then obtain D :: "'f gclause" where "𝒞 = Some D"
      by blast

    hence D_in: "D |∈| iefac  |`| (N |∪| Uer)"
      by (metis 𝒞 = Some D invars7)

    have "is_least_nonskipped_clause N Uer  Γ D"
      using match_hyps 𝒞 = Some D by metis

    moreover have D_not_false: "¬ trail_false_cls Γ D"
      using D_in step_hyps by metis

    moreover have "¬ ord_res_8_can_produce N Uer  Γ D"
    proof (rule notI)
      assume "ord_res_8_can_produce N Uer  Γ D"
      thus False
      proof (cases N Uer  Γ D rule: ord_res_8_can_produce.cases)
        case (1 L')
        thus ?thesis
          by (metis A_in A_least A_undef D_in atm_of_in_atms_of_clssI
              atoms_of_trail_lt_atom_of_propagatable_literal clause_could_propagate_def invars7
              linorder_lit.is_maximal_in_mset_iff literal.collapse(1) step_hyps(4))
      qed
    qed

    moreover have "¬ ord_res_8_can_factorize N Uer  Γ D"
    proof (rule notI)
      assume "ord_res_8_can_factorize N Uer  Γ D"
      thus False
      proof (cases N Uer  Γ D rule: ord_res_8_can_factorize.cases)
        case (1 L')
        thus False
          by (metis A_in A_least A_undef D_in atm_of_in_atms_of_clssI
              atoms_of_trail_lt_atom_of_propagatable_literal clause_could_propagate_def invars7
              linorder_lit.is_maximal_in_mset_iff literal.collapse(1) step_hyps(4))
      qed
    qed

    ultimately have "ord_res_8_can_decide_neg N Uer  Γ D 
      ord_res_8_can_skip_undefined_neg N Uer  Γ D 
      ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D"
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
      by argo

    then obtain 𝒞' where first_step7: "ord_res_7 N (Uer, , Γ, Some D) (Uer, , Γ', 𝒞')"
    proof (elim disjE; atomize_elim)
      assume "ord_res_8_can_decide_neg N Uer  Γ D"
      thus "𝒞'. ord_res_7 N (Uer, , Γ, Some D) (Uer, , Γ', 𝒞')"
      proof (cases N Uer  Γ D rule: ord_res_8_can_decide_neg.cases)
        case hyps: (1 L A')
        hence "A' = A"
          by (smt (verit, del_insts) Γ_lower_set linorder_trm.is_least_in_ffilter_iff
              linorder_trm.neq_iff linorder_trm.not_in_lower_setI linorder_trm.order.strict_trans
              step_hyps(3))
        thus ?thesis
          using hyps Γ' = (Neg A, None) # Γ
          using ord_res_7.decide_neg[of Γ D _ N Uer A Γ' ] by blast
      qed
    next
      assume "ord_res_8_can_skip_undefined_neg N Uer  Γ D"
      thus "𝒞'. ord_res_7 N (Uer, , Γ, Some D) (Uer, , Γ', 𝒞')"
      proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
        case hyps: (1 L)
        hence "L = Neg A"
          by (smt (verit) A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI literal.disc(2) literal.expand literal.sel(2)
              trail_defined_lit_iff_trail_defined_atm)
        thus ?thesis
          using hyps Γ' = (Neg A, None) # Γ
          using ord_res_7.skip_undefined_neg by blast
      qed
    next
      assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D"
      thus "𝒞'. ord_res_7 N (Uer, , Γ, Some D) (Uer, , Γ', 𝒞')"
      proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
        case hyps: (1 L)
        hence "L = Pos A"
          by (smt (verit, best) A_in A_least A_undef D_in atm_of_in_atms_of_clssI invars7
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI literal.disc(1) literal.expand literal.sel(1)
              trail_defined_lit_iff_trail_defined_atm)
        thus ?thesis
          using hyps Γ' = (Neg A, None) # Γ
          using ord_res_7.skip_undefined_pos_ultimate by fastforce
      qed
    qed

    moreover have "ord_res_7_invars N (Uer, , Γ', 𝒞')"
      using 𝒞 = Some D first_step7 match_hyps(3) ord_res_7_preserves_invars by blast

    ultimately obtain 𝒞'' where
      following_steps7: "(ord_res_7 N)** (Uer, , Γ', 𝒞') (Uer, , Γ', 𝒞'')" and
      no_more_step7: "(𝒞'''. ord_res_7 N (Uer, , Γ', 𝒞'') (Uer, , Γ', 𝒞'''))"
      using MAGIC6 by metis

    show ?thesis
    proof (intro exI conjI)
      have "(ord_res_7 N)++ (Uer, , Γ, 𝒞) (Uer, , Γ', 𝒞'')"
        unfolding 𝒞 = Some D
        using first_step7 following_steps7 by simp

      thus "(constant_context ord_res_7)++ S7 (N, Uer, , Γ', 𝒞'')"
        unfolding S7_def by (simp add: tranclp_constant_context)

      show "ord_res_7_matches_ord_res_8 (N, Uer, , Γ', 𝒞'') S8'"
        unfolding S8'_def s8' = (Uer, , Γ')
      proof (intro ord_res_7_matches_ord_res_8.intros allI)
        show "ord_res_7_invars N (Uer, , Γ', 𝒞'')"
          using ord_res_7_invars N (Uer, , Γ', 𝒞') following_steps7
            rtranclp_ord_res_7_preserves_ord_res_7_invars by blast

        show "ord_res_8_invars N (Uer, , Γ')"
          using invars_s8' step_hyps(1) by blast

        fix C :: "'f gclause"
        show "𝒞'' = Some C  is_least_nonskipped_clause N Uer  Γ' C"
          using MAGIC5 ord_res_7_invars N (Uer, , Γ', 𝒞'') no_more_step7 by metis
      qed
    qed
  next
    case step_hyps: (propagate A C Γ')

    have
      A_in: "A |∈| atms_of_clss (N |∪| Uer)" and
      A_undef: "A |∉| trail_atms Γ" and
      A_least: "y|∈|atms_of_clss (N |∪| Uer). y  A  (A1|∈|trail_atms Γ. A1 t y)  A t y"
      using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto

    have
      C_in: "C |∈| iefac  |`| (N |∪| Uer)" and
      C_can_prop:"clause_could_propagate Γ C (Pos A)" and
      C_least: "D |∈| iefac  |`| (N |∪| Uer).
        D  C  clause_could_propagate Γ D (Pos A)  C c D"
      using step_hyps unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo

    hence
      Pos_A_undef: "¬ trail_defined_lit Γ (Pos A)" and
      C_max_lit: "linorder_lit.is_maximal_in_mset C (Pos A)" and
      C_almost_false: "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
      unfolding atomize_conj clause_could_propagate_def by argo

    have "is_least_nonskipped_clause N Uer  Γ C"
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
    proof (intro conjI ballI impI)
      show "C |∈| iefac  |`| (N |∪| Uer)"
        using C_in .
    next
      have "ord_res_8_can_produce N Uer  Γ C"
      proof (rule ord_res_8_can_produce.intros)
        show "¬ trail_false_cls Γ C"
          using step_hyps C_in by metis
      next
        show "ord_res.is_maximal_lit (Pos A) C"
          using step_hyps by blast
      next
        show "¬ (Aa|∈|atms_of_clss (N |∪| Uer). Aa t atm_of (Pos A)  Aa |∉| trail_atms Γ)"
          unfolding literal.sel
          using step_hyps
          by (smt (verit, ccfv_threshold) Γ_lower_set linorder_trm.dual_order.asym
              linorder_trm.is_least_in_ffilter_iff linorder_trm.is_lower_set_iff
              linorder_trm.neq_iff)
      next
        show "¬ trail_defined_lit Γ (Pos A)"
          using A_undef unfolding trail_defined_lit_iff_trail_defined_atm literal.sel .
      next
        show "is_pos (Pos A)"
          by simp
      next
        show "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
          using clause_could_propagate Γ C (Pos A)
          unfolding clause_could_propagate_def by argo
      next
        show "ord_res.is_strictly_maximal_lit (Pos A) C"
          using step_hyps by argo
      qed

      thus "trail_false_cls Γ C 
        ord_res_8_can_decide_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C 
        ord_res_8_can_produce N Uer  Γ C  ord_res_8_can_factorize N Uer  Γ C"
        by argo
    next
      fix D :: "'f gclause"
      assume
        D_in: "D |∈| iefac  |`| (N |∪| Uer)" and
        D_neq: "D  C" and
        D_spec_disj: "trail_false_cls Γ D 
          ord_res_8_can_decide_neg N Uer  Γ D 
          ord_res_8_can_skip_undefined_neg N Uer  Γ D 
          ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D 
          ord_res_8_can_produce N Uer  Γ D 
          ord_res_8_can_factorize N Uer  Γ D"

      hence "
        ord_res_8_can_decide_neg N Uer  Γ D 
        ord_res_8_can_skip_undefined_neg N Uer  Γ D 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D 
        ord_res_8_can_produce N Uer  Γ D 
        ord_res_8_can_factorize N Uer  Γ D"
        using step_hyps D_in by metis

      thus "C c D"
      proof (elim disjE)
        assume "ord_res_8_can_decide_neg N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_decide_neg.cases)
          case hyps: (1 L A')
          hence "A = A'"
            using step_hyps
            by (smt (verit, del_insts) Γ_lower_set linorder_trm.antisym_conv3
                linorder_trm.dual_order.strict_implies_not_eq linorder_trm.dual_order.strict_trans
                linorder_trm.is_least_in_ffilter_iff linorder_trm.not_in_lower_setI)
          hence "A t atm_of L"
            using hyps
            unfolding linorder_trm.is_least_in_ffilter_iff
            by argo
          hence "Pos A l L"
            by (cases L) simp_all
          thus ?thesis
            using C_max_lit linorder_lit.is_maximal_in_mset D L
            by (metis linorder_lit.multp_if_maximal_less_that_maximal)
        qed
      next
        assume "ord_res_8_can_skip_undefined_neg N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
          case hyps: (1 L)
          hence "atm_of L = A"
            using step_hyps
            by (smt (verit, best)
                A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
                linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
                linorder_trm.not_in_lower_setI trail_defined_lit_iff_trail_defined_atm)
          hence "Pos A l L"
            using is_neg L by (cases L) simp_all
          thus ?thesis
            using C_max_lit linorder_lit.is_maximal_in_mset D L
            by (metis linorder_lit.multp_if_maximal_less_that_maximal)
        qed
      next
        assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
          case hyps: (1 L)
          thus ?thesis
            by (meson C_in D_neq linorder_cls.linorder_cases)
        qed
      next
        assume "ord_res_8_can_produce N Uer  Γ D"

        then obtain L where
          D_max_lit: "ord_res.is_maximal_lit L D" and
          "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)" and
          L_undef: "¬ trail_defined_lit Γ L" and
          D_almost_false: "trail_false_cls Γ {#K ∈# D. K  L#}" and
          "is_pos L"
          by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)

        hence "atm_of L = A"
          using step_hyps
          by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
              trail_defined_lit_iff_trail_defined_atm)

        hence "L = Pos A"
          using is_pos L by (cases L) simp_all

        hence "clause_could_propagate Γ D (Pos A)"
          unfolding clause_could_propagate_def
          using D_almost_false D_max_lit L_undef by metis

        thus "C c D"
          using D_in D_neq C_least by metis
      next
        assume "ord_res_8_can_factorize N Uer  Γ D"

        then obtain L where
          D_max_lit: "ord_res.is_maximal_lit L D" and
          "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)" and
          L_undef: "¬ trail_defined_lit Γ L" and
          D_almost_false: "trail_false_cls Γ {#K ∈# D. K  L#}" and
          "is_pos L"
          by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)

        hence "atm_of L = A"
          using step_hyps
          by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
              trail_defined_lit_iff_trail_defined_atm)

        hence "L = Pos A"
          using is_pos L by (cases L) simp_all

        hence "clause_could_propagate Γ D (Pos A)"
          unfolding clause_could_propagate_def
          using D_almost_false D_max_lit L_undef by metis

        thus "C c D"
          using D_in D_neq C_least by metis
      qed
    qed

    hence "𝒞 = Some C"
      using match_hyps by metis

    define 𝒞' where
      "𝒞' = The_optional (linorder_cls.is_least_in_fset
        (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"

    have first_step7: "ord_res_7 N (Uer, , Γ, Some C) (Uer, , Γ', 𝒞')"
    proof (rule ord_res_7.production)
      show "¬ trail_false_cls Γ C"
        using C_in step_hyps(2) by blast
    next
      show "ord_res.is_maximal_lit (Pos A) C"
        using C_max_lit by force
    next
      show "¬ (Aa|∈|atms_of_clss (N |∪| Uer). Aa t atm_of (Pos A)  Aa |∉| trail_atms Γ)"
        by (metis A_least Γ_lower_set linorder_trm.dual_order.asym linorder_trm.neq_iff
            linorder_trm.not_in_lower_setI literal.sel(1))
    next
      show "¬ trail_defined_lit Γ (Pos A)"
        using Pos_A_undef .
    next
      show "is_pos (Pos A)"
        by simp
    next
      show "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
        using C_almost_false .
    next
      show "ord_res.is_strictly_maximal_lit (Pos A) C"
        using step_hyps by argo
    next
      show "Γ' = (Pos A, Some C) # Γ"
        using step_hyps by argo
    next
      show "𝒞' = The_optional (linorder_cls.is_least_in_fset (ffilter ((≺c) C) (iefac  |`| (N |∪| Uer))))"
        using 𝒞'_def .
    qed

    moreover have "ord_res_7_invars N (Uer, , Γ', 𝒞')"
      using 𝒞 = Some C first_step7 match_hyps(3) ord_res_7_preserves_invars by blast

    ultimately obtain 𝒞'' where
      following_steps7: "(ord_res_7 N)** (Uer, , Γ', 𝒞') (Uer, , Γ', 𝒞'')" and
      no_more_step7: "(𝒞'''. ord_res_7 N (Uer, , Γ', 𝒞'') (Uer, , Γ', 𝒞'''))"
      using MAGIC6 by metis

    show ?thesis
    proof (intro exI conjI)
      have "(ord_res_7 N)++ (Uer, , Γ, 𝒞) (Uer, , Γ', 𝒞'')"
        unfolding 𝒞 = Some C
        using first_step7 following_steps7 by simp

      thus "(constant_context ord_res_7)++ S7 (N, Uer, , Γ', 𝒞'')"
        unfolding S7_def by (simp add: tranclp_constant_context)

      show "ord_res_7_matches_ord_res_8 (N, Uer, , Γ', 𝒞'') S8'"
        unfolding S8'_def s8' = (Uer, , Γ')
      proof (intro ord_res_7_matches_ord_res_8.intros allI)
        show "ord_res_7_invars N (Uer, , Γ', 𝒞'')"
          using ord_res_7_invars N (Uer, , Γ', 𝒞') following_steps7
            rtranclp_ord_res_7_preserves_ord_res_7_invars by blast

        show "ord_res_8_invars N (Uer, , Γ')"
          using invars_s8' step_hyps(1) by blast

        fix C :: "'f gclause"
        show "𝒞'' = Some C  is_least_nonskipped_clause N Uer  Γ' C"
          using MAGIC5 ord_res_7_invars N (Uer, , Γ', 𝒞'') no_more_step7 by metis
      qed
    qed
  next
    case step_hyps: (factorize A C ℱ')

    have
      A_in: "A |∈| atms_of_clss (N |∪| Uer)" and
      A_undef: "A |∉| trail_atms Γ" and
      A_least: "y|∈|atms_of_clss (N |∪| Uer). y  A  (A1|∈|trail_atms Γ. A1 t y)  A t y"
      using step_hyps(3) unfolding linorder_trm.is_least_in_fset_iff by auto

    have
      C_in: "C |∈| iefac  |`| (N |∪| Uer)" and
      C_can_prop:"clause_could_propagate Γ C (Pos A)" and
      C_least: "D |∈| iefac  |`| (N |∪| Uer).
        D  C  clause_could_propagate Γ D (Pos A)  C c D"
      using step_hyps unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo

    hence
      Pos_A_undef: "¬ trail_defined_lit Γ (Pos A)" and
      C_max_lit: "linorder_lit.is_maximal_in_mset C (Pos A)" and
      C_almost_false: "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
      unfolding atomize_conj clause_could_propagate_def by argo

    have C_not_false: "¬ trail_false_cls Γ C"
      using C_in step_hyps by metis

    have no_undef_atm_lt_A:
      "¬ (Aa|∈|atms_of_clss (N |∪| Uer). Aa t A  Aa |∉| trail_atms Γ)"
      by (metis A_least Γ_lower_set linorder_trm.dual_order.asym linorder_trm.neq_iff
          linorder_trm.not_in_lower_setI)

    have "is_least_nonskipped_clause N Uer  Γ C"
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
    proof (intro conjI ballI impI)
      show "C |∈| iefac  |`| (N |∪| Uer)"
        using C_in .
    next
      have "ord_res_8_can_factorize N Uer  Γ C"
      proof (intro ord_res_8_can_factorize.intros)
        show "¬ trail_false_cls Γ C"
          using C_not_false .
      next
        show "ord_res.is_maximal_lit (Pos A) C"
          using C_max_lit .
      next
        show "¬ (Aa|∈|atms_of_clss (N |∪| Uer). Aa t atm_of (Pos A)  Aa |∉| trail_atms Γ)"
          using no_undef_atm_lt_A by simp
      next
        show "¬ trail_defined_lit Γ (Pos A)"
          using Pos_A_undef .
      next
        show "is_pos (Pos A)"
          by simp
      next
        show "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
          using C_almost_false .
      next
        show "¬ ord_res.is_strictly_maximal_lit (Pos A) C"
          using step_hyps by argo
      qed

      thus "trail_false_cls Γ C 
        ord_res_8_can_decide_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_neg N Uer  Γ C 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ C 
        ord_res_8_can_produce N Uer  Γ C 
        ord_res_8_can_factorize N Uer  Γ C"
        using ord_res_8_can_factorize N Uer  Γ C by argo
    next
      fix D :: "'f gclause"
      assume
        D_in: "D |∈| iefac  |`| (N |∪| Uer)" and
        D_neq: "D  C" and
        D_spec_disj: "trail_false_cls Γ D 
          ord_res_8_can_decide_neg N Uer  Γ D 
          ord_res_8_can_skip_undefined_neg N Uer  Γ D 
          ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D 
          ord_res_8_can_produce N Uer  Γ D 
          ord_res_8_can_factorize N Uer  Γ D"

      hence "
        ord_res_8_can_decide_neg N Uer  Γ D 
        ord_res_8_can_skip_undefined_neg N Uer  Γ D 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D 
        ord_res_8_can_produce N Uer  Γ D 
        ord_res_8_can_factorize N Uer  Γ D"
        using step_hyps D_in by metis

      thus "C c D"
      proof (elim disjE)
        assume "ord_res_8_can_decide_neg N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_decide_neg.cases)
          case hyps: (1 L A')
          hence "A = A'"
            using step_hyps
            by (smt (verit, del_insts) Γ_lower_set linorder_trm.antisym_conv3
                linorder_trm.dual_order.strict_implies_not_eq linorder_trm.dual_order.strict_trans
                linorder_trm.is_least_in_ffilter_iff linorder_trm.not_in_lower_setI)
          hence "A t atm_of L"
            using hyps
            unfolding linorder_trm.is_least_in_ffilter_iff
            by argo
          hence "Pos A l L"
            by (cases L) simp_all
          thus ?thesis
            using C_max_lit linorder_lit.is_maximal_in_mset D L
            by (metis linorder_lit.multp_if_maximal_less_that_maximal)
        qed
      next
        assume "ord_res_8_can_skip_undefined_neg N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_neg.cases)
          case hyps: (1 L)
          hence "atm_of L = A"
            using step_hyps
            by (smt (verit, best)
                A_in A_least A_undef D_in Γ_lower_set atm_of_in_atms_of_clssI
                linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
                linorder_trm.not_in_lower_setI trail_defined_lit_iff_trail_defined_atm)
          hence "Pos A l L"
            using is_neg L by (cases L) simp_all
          thus ?thesis
            using C_max_lit linorder_lit.is_maximal_in_mset D L
            by (metis linorder_lit.multp_if_maximal_less_that_maximal)
        qed
      next
        assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ D"
        thus "C c D"
        proof (cases N Uer  Γ D rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
          case hyps: (1 L)
          thus ?thesis
            by (meson C_in D_neq linorder_cls.linorder_cases)
        qed
      next
        assume "ord_res_8_can_produce N Uer  Γ D"

        then obtain L where
          D_max_lit: "ord_res.is_maximal_lit L D" and
          "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)" and
          L_undef: "¬ trail_defined_lit Γ L" and
          D_almost_false: "trail_false_cls Γ {#K ∈# D. K  L#}" and
          "is_pos L"
          by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)

        hence "atm_of L = A"
          using step_hyps
          by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
              trail_defined_lit_iff_trail_defined_atm)

        hence "L = Pos A"
          using is_pos L by (cases L) simp_all

        hence "clause_could_propagate Γ D (Pos A)"
          unfolding clause_could_propagate_def
          using D_almost_false D_max_lit L_undef by metis

        thus "C c D"
          using D_in D_neq C_least by metis
      next
        assume "ord_res_8_can_factorize N Uer  Γ D"

        then obtain L where
          D_max_lit: "ord_res.is_maximal_lit L D" and
          "¬ (A|∈|atms_of_clss (N |∪| Uer). A t atm_of L  A |∉| trail_atms Γ)" and
          L_undef: "¬ trail_defined_lit Γ L" and
          D_almost_false: "trail_false_cls Γ {#K ∈# D. K  L#}" and
          "is_pos L"
          by (auto elim: ord_res_8_can_factorize.cases ord_res_8_can_produce.cases)

        hence "atm_of L = A"
          using step_hyps
          by (smt (verit, ccfv_SIG) A_in A_least A_undef D_in Γ_lower_set
              linorder_lit.is_maximal_in_mset_iff linorder_trm.antisym_conv3
              linorder_trm.not_in_lower_setI atm_of_in_atms_of_clssI
              trail_defined_lit_iff_trail_defined_atm)

        hence "L = Pos A"
          using is_pos L by (cases L) simp_all

        hence "clause_could_propagate Γ D (Pos A)"
          unfolding clause_could_propagate_def
          using D_almost_false D_max_lit L_undef by metis

        thus "C c D"
          using D_in D_neq C_least by metis
      qed
    qed

    hence "𝒞 = Some C"
      using match_hyps by metis

    define 𝒞' where
      "𝒞' = Some (efac C)"

    have first_step7: "ord_res_7 N (Uer, , Γ, Some C) (Uer, ℱ', Γ, 𝒞')"
      unfolding 𝒞'_def
    proof (rule ord_res_7.factoring)
      show "¬ trail_false_cls Γ C"
        using C_not_false .
    next
      show "ord_res.is_maximal_lit (Pos A) C"
        using C_max_lit .
    next
      show "¬ (Aa|∈|atms_of_clss (N |∪| Uer). Aa t atm_of (Pos A)  Aa |∉| trail_atms Γ)"
        using no_undef_atm_lt_A by simp
    next
      show "¬ trail_defined_lit Γ (Pos A)"
        using Pos_A_undef .
    next
      show "is_pos (Pos A)"
        by simp
    next
      show "trail_false_cls Γ {#K ∈# C. K  Pos A#}"
        using C_almost_false .
    next
      show "¬ ord_res.is_strictly_maximal_lit (Pos A) C"
        using step_hyps by argo
    next
      show "ℱ' = finsert C "
        using step_hyps by argo
    qed

    moreover have "ord_res_7_invars N (Uer, ℱ', Γ, 𝒞')"
      using 𝒞 = Some C first_step7 match_hyps(3) ord_res_7_preserves_invars by blast

    ultimately obtain 𝒞'' where
      following_steps7: "(ord_res_7 N)** (Uer, ℱ', Γ, 𝒞') (Uer, ℱ', Γ, 𝒞'')" and
      no_more_step7: "(𝒞'''. ord_res_7 N (Uer, ℱ', Γ, 𝒞'') (Uer, ℱ', Γ, 𝒞'''))"
      using MAGIC6 by metis

    show ?thesis
    proof (intro exI conjI)
      have "(ord_res_7 N)++ (Uer, , Γ, 𝒞) (Uer, ℱ', Γ, 𝒞'')"
        unfolding 𝒞 = Some C
        using first_step7 following_steps7 by simp

      thus "(constant_context ord_res_7)++ S7 (N, Uer, ℱ', Γ, 𝒞'')"
        unfolding S7_def by (simp add: tranclp_constant_context)

      show "ord_res_7_matches_ord_res_8 (N, Uer, ℱ', Γ, 𝒞'') S8'"
        unfolding S8'_def s8' = (Uer, ℱ', Γ)
      proof (intro ord_res_7_matches_ord_res_8.intros allI)
        show "ord_res_7_invars N (Uer, ℱ', Γ, 𝒞'')"
          using ord_res_7_invars N (Uer, ℱ', Γ, 𝒞') following_steps7
            rtranclp_ord_res_7_preserves_ord_res_7_invars by blast

        show "ord_res_8_invars N (Uer, ℱ', Γ)"
          using invars_s8' step_hyps(1) by blast

        fix C :: "'f gclause"
        show "𝒞'' = Some C  is_least_nonskipped_clause N Uer ℱ' Γ C"
          using MAGIC5 ord_res_7_invars N (Uer, ℱ', Γ, 𝒞'') no_more_step7 by metis
      qed
    qed
  next
    case step_hyps: (resolution E A D Uer' Γ')

    note E_max_lit = ord_res.is_maximal_lit (Neg A) E

    have
      E_in: "E |∈| iefac  |`| (N |∪| Uer)" and
      E_false: "trail_false_cls Γ E" and
      E_least: "F |∈| iefac  |`| (N |∪| Uer). F  E  trail_false_cls Γ F  E c F"
      using step_hyps
      unfolding atomize_conj
      unfolding linorder_cls.is_least_in_ffilter_iff
      by argo

    have "is_least_nonskipped_clause N Uer  Γ E"
      unfolding is_least_nonskipped_clause_def
      unfolding linorder_cls.is_least_in_ffilter_iff
    proof (intro conjI ballI impI)
      show "E |∈| iefac  |`| (N |∪| Uer)"
        using E_in .
    next
      show "trail_false_cls Γ E 
        ord_res_8_can_decide_neg N Uer  Γ E 
        ord_res_8_can_skip_undefined_neg N Uer  Γ E 
        ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ E 
        ord_res_8_can_produce N Uer  Γ E 
        ord_res_8_can_factorize N Uer  Γ E"
        using E_false by argo
    next
      fix F :: "'f gclause"
      assume
        F_in: "F |∈| iefac  |`| (N |∪| Uer)" and
        F_neq: "F  E" and
        D_spec_disj: "trail_false_cls Γ F 
          ord_res_8_can_decide_neg N Uer  Γ F 
          ord_res_8_can_skip_undefined_neg N Uer  Γ F 
          ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ F 
          ord_res_8_can_produce N Uer  Γ F 
          ord_res_8_can_factorize N Uer  Γ F"

      show "E c F"
        using D_spec_disj
      proof (elim disjE)
        assume "trail_false_cls Γ F"
        thus "E c F"
          using E_least F_in F_neq by metis
      next
        assume "ord_res_8_can_decide_neg N Uer  Γ F"
        thus "E c F"
        proof (cases N Uer  Γ F rule: ord_res_8_can_decide_neg.cases)
          case hyps: (1 L' A')
          thus ?thesis
            using no_undef_atom_le_max_lit_if_lt_false_clause[
                OF Γ_lower_set E_in E_false E_max_lit F_in ord_res.is_maximal_lit L' F]
            by (metis (no_types, lifting) F_neq linorder_cls.neq_iff
                linorder_trm.is_least_in_ffilter_iff reflclp_iff)
        qed
      next
        assume "ord_res_8_can_skip_undefined_neg N Uer  Γ F"
        thus "E c F"
        proof (cases N Uer  Γ F rule: ord_res_8_can_skip_undefined_neg.cases)
          case (1 L')
          thus ?thesis
            using no_undef_atom_le_max_lit_if_lt_false_clause[
                OF Γ_lower_set E_in E_false E_max_lit F_in ord_res.is_maximal_lit L' F]
            by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
                linorder_lit.is_maximal_in_mset_iff reflclp_iff
                trail_defined_lit_iff_trail_defined_atm)
        qed
      next
        assume "ord_res_8_can_skip_undefined_pos_ultimate N Uer  Γ F"
        thus "E c F"
        proof (cases N Uer  Γ F rule: ord_res_8_can_skip_undefined_pos_ultimate.cases)
          case (1 L')
          thus ?thesis
            using no_undef_atom_le_max_lit_if_lt_false_clause[
                OF Γ_lower_set E_in E_false E_max_lit F_in ord_res.is_maximal_lit L' F]
            by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
                linorder_lit.is_maximal_in_mset_iff reflclp_iff
                trail_defined_lit_iff_trail_defined_atm)
        qed
      next
        assume "ord_res_8_can_produce N Uer  Γ F"
        thus "E c F"
        proof (cases N Uer  Γ F rule: ord_res_8_can_produce.cases)
          case (1 L')
          thus ?thesis
            using no_undef_atom_le_max_lit_if_lt_false_clause[
                OF Γ_lower_set E_in E_false E_max_lit F_in ord_res.is_maximal_lit L' F]
            by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
                linorder_lit.is_maximal_in_mset_iff reflclp_iff
                trail_defined_lit_iff_trail_defined_atm)
        qed
      next
        assume "ord_res_8_can_factorize N Uer  Γ F"
        thus "E c F"
        proof (cases N Uer  Γ F rule: ord_res_8_can_factorize.cases)
          case (1 L')
          thus ?thesis
            using no_undef_atom_le_max_lit_if_lt_false_clause[
                OF Γ_lower_set E_in E_false E_max_lit F_in ord_res.is_maximal_lit L' F]
            by (metis F_in F_neq atm_of_in_atms_of_clssI linorder_cls.not_less_iff_gr_or_eq
                linorder_lit.is_maximal_in_mset_iff reflclp_iff
                trail_defined_lit_iff_trail_defined_atm)
        qed
      qed
    qed

    hence "𝒞 = Some E"
      using match_hyps by metis

    obtain 𝒞' where first_step7: "ord_res_7 N (Uer, , Γ, Some E) (Uer', , Γ', 𝒞')"
    proof atomize_elim
      have "(Pos A, Some D)  set Γ"
        using map_of Γ (Pos A) = Some (Some D) by (metis map_of_SomeD)

      hence D_almost_false: "trail_false_cls Γ {#K ∈# D. K  Pos A#}"
        using ord_res_7_invars_implies_propagated_clause_almost_false
          ord_res_7_invars N (Uer, , Γ, 𝒞) by metis

      have eres_false: "trail_false_cls Γ (eres D E)"
        unfolding trail_false_cls_def
      proof (intro ballI)
        fix K :: "'f gliteral"
        assume "K ∈# eres D E"
        hence "K ∈# D  K  Pos A  K ∈# E"
          using strong_lit_in_one_of_resolvents_if_in_eres[OF E_max_lit] by simp
        thus "trail_false_lit Γ K"
        proof (elim disjE conjE)
          assume "K ∈# D" and "K  Pos A"
          thus "trail_false_lit Γ K"
            using D_almost_false unfolding trail_false_cls_def by simp
        next
          assume "K ∈# E"
          thus "trail_false_lit Γ K"
            using E_false unfolding trail_false_cls_def by simp
        qed
      qed

      show "𝒞'. ord_res_7 N (Uer, , Γ, Some E) (Uer', , Γ', 𝒞')"
      proof (cases "eres D E = {#}")
        case True
        hence "Ln. K. ord_res.is_maximal_lit K (eres D E)  atm_of K t atm_of (fst Ln)"
          unfolding linorder_lit.is_maximal_in_mset_iff
          by simp
        hence "Γ' = dropWhile (λLn. True) Γ"
          using step_hyps by meson
        hence "Γ' = []"
          by simp
        show ?thesis
        proof (intro exI)
          show "ord_res_7 N (Uer, , Γ, Some E) (Uer', , Γ', Some {#})"
            using ord_res_7.resolution_bot[OF E_false E_max_lit]
              map_of Γ (Pos A) = Some (Some D) Uer' = finsert (eres D E) Uer True Γ' = []
            by simp
        qed
      next
        case False
        then obtain K where eres_max_lit: "ord_res.is_maximal_lit K (eres D E)"
          using linorder_lit.ex_maximal_in_mset by presburger
        hence "Ln. (K. ord_res.is_maximal_lit K (eres D E)  atm_of K t atm_of (fst Ln)) 
          atm_of K t atm_of (fst Ln)"
          by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')
        hence Γ'_eq: "Γ' = dropWhile (λLn. atm_of K t atm_of (fst Ln)) Γ"
          using step_hyps by meson
        show ?thesis
        proof (cases K)
          case (Pos AK)
          hence K_pos: "is_pos K"
            by simp

          then show ?thesis
            using ord_res_7.resolution_pos[OF E_false E_max_lit _ _ _ False Γ'_eq
                eres_max_lit K_pos]
            using step_hyps by fastforce
        next
          case (Neg AK)

          hence K_neg: "is_neg K"
            by simp

          have "trail_false_lit Γ K"
            using eres_false eres_max_lit
            unfolding linorder_lit.is_maximal_in_mset_iff trail_false_cls_def by metis

          hence "opt. (- K, opt)  set Γ"
            unfolding trail_false_lit_def by auto

          moreover have "Lnset Γ. is_neg (fst Ln) = (snd Ln = None)"
            using invars7 by argo

          ultimately obtain C where "(- K, Some C)  set Γ"
            unfolding Neg uminus_Neg by fastforce

          hence "map_of Γ (- K) = Some (Some C)"
          proof (rule map_of_is_SomeI[rotated])
            show "distinct (map fst Γ)"
              using Γ_consistent
              by (metis distinct_lits_if_trail_consistent)
          qed

          then show ?thesis
            using ord_res_7.resolution_neg[OF E_false E_max_lit _ _ _ False Γ'_eq
                eres_max_lit K_neg]
            using step_hyps by fastforce
        qed
      qed
    qed

    moreover have "ord_res_7_invars N (Uer', , Γ', 𝒞')"
      using 𝒞 = Some E first_step7 match_hyps(3) ord_res_7_preserves_invars by blast

    ultimately obtain 𝒞'' where
      following_steps7: "(ord_res_7 N)** (Uer', , Γ', 𝒞') (Uer', , Γ', 𝒞'')" and
      no_more_step7: "(𝒞'''. ord_res_7 N (Uer', , Γ', 𝒞'') (Uer', , Γ', 𝒞'''))"
      using MAGIC6 by metis

    show ?thesis
    proof (intro exI conjI)
      have "(ord_res_7 N)++ (Uer, , Γ, 𝒞) (Uer', , Γ', 𝒞'')"
        unfolding 𝒞 = Some E
        using first_step7 following_steps7 by simp

      thus "(constant_context ord_res_7)++ S7 (N, Uer', , Γ', 𝒞'')"
        unfolding S7_def by (simp add: tranclp_constant_context)

      show "ord_res_7_matches_ord_res_8 (N, Uer', , Γ', 𝒞'') S8'"
        unfolding S8'_def s8' = (Uer', , Γ')
      proof (intro ord_res_7_matches_ord_res_8.intros allI)
        show "ord_res_7_invars N (Uer', , Γ', 𝒞'')"
          using ord_res_7_invars N (Uer', , Γ', 𝒞') following_steps7
            rtranclp_ord_res_7_preserves_ord_res_7_invars by blast

        show "ord_res_8_invars N (Uer', , Γ')"
          using invars_s8' step_hyps(1) by blast

        fix C :: "'f gclause"
        show "𝒞'' = Some C  is_least_nonskipped_clause N Uer'  Γ' C"
          using MAGIC5 ord_res_7_invars N (Uer', , Γ', 𝒞'') no_more_step7 by metis
      qed
    qed
  qed
qed

theorem bisimulation_ord_res_7_ord_res_8:
  defines "match  λ_. ord_res_7_matches_ord_res_8"
  shows "(MATCH :: nat × nat  'f ord_res_7_state  'f ord_res_8_state  bool) .
    bisimulation
      (constant_context ord_res_7) (constant_context ord_res_8)
      ord_res_7_final ord_res_8_final
       MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_7_state ⇒ 'f ord_res_8_state ⇒ bool) ℛfb.
    bisimulation
      (constant_context ord_res_7) ord_res_7_final
      (constant_context ord_res_8) ord_res_8_final
      MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique (constant_context ord_res_7)"
    using right_unique_constant_context right_unique_ord_res_7 by metis
next
  show "right_unique (constant_context ord_res_8)"
    using right_unique_constant_context right_unique_ord_res_8 by metis
next
  show "S. ord_res_7_final S  (S'. constant_context ord_res_7 S S')"
    by (metis finished_def ord_res_7_semantics.final_finished)
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_8 S S')"
    by (metis finished_def ord_res_8_semantics.final_finished)
next
  show "i S7 S8. match i S7 S8  ord_res_7_final S7  ord_res_8_final S8"
    unfolding match_def
    using ord_res_7_final_iff_ord_res_8_final by metis
next
  show "i S7 S8. match i S7 S8 
       safe_state (constant_context ord_res_7) ord_res_7_final S7 
       safe_state (constant_context ord_res_8) ord_res_8_final S8"
  proof (intro allI impI conjI)
    fix i S7 S8
    assume match: "match i S7 S8"
    show "safe_state (constant_context ord_res_7) ord_res_7_final S7"
      using match[unfolded match_def]
      using ord_res_7_safe_state_if_invars
      using ord_res_7_matches_ord_res_8.simps by auto

    show "safe_state (constant_context ord_res_8) ord_res_8_final S8"
      using match[unfolded match_def]
      using ord_res_8_safe_state_if_invars
      using ord_res_7_matches_ord_res_8.simps by auto
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i S7 S8 S8'. match i S7 S8  constant_context ord_res_8 S8 S8' 
    (i' S7'. (constant_context ord_res_7)++ S7 S7'  match i' S7' S8') 
    (i'. match i' S7 S8'  False)"
    unfolding match_def
    using backward_simulation_between_7_and_8 by metis
qed

end


section ‹ORD-RES-9 (factorize when propagating)›

type_synonym 'f ord_res_9_state =
  "'f gclause fset ×'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list"

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_8_matches_ord_res_9 :: "'f ord_res_8_state  'f ord_res_9_state  bool" where
  "ord_res_8_invars N (Uer, , Γ) 
    ord_res_8_matches_ord_res_9 (N, Uer, , Γ) (N, Uer, , Γ)"

lemma ord_res_8_final_iff_ord_res_9_final:
  fixes S8 S9
  assumes match: "ord_res_8_matches_ord_res_9 S8 S9"
  shows "ord_res_8_final S8  ord_res_8_final S9"
  using match
proof (cases S8 S9 rule: ord_res_8_matches_ord_res_9.cases)
  case (1 N Uer  Γ)
  then show ?thesis
    by argo
qed

lemma backward_simulation_between_8_and_9:
  fixes S8 S9 S9'
  assumes match: "ord_res_8_matches_ord_res_9 S8 S9" and step: "constant_context ord_res_9 S9 S9'"
  shows "S8'. (constant_context ord_res_8)++ S8 S8'  ord_res_8_matches_ord_res_9 S8' S9'"
  using match
proof (cases S8 S9 rule: ord_res_8_matches_ord_res_9.cases)
  case match_hyps: (1 N Uer  Γ)

  note S8_def = S8 = (N, Uer, , Γ)
  note S9_def = S9 = (N, Uer, , Γ)
  note invars = ord_res_8_invars N (Uer, , Γ)

  obtain s9' where S9'_def: "S9' = (N, s9')" and step': "ord_res_9 N (Uer, , Γ) s9'"
    using step unfolding S9_def
    using constant_context.cases by blast

  have "ord_res_8 N (Uer, , Γ) s9'  (ord_res_8 N OO ord_res_8 N) (Uer, , Γ) s9'"
    using step' ord_res_9_is_one_or_two_ord_res_9_steps by metis

  hence steps8: "(ord_res_8 N)++ (Uer, , Γ) s9'"
    by auto

  show ?thesis
  proof (intro exI conjI)
    show "(constant_context ord_res_8)++ S8 (N, s9')"
      unfolding S8_def
      using steps8 by (simp add: tranclp_constant_context)
  next
    have "ord_res_8_invars N s9'"
      using invars steps8 tranclp_ord_res_8_preserves_invars by metis

    thus "ord_res_8_matches_ord_res_9 (N, s9') S9'"
      unfolding S9'_def
      by (metis ord_res_8_matches_ord_res_9.intros prod_cases3)
  qed
qed

theorem bisimulation_ord_res_8_ord_res_9:
  defines "match  λ_. ord_res_8_matches_ord_res_9"
  shows "(MATCH :: nat × nat  'f ord_res_8_state  'f ord_res_9_state  bool) .
    bisimulation
      (constant_context ord_res_8) (constant_context ord_res_9)
      ord_res_8_final ord_res_8_final
       MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_8_state ⇒ 'f ord_res_9_state ⇒ bool) ℛfb.
    bisimulation
      (constant_context ord_res_8) ord_res_8_final
      (constant_context ord_res_9) ord_res_8_final
      MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique (constant_context ord_res_8)"
    using right_unique_constant_context right_unique_ord_res_8 by metis
next
  show "right_unique (constant_context ord_res_9)"
    using right_unique_constant_context right_unique_ord_res_9 by metis
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_8 S S')"
    by (metis finished_def ord_res_8_semantics.final_finished)
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_9 S S')"
    by (metis finished_def ord_res_9_semantics.final_finished)
next
  show "i S8 S9. match i S8 S9  ord_res_8_final S8  ord_res_8_final S9"
    unfolding match_def
    using ord_res_8_final_iff_ord_res_9_final by metis
next
  show "i S8 S9. match i S8 S9 
       safe_state (constant_context ord_res_8) ord_res_8_final S8 
       safe_state (constant_context ord_res_9) ord_res_8_final S9"
  proof (intro allI impI conjI)
    fix i S8 S9
    assume match: "match i S8 S9"
    show "safe_state (constant_context ord_res_8) ord_res_8_final S8"
      using match[unfolded match_def]
      using ord_res_8_safe_state_if_invars
      using ord_res_8_matches_ord_res_9.simps by auto

    show "safe_state (constant_context ord_res_9) ord_res_8_final S9"
      using match[unfolded match_def]
      using ord_res_9_safe_state_if_invars
      using ord_res_8_matches_ord_res_9.simps by auto
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i S8 S9 S9'. match i S8 S9  constant_context ord_res_9 S9 S9' 
    (i' S8'. (constant_context ord_res_8)++ S8 S8'  match i' S8' S9') 
    (i'. match i' S8 S9'  False)"
    unfolding match_def
    using backward_simulation_between_8_and_9 by metis
qed

end


section ‹ORD-RES-10 (propagate iff a conflict is produced)›

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_9_matches_ord_res_10 :: "'f ord_res_9_state  'f ord_res_10_state  bool" where
  "ord_res_8_invars N (Uer, , Γ9) 
    ord_res_10_invars N (Uer, , Γ10) 
    list_all2 (λx y. fst x = fst y) Γ9 Γ10 
    list_all2 (λx y. snd y  None  x = y) Γ9 Γ10 
    ord_res_9_matches_ord_res_10 (N, Uer, , Γ9) (N, Uer, , Γ10)"

lemma ord_res_9_final_iff_ord_res_10_final:
  fixes S9 S10
  assumes match: "ord_res_9_matches_ord_res_10 S9 S10"
  shows "ord_res_8_final S9  ord_res_8_final S10"
  using match
proof (cases S9 S10 rule: ord_res_9_matches_ord_res_10.cases)
  case match_hyps: (1 N Uer  Γ9 Γ10)
  then show ?thesis
    using trail_atms_eq_trail_atms_if_same_lits[OF list_all2 (λx y. fst x = fst y) Γ9 Γ10]
    using trail_false_cls_eq_trail_false_cls_if_same_lits[OF list_all2 (λx y. fst x = fst y) Γ9 Γ10]
    unfolding ord_res_8_final.simps
    by simp
qed

lemma backward_simulation_between_9_and_10:
  fixes S9 S10 S10'
  assumes
    match: "ord_res_9_matches_ord_res_10 S9 S10" and
    step: "constant_context ord_res_10 S10 S10'"
  shows "S9'. (constant_context ord_res_9)++ S9 S9'  ord_res_9_matches_ord_res_10 S9' S10'"
  using match
proof (cases S9 S10 rule: ord_res_9_matches_ord_res_10.cases)
  case match_hyps: (1 N Uer  Γ9 Γ10)

  note S9_def = S9 = (N, Uer, , Γ9)
  note S10_def = S10 = (N, Uer, , Γ10)
  note invars9 = ord_res_8_invars N (Uer, , Γ9)
  note invars10 = ord_res_10_invars N (Uer, , Γ10)

  have "trail_atms Γ9 = trail_atms Γ10"
    using list_all2 (λx y. fst x = fst y) Γ9 Γ10 trail_atms_eq_trail_atms_if_same_lits
    by metis

  have "trail_false_lit Γ9 = trail_false_lit Γ10"
    using list_all2 (λx y. fst x = fst y) Γ9 Γ10 trail_false_lit_eq_trail_false_lit_if_same_lits
    by metis

  have "trail_false_cls Γ9 = trail_false_cls Γ10"
    using list_all2 (λx y. fst x = fst y) Γ9 Γ10 trail_false_cls_eq_trail_false_cls_if_same_lits
    by metis

  have "trail_defined_lit Γ9 = trail_defined_lit Γ10"
    using list_all2 (λx y. fst x = fst y) Γ9 Γ10
      trail_defined_lit_eq_trail_defined_lit_if_same_lits by metis

  have "trail_defined_cls Γ9 = trail_defined_cls Γ10"
    using list_all2 (λx y. fst x = fst y) Γ9 Γ10
      trail_defined_cls_eq_trail_defined_cls_if_same_lits by metis

  have "clause_could_propagate Γ9 = clause_could_propagate Γ10"
    unfolding clause_could_propagate_def
    unfolding trail_defined_lit Γ9 = trail_defined_lit Γ10
    unfolding trail_false_cls Γ9 = trail_false_cls Γ10 ..

  have Γ9_sorted: "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) Γ9"
    using invars9[unfolded ord_res_8_invars_def trail_is_sorted_def, simplified] by argo

  obtain s10' where "S10' = (N, s10')" and step10: "ord_res_10 N (Uer, , Γ10) s10'"
    using step unfolding S10_def by (auto elim: constant_context.cases)

  show ?thesis
    using step10
  proof (cases N "(Uer, , Γ10)" s10' rule: ord_res_10.cases)
    case step_hyps: (decide_neg A Γ10')

    define Γ9' where
      "Γ9' = (Neg A, None) # Γ9"

    show ?thesis
    proof (intro exI conjI)
      have step9: "ord_res_9 N (Uer, , Γ9) (Uer, , Γ9')"
      proof (rule ord_res_9.decide_neg)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer). trail_false_cls Γ9 C)"
          using step_hyps trail_false_cls Γ9 = trail_false_cls Γ10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer). A1 |∈| trail_atms Γ9. A1 t A2|} A"
          using step_hyps trail_atms Γ9 = trail_atms Γ10 by metis
      next
        show "¬ (C|∈|iefac  |`| (N |∪| Uer). clause_could_propagate Γ9 C (Pos A))"
          using step_hyps clause_could_propagate Γ9 = clause_could_propagate Γ10 by metis
      next
        show "Γ9' = (Neg A, None) # Γ9"
          using Γ9'_def .
      qed

      thus "(constant_context ord_res_9)++ S9 (N, Uer, , Γ9')"
        unfolding S9_def by (auto intro: constant_context.intros)

      show "ord_res_9_matches_ord_res_10 (N, Uer, , Γ9') S10'"
        unfolding S10' = (N, s10') s10' = (Uer, , Γ10')
      proof (rule ord_res_9_matches_ord_res_10.intros)
        show "ord_res_8_invars N (Uer, , Γ9')"
          using invars9 step9 ord_res_9_preserves_invars by metis 
      next
        show "ord_res_10_invars N (Uer, , Γ10')"
          using invars10 step10 ord_res_10_preserves_invars s10' = (Uer, , Γ10') by metis
      next
        show "list_all2 (λx y. fst x = fst y) Γ9' Γ10'"
          unfolding Γ9' = (Neg A, None) # Γ9 Γ10' = (Neg A, None) # Γ10
          using list_all2 (λx y. fst x = fst y) Γ9 Γ10 by simp
      next
        show "list_all2 (λx y. snd y  None  x = y) Γ9' Γ10'"
          unfolding Γ9' = (Neg A, None) # Γ9 Γ10' = (Neg A, None) # Γ10
          using list_all2 (λx y. snd y  None  x = y) Γ9 Γ10
          by simp
      qed
    qed
  next
    case step_hyps: (decide_pos A C Γ10' ℱ')

    define Γ9' where
      "Γ9' = (Pos A, Some (efac C)) # Γ9"
    
    show ?thesis
    proof (intro exI conjI)
      have step9: "ord_res_9 N (Uer, , Γ9) (Uer, ℱ', Γ9')"
      proof (rule ord_res_9.propagate)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer). trail_false_cls Γ9 C)"
          using step_hyps trail_false_cls Γ9 = trail_false_cls Γ10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer). A1|∈|trail_atms Γ9. A1 t A2|} A"
          using step_hyps trail_atms Γ9 = trail_atms Γ10 by metis
      next
        show "linorder_cls.is_least_in_fset
          {|C |∈| iefac  |`| (N |∪| Uer). clause_could_propagate Γ9 C (Pos A)|} C"
          using step_hyps clause_could_propagate Γ9 = clause_could_propagate Γ10 by metis
      next
        show "Γ9' = (Pos A, Some (efac C)) # Γ9"
          using Γ9'_def .
      next
        show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then  else finsert C )"
          using step_hyps by argo
      qed

      thus "(constant_context ord_res_9)++ S9 (N, Uer, ℱ', Γ9')"
        unfolding S9_def by (auto intro: constant_context.intros)

      show "ord_res_9_matches_ord_res_10 (N, Uer, ℱ', Γ9') S10'"
        unfolding S10' = (N, s10') s10' = (Uer, ℱ', Γ10')
      proof (rule ord_res_9_matches_ord_res_10.intros)
        show "ord_res_8_invars N (Uer, ℱ', Γ9')"
          using invars9 step9 ord_res_9_preserves_invars by metis 
      next
        show "ord_res_10_invars N (Uer, ℱ', Γ10')"
          using invars10 step10 ord_res_10_preserves_invars s10' = (Uer, ℱ', Γ10') by metis
      next
        show "list_all2 (λx y. fst x = fst y) Γ9' Γ10'"
          unfolding Γ9' = (Pos A, Some (efac C)) # Γ9 Γ10' = (Pos A, None) # Γ10
          using list_all2 (λx y. fst x = fst y) Γ9 Γ10 by simp
      next
        show "list_all2 (λx y. snd y  None  x = y) Γ9' Γ10'"
          unfolding Γ9' = (Pos A, Some (efac C)) # Γ9 Γ10' = (Pos A, None) # Γ10
          using list_all2 (λx y. snd y  None  x = y) Γ9 Γ10
          by simp
      qed
    qed
  next
    case step_hyps: (propagate A C Γ10' ℱ')

    define Γ9' where
      "Γ9' = (Pos A, Some (efac C)) # Γ9"
    
    show ?thesis
    proof (intro exI conjI)
      have step9: "ord_res_9 N (Uer, , Γ9) (Uer, ℱ', Γ9')"
      proof (rule ord_res_9.propagate)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer). trail_false_cls Γ9 C)"
          using step_hyps trail_false_cls Γ9 = trail_false_cls Γ10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer). A1|∈|trail_atms Γ9. A1 t A2|} A"
          using step_hyps trail_atms Γ9 = trail_atms Γ10 by metis
      next
        show "linorder_cls.is_least_in_fset
          {|C |∈| iefac  |`| (N |∪| Uer). clause_could_propagate Γ9 C (Pos A)|} C"
          using step_hyps clause_could_propagate Γ9 = clause_could_propagate Γ10 by metis
      next
        show "Γ9' = (Pos A, Some (efac C)) # Γ9"
          using Γ9'_def .
      next
        show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then  else finsert C )"
          using step_hyps by argo
      qed

      thus "(constant_context ord_res_9)++ S9 (N, Uer, ℱ', Γ9')"
        unfolding S9_def by (auto intro: constant_context.intros)

      show "ord_res_9_matches_ord_res_10 (N, Uer, ℱ', Γ9') S10'"
        unfolding S10' = (N, s10') s10' = (Uer, ℱ', Γ10')
      proof (rule ord_res_9_matches_ord_res_10.intros)
        show "ord_res_8_invars N (Uer, ℱ', Γ9')"
          using invars9 step9 ord_res_9_preserves_invars by metis 
      next
        show "ord_res_10_invars N (Uer, ℱ', Γ10')"
          using invars10 step10 ord_res_10_preserves_invars s10' = (Uer, ℱ', Γ10') by metis
      next
        show "list_all2 (λx y. fst x = fst y) Γ9' Γ10'"
          unfolding Γ9' = (Pos A, Some (efac C)) # Γ9 Γ10' = (Pos A, Some (efac C)) # Γ10
          using list_all2 (λx y. fst x = fst y) Γ9 Γ10 by simp
      next
        show "list_all2 (λx y. snd y  None  x = y) Γ9' Γ10'"
          unfolding Γ9' = (Pos A, Some (efac C)) # Γ9 Γ10' = (Pos A, Some (efac C)) # Γ10
          using list_all2 (λx y. snd y  None  x = y) Γ9 Γ10
          by simp
      qed
    qed
  next
    case step_hyps: (resolution D A C Uer' Γ10')

    have "Ln Γ'. Γ10 = Ln # Γ' 
      (snd Ln  None) = fBex (iefac  |`| (N |∪| Uer)) (trail_false_cls Γ10) 
      (xset Γ'. snd x = None)"
      using invars10 by (simp add: ord_res_10_invars.simps)

    then obtain Γ10'' where "Γ10 = (Pos A, Some C) # Γ10''"
      using map_of Γ10 (Pos A) = Some (Some C)
      by (metis list.set_cases map_of_SomeD not_Some_eq snd_conv)

    then obtain Γ9'' where "Γ9 = (Pos A, Some C) # Γ9''"
      using list_all2 (λx y. snd y  None  x = y) Γ9 Γ10
      by (smt (verit, best) list_all2_Cons2 option.discI snd_conv)

    define Γ9' where
      "Γ9' = dropWhile (λLn. K. ord_res.is_maximal_lit K (eres C D) 
        atm_of K t atm_of (fst Ln)) Γ9"
    
    show ?thesis
    proof (intro exI conjI)
      have step9: "ord_res_9 N (Uer, , Γ9) (Uer', , Γ9')"
      proof (rule ord_res_9.resolution)
        show "linorder_cls.is_least_in_fset (ffilter (trail_false_cls Γ9) (iefac  |`| (N |∪| Uer))) D"
          using step_hyps trail_false_cls Γ9 = trail_false_cls Γ10 by argo
      next
        show "ord_res.is_maximal_lit (Neg A) D"
          using step_hyps by argo
      next
        show "map_of Γ9 (Pos A) = Some (Some C)"
          unfolding Γ9 = (Pos A, Some C) # Γ9'' by simp
      next
        show "Uer' = finsert (eres C D) Uer"
          using step_hyps by argo
      next
        show "Γ9' = dropWhile (λLn. K. ord_res.is_maximal_lit K (eres C D) 
          atm_of K t atm_of (fst Ln)) Γ9"
          using Γ9'_def .
      qed

      thus "(constant_context ord_res_9)++ S9 (N, Uer', , Γ9')"
        unfolding S9_def by (auto intro: constant_context.intros)

      show "ord_res_9_matches_ord_res_10 (N, Uer', , Γ9') S10'"
        unfolding S10' = (N, s10') s10' = (Uer', , Γ10')
      proof (rule ord_res_9_matches_ord_res_10.intros)
        show "ord_res_8_invars N (Uer', , Γ9')"
          using invars9 step9 ord_res_9_preserves_invars by metis 
      next
        show "ord_res_10_invars N (Uer', , Γ10')"
          using invars10 step10 ord_res_10_preserves_invars s10' = (Uer', , Γ10') by metis
      next
        define P :: "'f gterm literal × 'f gterm literal multiset option  bool" where
          "P  λLn. K. ord_res.is_maximal_lit K (eres C D)  atm_of K t atm_of (fst Ln)"

        have "length (takeWhile P Γ9) = length (takeWhile P Γ10)"
          using list_all2 (λx y. fst x = fst y) Γ9 Γ10
        proof (induction Γ9 Γ10 rule: list.rel_induct)
          case Nil
          show ?case
            by simp
        next
          case (Cons x xs y ys)
          then show ?case
            by (simp add: P_def)
        qed

        moreover have "Γ9 = takeWhile P Γ9 @ Γ9'"
          unfolding takeWhile_dropWhile_id
          unfolding P_def Γ9' = dropWhile _ Γ9 by simp

        moreover have "Γ10 = takeWhile P Γ10 @ Γ10'"
          unfolding takeWhile_dropWhile_id
          unfolding P_def Γ10' = dropWhile _ Γ10 by simp

        ultimately have "Q. list_all2 Q Γ9 Γ10 
          (list_all2 Q (takeWhile P Γ9) (takeWhile P Γ10)  list_all2 Q Γ9' Γ10')"
          using list_all2_append by metis

        thus
          "list_all2 (λx y. fst x = fst y) Γ9' Γ10'"
          "list_all2 (λx y. snd y  None  x = y) Γ9' Γ10'"
          unfolding atomize_conj
          using list_all2 (λx y. fst x = fst y) Γ9 Γ10
          using list_all2 (λx y. snd y  None  x = y) Γ9 Γ10
          by (simp only:)
      qed
    qed
  qed
qed

theorem bisimulation_ord_res_9_ord_res_10:
  defines "match  λ_. ord_res_9_matches_ord_res_10"
  shows "(MATCH :: nat × nat  'f ord_res_8_state  'f ord_res_9_state  bool) .
    bisimulation
      (constant_context ord_res_9) (constant_context ord_res_10)
      ord_res_8_final ord_res_8_final
       MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_8_state ⇒ 'f ord_res_9_state ⇒ bool) ℛfb.
    bisimulation
      (constant_context ord_res_9) ord_res_8_final
      (constant_context ord_res_10) ord_res_8_final
      MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_backward_simulation)
  show "right_unique (constant_context ord_res_9)"
    using right_unique_constant_context right_unique_ord_res_9 by metis
next
  show "right_unique (constant_context ord_res_10)"
    using right_unique_constant_context right_unique_ord_res_10 by metis
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_9 S S')"
    by (metis finished_def ord_res_9_semantics.final_finished)
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_10 S S')"
    by (metis finished_def ord_res_10_semantics.final_finished)
next
  show "i S9 S10. match i S9 S10  ord_res_8_final S9  ord_res_8_final S10"
    unfolding match_def
    using ord_res_9_final_iff_ord_res_10_final by metis
next
  show "i S9 S10. match i S9 S10 
       safe_state (constant_context ord_res_9) ord_res_8_final S9 
       safe_state (constant_context ord_res_10) ord_res_8_final S10"
  proof (intro allI impI conjI)
    fix i S9 S10
    assume match: "match i S9 S10"
    show "safe_state (constant_context ord_res_9) ord_res_8_final S9"
      using match[unfolded match_def]
      using ord_res_9_safe_state_if_invars
      using ord_res_9_matches_ord_res_10.simps by auto

    show "safe_state (constant_context ord_res_10) ord_res_8_final S10"
      using match[unfolded match_def]
      using ord_res_10_safe_state_if_invars
      using ord_res_9_matches_ord_res_10.simps by auto
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i S9 S10 S10'. match i S9 S10  constant_context ord_res_10 S10 S10' 
    (i' S9'. (constant_context ord_res_9)++ S9 S9'  match i' S9' S10') 
    (i'. match i' S9 S10'  False)"
    unfolding match_def
    using backward_simulation_between_9_and_10 by metis
qed

end


section ‹ORD-RES-11 (SCL strategy)›

context simulation_SCLFOL_ground_ordered_resolution begin

inductive ord_res_10_matches_ord_res_11 :: "'f ord_res_10_state  'f ord_res_11_state  bool" where
  "ord_res_10_invars N (Uer10, , Γ) 
    ord_res_11_invars N (Uer11, , Γ, 𝒞) 
    Uer11 = Uer10 - {|{#}|} 
    if {#} |∈| iefac  |`| (N |∪| Uer10) then Γ = []  𝒞 = Some {#} else 𝒞 = None 
    ord_res_10_matches_ord_res_11 (N, Uer10, , Γ) (N, Uer11, , Γ, 𝒞)"

lemma ord_res_10_final_iff_ord_res_11_final:
  fixes S10 S11
  assumes match: "ord_res_10_matches_ord_res_11 S10 S11"
  shows "ord_res_8_final S10  ord_res_11_final S11"
  using match
proof (cases S10 S11 rule: ord_res_10_matches_ord_res_11.cases)
  case match_hyps: (1 N Uer10  Γ Uer11 𝒞)
  show ?thesis
  proof (rule iffI)
    assume "ord_res_8_final S10"
    thus "ord_res_11_final S11"
      unfolding S10 = _
    proof (cases "(N, Uer10, , Γ)" rule: ord_res_8_final.cases)
      case model_found
      hence "{#} |∉| iefac  |`| (N |∪| Uer10)"
        using trail_false_cls_mempty by blast
      hence "𝒞 = None"
        using match_hyps by argo
      moreover have "Uer11 = Uer10"
        using match_hyps
        by (metis {#} |∉| iefac  |`| (N |∪| Uer10) fimage_eqI finsert_fminus1 finsert_iff
            fminus_finsert_absorb funionI2 iefac_mempty)
      ultimately show ?thesis
        unfolding S11 = _
        using model_found
        using ord_res_11_final.model_found
        by metis
    next
      case contradiction_found
      hence "Γ = []  𝒞 = Some {#}"
        using match_hyps by argo
      thus ?thesis
        unfolding S11 = _
        using ord_res_11_final.contradiction_found by metis
    qed
  next
    assume "ord_res_11_final S11"
    thus "ord_res_8_final S10"
      unfolding S11 = _
    proof (cases "(N, Uer11, , Γ, 𝒞)" rule: ord_res_11_final.cases)
      case model_found
      show ?thesis
        unfolding S10 = _
      proof (rule ord_res_8_final.model_found)
        show "¬ (A |∈| atms_of_clss (N |∪| Uer10). A |∉| trail_atms Γ)"
          by (metis (no_types, lifting) fimage_iff fminus_finsert_absorb fminus_idemp funionCI
              iefac_mempty local.model_found(1,2) match_hyps(5,6) option.simps(3))
      next
        show "¬ (C |∈| iefac  |`| (N |∪| Uer10). trail_false_cls Γ C)"
          by (metis finsertCI finsert_fminus1 fminus_finsert_absorb funionI2 iefac_mempty
              local.model_found(1,3) match_hyps(5,6) option.simps(3) rev_fimage_eqI)
      qed
    next
      case contradiction_found
      show ?thesis
        unfolding S10 = _
      proof (rule ord_res_8_final.contradiction_found)
        show "{#} |∈| iefac  |`| (N |∪| Uer10)"
          using match_hyps contradiction_found
          by auto
      qed
    qed
  qed
qed

lemma forward_simulation_between_10_and_11:
  fixes S10 S11 S10'
  assumes
    match: "ord_res_10_matches_ord_res_11 S10 S11" and
    step: "constant_context ord_res_10 S10 S10'"
  shows "S11'. (constant_context ord_res_11)++ S11 S11'  ord_res_10_matches_ord_res_11 S10' S11'"
  using match
proof (cases S10 S11 rule: ord_res_10_matches_ord_res_11.cases)
  case match_hyps: (1 N Uer10  Γ Uer11 𝒞)

  note S10_def = S10 = (N, Uer10, , Γ)
  note S11_def = S11 = (N, Uer11, , Γ, 𝒞)
  note invars10 = ord_res_10_invars N (Uer10, , Γ)
  note invars11 = ord_res_11_invars N (Uer11, , Γ, 𝒞)

  have mempty_not_in_if_no_false_cls: "{#} |∉| iefac  |`| (N |∪| Uer10)"
    if "¬ fBex (iefac  |`| (N |∪| Uer10)) (trail_false_cls Γ)"
    using that by force

  have 𝒞_eq_None_if_no_false_cls: "𝒞 = None"
    if "¬ fBex (iefac  |`| (N |∪| Uer10)) (trail_false_cls Γ)"
    using match_hyps mempty_not_in_if_no_false_cls[OF that] by argo

  have mempty_not_in_if: "{#} |∉| N |∪| Uer10"
    if "¬ fBex (iefac  |`| (N |∪| Uer10)) (trail_false_cls Γ)"
    using that
    by (metis (no_types, opaque_lifting) fimageI iefac_mempty trail_false_cls_mempty)

  have Uer11_eq_Uer10_if: "Uer11 = Uer10"
    if "¬ fBex (iefac  |`| (N |∪| Uer10)) (trail_false_cls Γ)"
    using mempty_not_in_if[OF that] Uer11 = Uer10 |-| {|{#}|}
    by (metis (no_types, opaque_lifting) finsertI1 finsert_ident fminusD2 funionCI
        funion_fempty_right funion_finsert_right funion_fminus_cancel2)

  obtain s10' where "S10' = (N, s10')" and step10: "ord_res_10 N (Uer10, , Γ) s10'"
    using step unfolding S10_def by (auto elim: constant_context.cases)

  show ?thesis
    using step10
  proof (cases N "(Uer10, , Γ)" s10' rule: ord_res_10.cases)
    case step_hyps: (decide_neg A Γ')

    have "𝒞 = None"
      using step_hyps 𝒞_eq_None_if_no_false_cls by argo

    have "{#} |∉| N |∪| Uer10"
      using step_hyps mempty_not_in_if by argo

    have "Uer11 = Uer10"
      using step_hyps Uer11_eq_Uer10_if by argo

    show ?thesis
    proof (intro exI conjI)
      have step11: "ord_res_11 N (Uer11, , Γ, None) (Uer11, , Γ', None)"
      proof (rule ord_res_11.decide_neg)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer11). trail_false_cls Γ C)"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer11). A1|∈|trail_atms Γ. A1 t A2|} A"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "¬ (C|∈|iefac  |`| (N |∪| Uer11). clause_could_propagate Γ C (Pos A))"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "Γ' = (Neg A, None) # Γ"
          using step_hyps by argo
      qed

      thus "(constant_context ord_res_11)++ S11 (N, Uer11, , Γ', None)"
        unfolding S11_def 𝒞 = None by (auto intro: constant_context.intros)

      show "ord_res_10_matches_ord_res_11 S10' (N, Uer11, , Γ', None)"
        unfolding S10' = (N, s10') s10' = _
      proof (rule ord_res_10_matches_ord_res_11.intros)
        show "ord_res_10_invars N (Uer10, , Γ')"
          using step10 s10' = _ invars10 ord_res_10_preserves_invars by metis
      next
        show "ord_res_11_invars N (Uer11, , Γ', None)"
          using step11 invars11 𝒞 = None ord_res_11_preserves_invars by metis
      next
        show "Uer11 = Uer10 |-| {|{#}|}"
          unfolding Uer11 = Uer10
          using {#} |∉| N |∪| Uer10 by simp
      next
        have "{#} |∉| iefac  |`| (N |∪| Uer10)"
          using {#} |∉| N |∪| Uer10 by (simp add: iefac_def)
        thus "if {#} |∈| iefac  |`| (N |∪| Uer10) then Γ' = []  None = Some {#} else None = None"
          by argo
      qed
    qed
  next
    case step_hyps: (decide_pos A C Γ' ℱ')
    
    have "𝒞 = None"
      using step_hyps 𝒞_eq_None_if_no_false_cls by argo

    have "{#} |∉| N |∪| Uer10"
      using step_hyps mempty_not_in_if by argo

    have "Uer11 = Uer10"
      using step_hyps Uer11_eq_Uer10_if by argo

    show ?thesis
    proof (intro exI conjI)
      have step11: "ord_res_11 N (Uer11, , Γ, None) (Uer11, ℱ', Γ', None)"
      proof (rule ord_res_11.decide_pos)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer11). trail_false_cls Γ C)"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer11). A1|∈|trail_atms Γ. A1 t A2|} A"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "linorder_cls.is_least_in_fset
          {|C |∈| iefac  |`| (N |∪| Uer11). clause_could_propagate Γ C (Pos A)|} C"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "Γ' = (Pos A, None) # Γ"
          using step_hyps by argo
      next
        show "¬ (C|∈|iefac  |`| (N |∪| Uer11). trail_false_cls Γ' C)"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then  else finsert C )"
          using step_hyps by argo
      qed

      thus "(constant_context ord_res_11)++ S11 (N, Uer11, ℱ', Γ', None)"
        unfolding S11_def 𝒞 = None by (auto intro: constant_context.intros)

      show "ord_res_10_matches_ord_res_11 S10' (N, Uer11, ℱ', Γ', None)"
        unfolding S10' = (N, s10') s10' = _
      proof (rule ord_res_10_matches_ord_res_11.intros)
        show "ord_res_10_invars N (Uer10, ℱ', Γ')"
          using step10 s10' = _ invars10 ord_res_10_preserves_invars by metis
      next
        show "ord_res_11_invars N (Uer11, ℱ', Γ', None)"
          using step11 invars11 𝒞 = None ord_res_11_preserves_invars by metis
      next
        show "Uer11 = Uer10 |-| {|{#}|}"
          unfolding Uer11 = Uer10
          using {#} |∉| N |∪| Uer10 by simp
      next
        have "{#} |∉| iefac ℱ' |`| (N |∪| Uer10)"
          using {#} |∉| N |∪| Uer10 by (simp add: iefac_def)
        thus "if {#} |∈| iefac ℱ' |`| (N |∪| Uer10) then Γ' = []  None = Some {#} else None = None"
          by argo
      qed
    qed
  next
    case step_hyps: (propagate A C Γ' ℱ')

    have "𝒞 = None"
      using step_hyps 𝒞_eq_None_if_no_false_cls by argo

    have "{#} |∉| N |∪| Uer10"
      using step_hyps mempty_not_in_if by argo

    have "Uer11 = Uer10"
      using step_hyps Uer11_eq_Uer10_if by argo

    show ?thesis
    proof (intro exI conjI)
      have step11: "ord_res_11 N (Uer11, , Γ, None) (Uer11, ℱ', Γ', None)"
      proof (rule ord_res_11.propagate)
        show "¬ (C|∈|iefac  |`| (N |∪| Uer11). trail_false_cls Γ C)"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "linorder_trm.is_least_in_fset
          {|A2 |∈| atms_of_clss (N |∪| Uer11). A1|∈|trail_atms Γ. A1 t A2|} A"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "linorder_cls.is_least_in_fset
          {|C |∈| iefac  |`| (N |∪| Uer11). clause_could_propagate Γ C (Pos A)|} C"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "Γ' = (Pos A, Some (efac C)) # Γ"
          using step_hyps by argo
      next
        show "C|∈|iefac  |`| (N |∪| Uer11). trail_false_cls Γ' C"
          using step_hyps unfolding Uer11 = Uer10 by argo
      next
        show "ℱ' = (if ord_res.is_strictly_maximal_lit (Pos A) C then  else finsert C )"
          using step_hyps by argo
      qed

      thus "(constant_context ord_res_11)++ S11 (N, Uer11, ℱ', Γ', None)"
        unfolding S11_def 𝒞 = None by (auto intro: constant_context.intros)

      show "ord_res_10_matches_ord_res_11 S10' (N, Uer11, ℱ', Γ', None)"
        unfolding S10' = (N, s10') s10' = _
      proof (rule ord_res_10_matches_ord_res_11.intros)
        show "ord_res_10_invars N (Uer10, ℱ', Γ')"
          using step10 s10' = _ invars10 ord_res_10_preserves_invars by metis
      next
        show "ord_res_11_invars N (Uer11, ℱ', Γ', None)"
          using step11 invars11 𝒞 = None ord_res_11_preserves_invars by metis
      next
        show "Uer11 = Uer10 |-| {|{#}|}"
          unfolding Uer11 = Uer10
          using {#} |∉| N |∪| Uer10 by simp
      next
        have "{#} |∉| iefac ℱ' |`| (N |∪| Uer10)"
          using {#} |∉| N |∪| Uer10 by (simp add: iefac_def)
        thus "if {#} |∈| iefac ℱ' |`| (N |∪| Uer10) then Γ' = []  None = Some {#} else None = None"
          by argo
      qed
    qed
  next
    case step_hyps: (resolution D A C Uer10' Γ')

    note D_max_lit = ord_res.is_maximal_lit (Neg A) D

    have "{#} |∉| iefac  |`| (N |∪| Uer10)"
      using linorder_cls.is_least_in_fset _ D linorder_lit.is_maximal_in_mset D _
      unfolding linorder_cls.is_least_in_ffilter_iff linorder_lit.is_maximal_in_mset_iff
      by (metis (no_types, lifting) empty_iff linorder_cls.leD mempty_lesseq_cls set_mset_empty
          trail_false_cls_mempty)

    have "𝒞 = None"
      using match_hyps {#} |∉| iefac  |`| (N |∪| Uer10) by argo

    have "Uer11 = Uer10"
      using match_hyps {#} |∉| iefac  |`| (N |∪| Uer10) by force

    have step11_conf: "ord_res_11 N (Uer11, , Γ, None) (Uer11, , Γ, Some D)"
    proof (rule ord_res_11.conflict)
      show "linorder_cls.is_least_in_fset
        (ffilter (trail_false_cls Γ) (iefac  |`| (N |∪| Uer11))) D"
        using step_hyps unfolding Uer11 = Uer10 by argo
    qed

    have Γ_spec: "Ln Γ'. Γ = Ln # Γ' 
      (snd Ln  None) = fBex (iefac  |`| (N |∪| Uer10)) (trail_false_cls Γ) 
      (xset Γ'. snd x = None)"
      using invars10 by (simp add: ord_res_10_invars.simps)

    then obtain Γ''' where "Γ = (Pos A, Some C) # Γ'''"
      using map_of Γ (Pos A) = Some (Some C)
      by (metis list.set_cases map_of_SomeD not_Some_eq snd_conv)

    have C_max_lit: "linorder_lit.is_greatest_in_mset C (Pos A)"
      using invars10 by (simp add: ord_res_10_invars.simps Γ = (Pos A, Some C) # Γ''')

    have "ord_res.ground_resolution D C ((D - {#Neg A#}) + (C - {#Pos A#}))"
    proof (rule ord_res.ground_resolutionI)
      show "D = add_mset (Neg A) (D - {#Neg A#})"
        using D_max_lit unfolding linorder_lit.is_maximal_in_mset_iff by simp
    next
      show "C = add_mset (Pos A) (C - {#Pos A#})"
        using C_max_lit unfolding linorder_lit.is_greatest_in_mset_iff by simp
    next
      show "C c D"
        using C_max_lit D_max_lit
        by (simp add: clause_lt_clause_if_max_lit_comp
            linorder_lit.is_greatest_in_mset_iff_is_maximal_and_count_eq_one)
    next
      show "{#} = {#}  ord_res.is_maximal_lit (Neg A) D  Neg A ∈# {#}"
        using D_max_lit by argo
    next
      show "{#} = {#}"
        by argo
    next
      show "ord_res.is_strictly_maximal_lit (Pos A) C"
        using C_max_lit .
    next
      show "remove1_mset (Neg A) D + remove1_mset (Pos A) C = (D - {#Neg A#}) + (C - {#Pos A#})"
        ..
    qed

    hence "eres C D  D"
      unfolding eres_ident_iff not_not ground_resolution_def by metis

    have D_false: "trail_false_cls Γ D"
      using step_hyps unfolding linorder_cls.is_least_in_ffilter_iff by argo

    have "clause_could_propagate Γ''' C (Pos A)"
      using invars10 Γ = (Pos A, Some C) # Γ''' by (simp add: ord_res_10_invars.simps)

    hence "trail_false_cls Γ''' {#L ∈# C. L  Pos A#}"
      unfolding clause_could_propagate_def by argo

    hence C_almost_false: "trail_false_cls Γ {#L ∈# C. L  Pos A#}"
      unfolding Γ = (Pos A, Some C) # Γ'''
      by (meson suffix_ConsD suffix_order.dual_order.refl trail_false_cls_if_trail_false_suffix)

    have eres_false: "trail_false_cls Γ (eres C D)"
      unfolding trail_false_cls_def
    proof (intro ballI)
      fix K :: "'f gliteral"
      assume "K ∈# eres C D"
      hence "K ∈# C  K  Pos A  K ∈# D"
        using strong_lit_in_one_of_resolvents_if_in_eres[OF D_max_lit] by simp
      thus "trail_false_lit Γ K"
      proof (elim disjE conjE)
        assume "K ∈# C" and "K  Pos A"
        thus "trail_false_lit Γ K"
          using C_almost_false unfolding trail_false_cls_def by simp
      next
        assume "K ∈# D"
        thus "trail_false_lit Γ K"
          using D_false unfolding trail_false_cls_def by simp
      qed
    qed

    have "Ln  set Γ. C. snd Ln = Some C  ord_res.is_strictly_maximal_lit (fst Ln) C"
      using invars10 by (simp add: ord_res_10_invars.simps)

    hence C_max_lit: "ord_res.is_strictly_maximal_lit (Pos A) C"
      unfolding Γ = (Pos A, Some C) # Γ''' by simp

    have steps11_reso: "(ord_res_11 N)** (Uer11, , Γ, Some D) (Uer11, , Γ, Some (eres C D))"
      unfolding Γ = (Pos A, Some C) # Γ'''
      using C_max_lit rtrancl_ord_res_11_all_resolution_steps by metis

    define strict_P :: "'f gterm literal × 'f gterm literal multiset option  bool" where
      "strict_P  λLn. K. ord_res.is_maximal_lit K (eres C D)  atm_of K t atm_of (fst Ln)"

    have "K ∈# eres C D. - K  fst ` set Γ"
      using eres_false unfolding trail_false_cls_def trail_false_lit_def .

    moreover have Γ_sorted: "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) Γ"
      using invars10 by (simp add: ord_res_10_invars.simps)

    ultimately have "dropWhile strict_P Γ = dropWhile (λLn. - fst Ln ∉# eres C D) Γ"
    proof (induction Γ)
      case Nil
      show ?case by simp
    next
      case (Cons Ln Γ)

      show ?case
      proof (cases "eres C D = {#}")
        case True
        thus ?thesis
          unfolding strict_P_def linorder_lit.is_maximal_in_mset_iff by simp
      next
        case False

        then obtain L where eres_max_lit: "ord_res.is_maximal_lit L (eres C D)"
          using linorder_lit.ex_maximal_in_mset by metis

        hence strict_P_Ln_iff: "strict_P Ln  atm_of L t atm_of (fst Ln)"
          unfolding strict_P_def
          by (metis linorder_lit.Uniq_is_maximal_in_mset the1_equality')

        show ?thesis
        proof (cases "atm_of L t atm_of (fst Ln)")
          case True

          moreover have "- fst Ln ∉# eres C D"
            using True
            by (smt (verit, best) atm_of_uminus eres_max_lit linorder_lit.dual_order.strict_trans
                linorder_lit.is_maximal_in_mset_iff linorder_lit.neq_iff linorder_trm.order.irrefl
                literal.exhaust_sel ord_res.less_lit_simps(4))

          moreover have "dropWhile strict_P Γ = dropWhile (λLn. - fst Ln ∉# eres C D) Γ"
          proof (rule Cons.IH)
            show "K∈#eres C D. - K  fst ` set Γ"
              using Cons.prems True - fst Ln ∉# eres C D
              using image_iff by fastforce
          next
            show "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) Γ"
              using Cons.prems by simp
          qed

          ultimately show ?thesis
            unfolding dropWhile.simps
            unfolding strict_P_Ln_iff
            by simp
        next
          case False

          hence "atm_of (fst Ln) t atm_of L"
            by order

          hence "atm_of (fst Ln) = atm_of L"
            using Cons.prems
            using atm_of_eq_atm_of eres_max_lit linorder_lit.is_maximal_in_mset_iff by fastforce

          hence "L = fst Ln  L = - fst Ln"
            by (metis atm_of_eq_atm_of)

          moreover have "- fst Ln  fst ` set Γ"
            using sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) (Ln # Γ)
            by fastforce

          moreover have "- L  fst ` (set (Ln # Γ))"
            using Cons.prems(1) eres_max_lit linorder_lit.is_maximal_in_mset_iff by blast

          ultimately have "- fst Ln ∈# eres C D"
            using eres_max_lit linorder_lit.is_maximal_in_mset_iff by auto

          then show ?thesis
            unfolding dropWhile.simps
            unfolding strict_P_Ln_iff
            using False
            by argo
        qed
      qed
    qed

    hence steps11_skip: "(ord_res_11 N)**
      (Uer11, , Γ, Some (eres C D))
      (Uer11, , dropWhile strict_P Γ, Some (eres C D))"
      using rtrancl_ord_res_11_all_skip_steps by metis

    have most_steps11: "(ord_res_11 N)++ (Uer11, , Γ, None)
     (Uer11, , dropWhile strict_P Γ, Some (eres C D))"
      using step11_conf steps11_reso steps11_skip by simp

    show ?thesis
    proof (cases "eres C D = {#}")
      case True

      hence "dropWhile strict_P Γ = []"
        unfolding strict_P_def eres C D = {#}
        unfolding linorder_lit.is_maximal_in_mset_iff
        by simp

      have "Γ' = []"
        unfolding Γ' = dropWhile _ Γ eres C D = {#}
        unfolding linorder_lit.is_maximal_in_mset_iff
        by simp

      show ?thesis
      proof (intro exI conjI)
        show "(constant_context ord_res_11)++ S11 (N, Uer11, , [], Some {#})"
          unfolding S11_def 𝒞 = None
          using most_steps11[unfolded dropWhile strict_P Γ = [] eres C D = {#}]
          using tranclp_constant_context by metis

        show "ord_res_10_matches_ord_res_11 S10' (N, Uer11, , [], Some {#})"
          unfolding S10' = (N, s10') s10' = _ Γ' = []
        proof (rule ord_res_10_matches_ord_res_11.intros)
          show "ord_res_10_invars N (Uer10', , [])"
            using step10 s10' = _ Γ' = [] invars10 ord_res_10_preserves_invars by metis
        next
          show "ord_res_11_invars N (Uer11, , [], Some {#})"
            using most_steps11 invars11
            unfolding 𝒞 = None dropWhile strict_P Γ = [] eres C D = {#}
            by (metis tranclp_ord_res_11_preserves_invars)
        next
          show "Uer11 = Uer10' |-| {|{#}|}"
            unfolding Uer11 = Uer10 Uer10' = finsert (eres C D) Uer10 eres C D = {#}
            using {#} |∉| iefac  |`| (N |∪| Uer10)
            by force
        next
          show "if {#} |∈| iefac  |`| (N |∪| Uer10') then
            [] = []  Some {#} = Some {#} else Some {#} = None"
            unfolding Uer10' = finsert (eres C D) Uer10 eres C D = {#}
            by simp
        qed
      qed
    next
      case False

      then obtain L where eres_max_lit: "ord_res.is_maximal_lit L (eres C D)"
        using linorder_lit.ex_maximal_in_mset by metis

      hence "L ∈# eres C D"
        unfolding linorder_lit.is_maximal_in_mset_iff by argo

      hence "L ∈# C  L  Pos A  L ∈# D  L  Neg A"
        using stronger_lit_in_one_of_resolvents_if_in_eres eres C D  D D_max_lit
        by (metis uminus_Neg)

      hence "L l Neg A"
        using C_max_lit D_max_lit
        unfolding linorder_lit.is_greatest_in_mset_iff linorder_lit.is_maximal_in_mset_iff
        by (metis C_max_lit D_max_lit L ∈# eres C D eres_lt_if ord_res.asymp_less_lit
            ord_res.less_lit_simps(3,4) ord_res.transp_less_lit in_remove1_mset_neq
            linorder_lit.less_than_maximal_if_multpHO linorder_lit.order.not_eq_order_implies_strict
            literal.disc(2) multp_eq_multpHO uminus_Neg)

      have "dropWhile strict_P Γ = dropWhile (λLn. atm_of L t atm_of (fst Ln)) Γ"
        unfolding strict_P_def
        using eres_max_lit
        by (metis (no_types) Uniq_D linorder_lit.Uniq_is_maximal_in_mset)

      also have " = (- L, None) # dropWhile (λLn. atm_of L t atm_of (fst Ln)) Γ"
      proof -
        have "- L  Pos A"
          using L l Neg A by (cases L) simp_all

        hence "- L  fst ` set Γ'''"
          using eres_false L ∈# eres C D
          unfolding trail_false_cls_def trail_false_lit_def 
          unfolding Γ = (Pos A, Some C) # Γ'''
          by auto

        hence "(- L, None)  set Γ'''"
          using Γ_spec unfolding Γ = (Pos A, Some C) # Γ''' by auto

        then obtain Γ0 Γ1 where "Γ''' = Γ1 @ (- L, None) # Γ0"
          by (meson split_list)

        hence "Γ = (Pos A, Some C) # Γ1 @ (- L, None) # Γ0"
          unfolding Γ = (Pos A, Some C) # Γ''' by (simp only:)

        have AAA: "Ln  set ((Pos A, Some C) # Γ1). atm_of L t atm_of (fst Ln)"
          using Γ_sorted unfolding Γ = (Pos A, Some C) # Γ1 @ (- L, None) # Γ0
          by (simp add: sorted_wrt_append)

        hence BBB: "Ln  set ((Pos A, Some C) # Γ1 @ [(- L, None)]). atm_of L t atm_of (fst Ln)"
          by simp

        have "Ln  set Γ0. atm_of (fst Ln) t atm_of L"
          using Γ_sorted unfolding Γ = (Pos A, Some C) # Γ1 @ (- L, None) # Γ0
          by (simp add: sorted_wrt_append)

        hence CCC: "Ln  set Γ0. ¬ atm_of L t atm_of (fst Ln)"
          using linorder_trm.leD by blast

        have "dropWhile (λLn. atm_of L t atm_of (fst Ln)) Γ = (- L, None) # Γ0"
          unfolding Γ = (Pos A, Some C) # Γ1 @ (- L, None) # Γ0
          using dropWhile_append2 AAA by simp

        also have " = (- L, None) # dropWhile (λLn. atm_of L t atm_of (fst Ln)) Γ0"
          using CCC by (simp add: dropWhile_ident_if_pred_always_false)

        also have " = (- L, None) # dropWhile (λLn. atm_of L t atm_of (fst Ln)) Γ"
          unfolding Γ = (Pos A, Some C) # Γ1 @ (- L, None) # Γ0
          using dropWhile_append2 BBB by simp

        finally show ?thesis .
      qed

      also have " = (- L, None) # Γ'"
        unfolding Γ' = dropWhile _ Γ
        using eres_max_lit
        by (metis (no_types) Uniq_D linorder_lit.Uniq_is_maximal_in_mset)

      finally have "dropWhile strict_P Γ = (- L, None) # Γ'" .

      have step10_back: "ord_res_11 N
        (Uer11, , dropWhile strict_P Γ, Some (eres C D))
        (finsert (eres C D) Uer11, , Γ', None)"
        unfolding dropWhile strict_P Γ = (- L, None) # Γ'
      proof (rule ord_res_11.backtrack)
        show "(- L, None) # Γ' = (- L, None) # Γ'" ..
      next
        show "- (- L) ∈# eres C D"
          unfolding uminus_of_uminus_id
          using eres_max_lit
          unfolding linorder_lit.is_maximal_in_mset_iff by argo
      qed

      hence all_steps11: "(ord_res_11 N)++ (Uer11, , Γ, None)
        (finsert (eres C D) Uer11, , Γ', None)"
        using most_steps11 by simp

      show ?thesis
      proof (intro exI conjI)
        show "(constant_context ord_res_11)++ S11 (N, finsert (eres C D) Uer11, , Γ', None)"
          unfolding S11_def 𝒞 = None
          using all_steps11 tranclp_constant_context by metis

        have "{#} |∉| iefac  |`| (N |∪| Uer10')"
          by (smt (verit, del_insts) False {#} |∉| iefac  |`| (N |∪| Uer10) fimage.rep_eq
              fimageE fimageI finsertE funion_iff iefac_def mempty_in_image_efac_iff step_hyps(5))

        show "ord_res_10_matches_ord_res_11 S10' (N, finsert (eres C D) Uer11, , Γ', None)"
          unfolding S10' = (N, s10') s10' = _
        proof (rule ord_res_10_matches_ord_res_11.intros)
          show "ord_res_10_invars N (Uer10', , Γ')"
            using step10 s10' = _ invars10 ord_res_10_preserves_invars by metis
        next
          show "ord_res_11_invars N (finsert (eres C D) Uer11, , Γ', None)"
            using all_steps11 invars11
            unfolding 𝒞 = None
            by (metis tranclp_ord_res_11_preserves_invars)
        next
          show "finsert (eres C D) Uer11 = Uer10' |-| {|{#}|}"
            unfolding Uer11 = Uer10 Uer10' = finsert (eres C D) Uer10
            using {#} |∉| iefac  |`| (N |∪| Uer10')
            using False Uer11 = Uer10 Uer11 = Uer10 |-| {|{#}|} by auto
        next
          show "if {#} |∈| iefac  |`| (N |∪| Uer10') then
            Γ' = []  None = Some {#} else None = None"
            using {#} |∉| iefac  |`| (N |∪| Uer10') by simp
        qed
      qed
    qed
  qed
qed

theorem bisimulation_ord_res_10_ord_res_11:
  defines "match  λ_. ord_res_10_matches_ord_res_11"
  shows "(MATCH :: nat × nat  'f ord_res_10_state  'f ord_res_11_state  bool) .
    bisimulation
      (constant_context ord_res_10) (constant_context ord_res_11)
      ord_res_8_final ord_res_11_final
       MATCH"
(* For AFP-devel
  shows "∃(MATCH :: nat × nat ⇒ 'f ord_res_10_state ⇒ 'f ord_res_11_state ⇒ bool) ℛfb.
    bisimulation
      (constant_context ord_res_10) ord_res_8_final
      (constant_context ord_res_11) ord_res_11_final
      MATCH ℛfb"
*)
proof (rule ex_bisimulation_from_forward_simulation)
  show "right_unique (constant_context ord_res_10)"
    using right_unique_constant_context right_unique_ord_res_10 by metis
next
  show "right_unique (constant_context ord_res_11)"
    using right_unique_constant_context right_unique_ord_res_11 by metis
next
  show "S. ord_res_8_final S  (S'. constant_context ord_res_10 S S')"
    by (metis finished_def ord_res_10_semantics.final_finished)
next
  show "S. ord_res_11_final S  (S'. constant_context ord_res_11 S S')"
    by (metis finished_def ord_res_11_semantics.final_finished)
next
  show "i S10 S11. match i S10 S11  ord_res_8_final S10  ord_res_11_final S11"
    unfolding match_def
    using ord_res_10_final_iff_ord_res_11_final by metis
next
  show "i S10 S11. match i S10 S11 
       safe_state (constant_context ord_res_10) ord_res_8_final S10 
       safe_state (constant_context ord_res_11) ord_res_11_final S11"
  proof (intro allI impI conjI)
    fix i S10 S11
    assume match: "match i S10 S11"
    show "safe_state (constant_context ord_res_10) ord_res_8_final S10"
      using match[unfolded match_def]
      using ord_res_10_safe_state_if_invars
      unfolding ord_res_10_matches_ord_res_11.simps by auto

    show "safe_state (constant_context ord_res_11) ord_res_11_final S11"
      using match[unfolded match_def]
      using ord_res_11_safe_state_if_invars
      using ord_res_10_matches_ord_res_11.simps by auto
  qed
next
  show "wfp (λ_ _. False)"
    by simp
next
  show "i S10 S11 S10'. match i S10 S11  constant_context ord_res_10 S10 S10' 
    (i' S11'. (constant_context ord_res_11)++ S11 S11'  match i' S10' S11') 
    (i'. match i' S10' S11  False)"
    unfolding match_def
    using forward_simulation_between_10_and_11 by metis
qed

end

lemma forward_simulation_composition:
  assumes
    "forward_simulation step1 step2 final1 final2 order1 match1"
    "forward_simulation step2 step3 final2 final3 order2 match2"
  defines "  λi i'. lex_prodp order2++ order1 (prod.swap i) (prod.swap i')"
  shows "forward_simulation step1 step3 final1 final3  (rel_comp match1 match2)"
proof intro_locales
  show "semantics step1 final1"
    using assms
    by (auto intro: forward_simulation.axioms)
next
  show "semantics step3 final3"
    using assms
    by (auto intro: forward_simulation.axioms)
next
  have "well_founded order1" "well_founded order2"
    using assms
    by (auto intro: forward_simulation.axioms)

  hence "wfp "
    unfolding ℛ_def
    by (metis (no_types, lifting) lex_prodp_wfP well_founded.wf wfP_trancl
        wfp_if_convertible_to_wfp)

  thus "well_founded "
    by unfold_locales
next
  show "forward_simulation_axioms step1 step3 final1 final3  (rel_comp match1 match2)"
  proof unfold_locales
    fix i s1 s3
    assume
      match: "rel_comp match1 match2 i s1 s3" and
      final: "final1 s1"
    obtain i1 i2 s2 where "match1 i1 s1 s2" and "match2 i2 s2 s3" and "i = (i1, i2)"
      using match unfolding rel_comp_def by auto
    thus "final3 s3"
      using final assms(1,2)[THEN forward_simulation.match_final]
      by simp
  next
    fix i s1 s3 s1'
    assume
      match: "rel_comp match1 match2 i s1 s3" and
      step: "step1 s1 s1'"
    obtain i1 i2 s2 where "match1 i1 s1 s2" and "match2 i2 s2 s3" and i_def: "i = (i1, i2)"
      using match unfolding rel_comp_def by auto
    from forward_simulation.simulation[OF assms(1) match1 i1 s1 s2 step]
    show "(i' s3'. step3++ s3 s3'  rel_comp match1 match2 i' s1' s3') 
       (i'. rel_comp match1 match2 i' s1' s3   i' i)"
      (is "(i' s1'. ?STEPS i' s1')  (i'. ?STALL i')")
    proof (elim disjE exE conjE)
      fix i1' s2'
      assume "step2++ s2 s2'" and "match1 i1' s1' s2'"
      from forward_simulation.lift_simulation_plus[OF assms(2) step2++ s2 s2' match2 i2 s2 s3]
      show ?thesis
      proof (elim disjE exE conjE)
        fix i2' s3'
        assume "step3++ s3 s3'" and "match2 i2' s2' s3'"
        hence "?STEPS (i1', i2') s3'"
          by (auto intro: match1 i1' s1' s2' simp: rel_comp_def)
        thus ?thesis by auto
      next
        fix i2''
        assume "match2 i2'' s2' s3" and "order2++ i2'' i2"
        hence "?STALL (i1', i2'')"
          unfolding rel_comp_def i_def comp_def prod.swap_def prod.sel
        proof (intro conjI)
          show "(match1 i1' OO match2 i2'') s1' s3"
            using match1 i1' s1' s2' match2 i2'' s2' s3
            by (auto simp add: relcompp_apply)
        next
          show " (i1', i2'') (i1, i2)"
            unfolding ℛ_def lex_prodp_def prod.swap_def prod.sel
            using order2++ i2'' i2 by argo
        qed
        thus ?thesis
          by metis
      qed
    next
      fix i1'
      assume "match1 i1' s1' s2" and "order1 i1' i1"
      hence "?STALL (i1', i2)"
        unfolding rel_comp_def i_def prod.sel
        using match2 i2 s2 s3 by (auto simp: ℛ_def lex_prodp_def)
      thus ?thesis
        by metis
    qed
  qed
qed

text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.forward_simulation_composition}
as it is available in theoryVeriComp.Simulation.›

type_synonym bisim_index_1_2 = "nat × nat"
type_synonym bisim_index_1_3 = "bisim_index_1_2 × (nat × nat)"
type_synonym bisim_index_1_4 = "bisim_index_1_3 × (nat × nat)"
type_synonym bisim_index_1_5 = "bisim_index_1_4 × (nat × nat)"
type_synonym bisim_index_1_6 = "bisim_index_1_5 × (nat × nat)"
type_synonym bisim_index_1_7 = "bisim_index_1_6 × (nat × nat)"
type_synonym bisim_index_1_8 = "bisim_index_1_7 × (nat × nat)"
type_synonym bisim_index_1_9 = "bisim_index_1_8 × (nat × nat)"
type_synonym bisim_index_1_10 = "bisim_index_1_9 × (nat × nat)"
type_synonym bisim_index_1_11 = "bisim_index_1_10 × (nat × nat)"

context simulation_SCLFOL_ground_ordered_resolution begin

theorem forward_simulation_ord_res_1_ord_res_11:
  obtains
    MATCH :: "bisim_index_1_11  'f ord_res_1_state  'f ord_res_11_state  bool" and
     :: "bisim_index_1_11  bisim_index_1_11  bool"
  where
    "forward_simulation
      ord_res_1 (constant_context ord_res_11)
      ord_res_1_final ord_res_11_final
       MATCH"
proof -
  have bi_to_fwd: "MATCH ( :: 'a  'a  bool).
    forward_simulation step1 step2 final1 final2  MATCH"
    if "MATCH ( :: 'a  'a  bool). bisimulation step1 step2 final1 final2  MATCH"
    for step1 step2 final1 final2
    using that by (auto intro: bisimulation.axioms)

  show ?thesis
    using that
  using bisimulation_ord_res_1_ord_res_2[THEN bi_to_fwd]
    bisimulation_ord_res_2_ord_res_3[THEN bi_to_fwd]
    bisimulation_ord_res_3_ord_res_4[THEN bi_to_fwd]
    bisimulation_ord_res_4_ord_res_5[THEN bi_to_fwd]
    bisimulation_ord_res_5_ord_res_6[THEN bi_to_fwd]
    bisimulation_ord_res_6_ord_res_7[THEN bi_to_fwd]
    bisimulation_ord_res_7_ord_res_8[THEN bi_to_fwd]
    bisimulation_ord_res_8_ord_res_9[THEN bi_to_fwd]
    bisimulation_ord_res_9_ord_res_10[THEN bi_to_fwd]
    bisimulation_ord_res_10_ord_res_11[THEN bi_to_fwd]
  using forward_simulation_composition by meson
qed

theorem backward_simulation_ord_res_1_ord_res_11:
  obtains
    MATCH :: "bisim_index_1_11  'f ord_res_1_state  'f ord_res_11_state  bool" and
     :: "bisim_index_1_11  bisim_index_1_11  bool"
  where
    "backward_simulation
      ord_res_1 (constant_context ord_res_11)
      ord_res_1_final ord_res_11_final
       MATCH"
proof -
  have bi_to_bwd: "MATCH ( :: 'a  'a  bool).
    backward_simulation step1 step2 final1 final2  MATCH"
    if "MATCH ( :: 'a  'a  bool). bisimulation step1 step2 final1 final2  MATCH"
    for step1 step2 final1 final2
    using that by (auto intro: bisimulation.axioms)

  show ?thesis
    using that
  using bisimulation_ord_res_1_ord_res_2[THEN bi_to_bwd]
    bisimulation_ord_res_2_ord_res_3[THEN bi_to_bwd]
    bisimulation_ord_res_3_ord_res_4[THEN bi_to_bwd]
    bisimulation_ord_res_4_ord_res_5[THEN bi_to_bwd]
    bisimulation_ord_res_5_ord_res_6[THEN bi_to_bwd]
    bisimulation_ord_res_6_ord_res_7[THEN bi_to_bwd]
    bisimulation_ord_res_7_ord_res_8[THEN bi_to_bwd]
    bisimulation_ord_res_8_ord_res_9[THEN bi_to_bwd]
    bisimulation_ord_res_9_ord_res_10[THEN bi_to_bwd]
    bisimulation_ord_res_10_ord_res_11[THEN bi_to_bwd]
  using backward_simulation_composition by meson
qed


(* For AFP-devel, uncomment this theorem; it depends on a changed definition of the locale
bisimulation.

theorem bisimulation_ord_res_1_ord_res_11:
  obtains
    MATCH :: "bisim_index_1_11 ⇒ 'f ord_res_1_state ⇒ 'f ord_res_11_state ⇒ bool" and
    ℛfb :: "bisim_index_1_11 ⇒ bisim_index_1_11 ⇒ bool"
  where
    "bisimulation
      ord_res_1 ord_res_1_final
      (constant_context ord_res_11) ord_res_11_final
      MATCH ℛfb"
  apply atomize_elim
  using bisimulation_ord_res_1_ord_res_2
    bisimulation_ord_res_2_ord_res_3
    bisimulation_ord_res_3_ord_res_4
    bisimulation_ord_res_4_ord_res_5
    bisimulation_ord_res_5_ord_res_6
    bisimulation_ord_res_6_ord_res_7
    bisimulation_ord_res_7_ord_res_8
    bisimulation_ord_res_8_ord_res_9
    bisimulation_ord_res_9_ord_res_10
    bisimulation_ord_res_10_ord_res_11
  using bisimulation_composition by meson

theorem
  obtains
    MATCH :: "bisim_index_1_11 ⇒ 'f ord_res_1_state ⇒ 'f ord_res_11_state ⇒ bool" and
    ℛfb :: "bisim_index_1_11 ⇒ bisim_index_1_11 ⇒ bool"
  where
    "bisimulation
      ord_res_1 ord_res_1_final
      (constant_context ord_res_11) ord_res_11_final
      MATCH ℛfb" and
    "⋀j S1 S11. MATCH j S1 S11 ⟹ ord_res_1_final S1 ⟷ ord_res_11_final S11"
  using bisimulation_ord_res_1_ord_res_11 bisimulation.agree_on_final
  by (metis (no_types, opaque_lifting))
*)


section ‹ORD-RES-11 is a regular SCL strategy›

definition gtrailelem_of_trailelem where
  "gtrailelem_of_trailelem  λ(L, opt).
    (lit_of_glit L, map_option (λC. (cls_of_gcls {#K ∈# C. K  L#}, lit_of_glit L, Var)) opt)"

fun state_of_gstate :: "_  ('f, 'v) SCL_FOL.state" where
  "state_of_gstate (UG, _, ΓG, 𝒞G) =
    (let
      Γ = map gtrailelem_of_trailelem ΓG;
      U = cls_of_gcls |`| UG;
      𝒞 = map_option (λCG. (cls_of_gcls CG, Var)) 𝒞G
    in (Γ, U, 𝒞))"


lemma fst_case_prod_simp: "fst (case p of (x, y)  (f x, g x y)) = f (fst p)"
  by (cases p) simp

lemma trail_false_cls_nonground_iff_trail_false_cls_ground:
  fixes ΓG and DG :: "'f gclause"
  fixes Γ :: "('f, 'v) SCL_FOL.trail" and D :: "('f, 'v) term clause"
  defines "Γ  map gtrailelem_of_trailelem ΓG" and "D  cls_of_gcls DG"
  shows "trail_false_cls Γ D  trail_false_cls ΓG DG"
proof -
  have "trail_false_cls Γ D  (L ∈# D. trail_false_lit Γ L)"
    unfolding trail_false_cls_def ..
  also have "  (LG ∈# DG. trail_false_lit Γ (lit_of_glit LG))"
    unfolding D_def cls_of_gcls_def by simp
  also have "  (LG ∈# DG. trail_false_lit ΓG LG)"
  proof -
    have "trail_false_lit Γ (lit_of_glit LG)  trail_false_lit ΓG LG"
      for LG :: "'f gterm literal"
    proof -
      have "trail_false_lit Γ (lit_of_glit LG)  - lit_of_glit LG  fst ` set Γ"
        unfolding trail_false_lit_def ..
      also have " 
        - (lit_of_glit LG :: ('f, 'v) term literal)  set (map (λx. lit_of_glit (fst x)) ΓG)"
        unfolding Γ_def image_set list.map_comp
        unfolding gtrailelem_of_trailelem_def
        unfolding list.map_comp
        unfolding comp_def fst_case_prod_simp ..
      also have "  (lit_of_glit (- LG) :: ('f, 'v) term literal)  lit_of_glit ` fst ` set ΓG"
        by (cases LG) (auto simp: lit_of_glit_def)
      also have "  - LG  fst ` set ΓG"
        using inj_image_mem_iff inj_lit_of_glit by metis
      also have "  trail_false_lit ΓG LG"
        unfolding trail_false_lit_def ..
      finally show "trail_false_lit Γ (lit_of_glit LG) = trail_false_lit ΓG LG" .
    qed
    thus ?thesis by metis
  qed
  also have "  trail_false_cls ΓG DG"
    unfolding trail_false_cls_def ..
  finally show ?thesis .
qed

theorem ord_res_11_is_strategy_for_regular_scl:
  fixes
    NG :: "'f gclause fset" and
    N :: "('f, 'v) term clause fset" and
    βG :: "'f gterm" and
    β :: "('f, 'v) term" and
    SG SG' :: "'f gclause fset × 'f gclause fset × ('f gliteral × 'f gclause option) list × 'f gclause option" and
    S S' :: "('f, 'v) SCL_FOL.state"
  defines
    "N  cls_of_gcls |`| NG" and
    "β  term_of_gterm βG" and
    "S  state_of_gstate SG" and
    "S'  state_of_gstate SG'"
  assumes
    ball_le_βG: "AG |∈| atms_of_clss NG. AG t βG" and
    run: "(ord_res_11 NG)** ({||}, {||}, [], None) SG" and
    step: "ord_res_11 NG SG SG'"
  shows
    "scl_fol.regular_scl N β S S'"
proof -
  have "ord_res_11_invars NG ({||}, {||}, [], None)"
    using ord_res_11_invars_initial_state .

  hence "ord_res_11_invars NG SG"
    using run rtranclp_ord_res_11_preserves_invars by metis

  obtain UG  ΓG 𝒞G where SG_def: "SG = (UG, , ΓG, 𝒞G)"
    by (metis surj_pair)

  obtain Γ U 𝒞 where S_def: "S = (Γ, U, 𝒞)"
    by (metis surj_pair)

  have Γ_def: "Γ = map gtrailelem_of_trailelem ΓG"
    using S_def SG_def S  state_of_gstate SG by simp

  have U_def: "U = cls_of_gcls |`| UG"
    using S_def SG_def S  state_of_gstate SG by simp

  have 𝒞_def: "𝒞 = map_option (λCG. (cls_of_gcls CG, Var)) 𝒞G"
    using S_def SG_def S  state_of_gstate SG by simp

  obtain ℱ' UG' :: "'f gclause fset" and ΓG' :: "_ list" and 𝒞G' :: "_ option" where
    SG'_def: "SG' = (UG', ℱ', ΓG', 𝒞G')"
    by (metis surj_pair)

  obtain Γ' :: "_ list" and U' :: "_ fset" and 𝒞' :: "_ option" where
    S'_def: "S' = (Γ', U', 𝒞')"
    by (metis surj_pair)

  have Γ'_def: "Γ' = map gtrailelem_of_trailelem ΓG'"
    using S'_def SG'_def S'  state_of_gstate SG' by simp

  have U'_def: "U' = cls_of_gcls |`| UG'"
    using S'_def SG'_def S'  state_of_gstate SG' by simp

  have 𝒞'_def: "𝒞' = map_option (λCG. (cls_of_gcls CG, Var)) 𝒞G'"
    using S'_def SG'_def S'  state_of_gstate SG' by simp

  have "atms_of_clss UG |⊆| atms_of_clss NG"
    using ord_res_11_invars NG SG[unfolded SG_def]
    unfolding ord_res_11_invars.simps by simp

  have "atms_of_clss (NG |∪| UG) = atms_of_clss NG |∪| atms_of_clss UG"
    by (simp add: atms_of_clss_def fimage_funion)

  also have " = atms_of_clss NG"
    using atms_of_clss UG |⊆| atms_of_clss NG by auto

  finally have "atms_of_clss (NG |∪| UG) = atms_of_clss NG" .

  have clauses_in_ℱ_have_pos_max_lit: "C|∈|. L. is_pos L  ord_res.is_maximal_lit L C"
    using ord_res_11_invars NG SG[unfolded SG_def ord_res_11_invars.simps]
    by simp

  have nex_conflict_if_nbex_trail_false:
    "¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)  ¬ Ex (scl_fol.conflict N β S)"
  proof (elim contrapos_nn exE)
    fix x :: "('f, 'v) state"
    assume "scl_fol.conflict N β S x"
    hence "fBex (NG |∪| UG) (trail_false_cls ΓG)"
      unfolding S_def
    proof (cases N β "(Γ, U, 𝒞)" x rule: scl_fol.conflict.cases)
      case (conflictI D γ)

      obtain DG where "DG |∈| NG |∪| UG" and D_def: "D = cls_of_gcls DG"
        using D |∈| N |∪| U
        unfolding N_def U_def by blast

      moreover have "trail_false_cls ΓG DG"
      proof -
        have "is_ground_cls D"
          using D = cls_of_gcls DG by simp
        hence "D  γ = D"
          by simp
        hence "trail_false_cls Γ D"
          using conflictI
          unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
          unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
          by argo

        thus ?thesis
          unfolding Γ_def D_def
          unfolding trail_false_cls_nonground_iff_trail_false_cls_ground .
      qed
      ultimately show ?thesis by metis
    qed

    thus "fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)"
      unfolding bex_trail_false_cls_simp .
  qed

  have nex_conflict_if_alread_in_conflict: "𝒞G = Some CG  ¬ Ex (scl_fol.conflict N β S)" for CG
    unfolding S_def 𝒞_def by (simp add: scl_fol.conflict.simps)

  have nex_conflict_if_no_clause_could_propagate_comp:
    "¬ Ex (scl_fol.conflict N β ((lit_of_glit LG, None) # Γ, U, 𝒞))"
    if
      nex_false_clause_wrt_ΓG: "¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)" and
      ball_lt_atm_LG: "x |∈| trail_atms ΓG. x t atm_of LG" and
      nex_clause_that_propagate: "¬ (C |∈| iefac  |`| (NG |∪| UG).
        clause_could_propagate ΓG C (- LG))"
    for LG
  proof (intro notI, elim exE)
    fix S'' :: "('f, 'v) SCL_FOL.state"
    assume "scl_fol.conflict N β ((lit_of_glit LG, None) # Γ, U, 𝒞) S''"
    thus "False"
    proof (cases N β "((lit_of_glit LG, None) # Γ, U, 𝒞)" S'' rule: scl_fol.conflict.cases)
      case (conflictI D γ)

      obtain DG where "DG |∈| NG |∪| UG" and D_def: "D = cls_of_gcls DG"
        using D |∈| N |∪| U N_def U_def by blast

      have "(lit_of_glit LG :: ('f, 'v) term literal, None) # Γ =
        (map gtrailelem_of_trailelem ((LG, None) # ΓG) :: ('f, 'v) SCL_FOL.trail)"
        by (simp add: Γ_def gtrailelem_of_trailelem_def)

      moreover have "D  γ = cls_of_gcls DG"
        unfolding D_def by simp

      ultimately have "trail_false_cls
        (map gtrailelem_of_trailelem ((LG, None) # ΓG) :: ('f, 'v) SCL_FOL.trail) (cls_of_gcls DG)"
        using SCL_FOL.trail_false_cls ((lit_of_glit LG, None) # Γ) (D  γ)
        unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
        unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
        by metis

      hence "trail_false_cls ((LG, None) # ΓG) DG"
        using trail_false_cls_nonground_iff_trail_false_cls_ground by blast

      hence "trail_false_cls ΓG {#KG ∈# DG. KG  - LG#}"
        unfolding trail_false_cls_def trail_false_lit_def
        by auto

      moreover have "ord_res.is_maximal_lit (- LG) DG"
        unfolding linorder_lit.is_maximal_in_mset_iff
      proof (intro conjI ballI impI)
        show "- LG ∈# DG"
          using DG |∈| NG |∪| UG trail_false_cls ((LG, None) # ΓG) DG subtrail_falseI
            nex_false_clause_wrt_ΓG
          unfolding bex_trail_false_cls_simp
          by blast
      next
        fix KG assume "KG ∈# DG" and "KG  - LG"
        hence "trail_false_lit ΓG KG"
          using trail_false_cls ΓG {#KG ∈# DG. KG  - LG#}
          unfolding trail_false_cls_def by simp
        hence "trail_defined_lit ΓG KG"
          by (simp add: trail_defined_lit_iff_true_or_false)
        hence "atm_of KG |∈| trail_atms ΓG"
          unfolding trail_defined_lit_iff_trail_defined_atm .
        hence "atm_of KG t atm_of LG"
          using ball_lt_atm_LG by metis
        hence "KG l - LG"
          by (cases LG; cases KG) simp_all
        thus "¬ - LG l KG"
          by order
      qed

      moreover have "¬ trail_defined_lit ΓG (- LG)"
        by (metis atm_of_uminus linorder_trm.less_irrefl that(2)
            trail_defined_lit_iff_trail_defined_atm)

      ultimately have "clause_could_propagate ΓG DG (- LG)"
        unfolding clause_could_propagate_def by argo

      hence "C|∈| NG |∪| UG. clause_could_propagate ΓG C (- LG)"
        using DG |∈| NG |∪| UG by metis

      hence False
         using nex_clause_that_propagate
         unfolding bex_clause_could_propagate_simp
         by contradiction

      thus ?thesis .
    qed
  qed

  show ?thesis
    using step unfolding SG_def SG'_def
  proof (cases NG "(UG, , ΓG, 𝒞G)" "(UG', ℱ', ΓG', 𝒞G')" rule: ord_res_11.cases)
    case step_hyps: (decide_neg AG)

    define A :: "('f, 'v) term" where
      "A = term_of_gterm AG"

    let ?f = "gtrailelem_of_trailelem"
    have "Γ' = map ?f ΓG'"
      unfolding Γ'_def ..
    also have " = map ?f ((Neg AG, None) # ΓG)"
      unfolding ΓG' = (Neg AG, None) # ΓG ..
    also have " = ?f (Neg AG, None) # map ?f ΓG"
      unfolding list.map ..
    also have " = ?f (Neg AG, None) # Γ"
      unfolding Γ_def ..
    also have " = (lit_of_glit (Neg AG), None) # Γ"
      unfolding gtrailelem_of_trailelem_def prod.case option.map ..
    also have " = (Neg (term_of_gterm AG), None) # Γ"
      unfolding lit_of_glit_def literal.map ..
    also have " = (Neg A, None) # Γ"
      unfolding A_def ..
    finally have "Γ' = decide_lit (Neg A) # Γ"
      unfolding decide_lit_def .

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    have "¬ Ex (scl_fol.conflict N β S)"
      using ¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG) nex_conflict_if_nbex_trail_false
      by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.decide N β S S'"
        unfolding S_def S'_def U' = U 𝒞_def 𝒞'_def 𝒞G = None 𝒞G' = None option.map
      proof (rule scl_fol.decideI')
        show "is_ground_lit (Neg A ⋅l Var)"
          by (simp add: A_def)
      next
        have "x |∈| trail_atms ΓG. x t AG"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        hence "AG |∉| trail_atms ΓG"
          by blast
        hence "AG  atm_of ` fst ` set ΓG"
          unfolding fset_trail_atms .
        hence "term_of_gterm AG  term_of_gterm ` atm_of ` fst ` set ΓG"
          using inj_image_mem_iff inj_term_of_gterm by metis
        hence "term_of_gterm AG  set (map (λx. term_of_gterm (atm_of (fst x))) ΓG)"
          unfolding image_set list.map_comp comp_def .
        hence "A  set (map (λx. atm_of (lit_of_glit (fst x))) ΓG)"
          unfolding A_def atm_of_lit_of_glit_conv .
        hence "A  atm_of ` fst ` set Γ"
          unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
            fst_case_prod_simp .
        hence "A |∉| trail_atms Γ"
          unfolding fset_trail_atms .
        hence "¬ trail_defined_lit Γ (Neg A ⋅l Var)"
          by (simp add: trail_defined_lit_iff_trail_defined_atm)
        thus "¬ SCL_FOL.trail_defined_lit Γ (Neg A ⋅l Var)"
          by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)
      next
        have "AG |∈| atms_of_clss (NG |∪| UG)"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by blast
        hence "AG |∈| atms_of_clss NG"
          unfolding atms_of_clss (NG |∪| UG) = atms_of_clss NG .
        hence "AG t βG"
          using ball_le_βG by metis
        moreover have "gterm_of_term A = AG"
          by (simp add: A_def)
        moreover have "gterm_of_term β = βG"
          by (simp add: β_def)
        ultimately have "gterm_of_term A t gterm_of_term β"
          by argo
        thus "less_B== (atm_of (Neg A) ⋅a Var) β"
          using inj_term_of_gterm[THEN injD]
          by (auto simp: less_B_def A_def β_def)
      next
        show "Γ' = trail_decide Γ (Neg A ⋅l Var)"
          using Γ' = decide_lit (Neg A) # Γ
          unfolding subst_lit_id_subst .
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.conflict N β S' S''"

      moreover have "S''. scl_fol.conflict N β S' S''"
      proof -
        have "¬ Ex (scl_fol.conflict N β ((lit_of_glit (Neg AG), None) # Γ, U, 𝒞))"
        proof (rule nex_conflict_if_no_clause_could_propagate_comp)
          show "¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)"
            using step_hyps by argo
        next
          show "x |∈| trail_atms ΓG. x t atm_of (Neg AG)"
            unfolding literal.sel
            using step_hyps linorder_trm.is_least_in_fset_iff by simp
        next
          show "¬ (C |∈| iefac  |`| (NG |∪| UG). clause_could_propagate ΓG C (- Neg AG))"
            using step_hyps by simp
        qed
        moreover have "lit_of_glit (Neg AG) = Neg A"
          unfolding A_def lit_of_glit_def literal.map ..
        ultimately show ?thesis
          unfolding S'_def Γ' = decide_lit (Neg A) # Γ decide_lit_def
          using 𝒞'_def 𝒞_def U' = U step_hyps(1,4) by argo
      qed

      ultimately show False
        by metis
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (decide_pos AG)

    define A :: "('f, 'v) term" where
      "A = term_of_gterm AG"

    let ?f = "gtrailelem_of_trailelem"
    have "Γ' = map ?f ΓG'"
      unfolding Γ'_def ..
    also have " = map ?f ((Pos AG, None) # ΓG)"
      unfolding ΓG' = (Pos AG, None) # ΓG ..
    also have " = ?f (Pos AG, None) # map ?f ΓG"
      unfolding list.map ..
    also have " = ?f (Pos AG, None) # Γ"
      unfolding Γ_def ..
    also have " = (lit_of_glit (Pos AG), None) # Γ"
      unfolding prod.case option.map gtrailelem_of_trailelem_def ..
    also have " = (Pos (term_of_gterm AG), None) # Γ"
      unfolding lit_of_glit_def literal.map ..
    also have " = (Pos A, None) # Γ"
      unfolding A_def ..
    finally have "Γ' = decide_lit (Pos A) # Γ"
      unfolding decide_lit_def .

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    have "¬ Ex (scl_fol.conflict N β S)"
      using step_hyps nex_conflict_if_nbex_trail_false by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.decide N β S S'"
        unfolding S_def S'_def U' = U 𝒞_def 𝒞'_def 𝒞G = None 𝒞G' = None option.map
      proof (rule scl_fol.decideI')
        show "is_ground_lit (Pos A ⋅l Var)"
          by (simp add: A_def)
      next
        have "x |∈| trail_atms ΓG. x t AG"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        hence "AG |∉| trail_atms ΓG"
          by blast
        hence "AG  atm_of ` fst ` set ΓG"
          unfolding fset_trail_atms .
        hence "term_of_gterm AG  term_of_gterm ` atm_of ` fst ` set ΓG"
          using inj_image_mem_iff inj_term_of_gterm by metis
        hence "term_of_gterm AG  set (map (λx. term_of_gterm (atm_of (fst x))) ΓG)"
          unfolding image_set list.map_comp comp_def .
        hence "A  set (map (λx. atm_of (lit_of_glit (fst x))) ΓG)"
          unfolding A_def atm_of_lit_of_glit_conv .
        hence "A  atm_of ` fst ` set Γ"
          unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
            fst_case_prod_simp .
        hence "A |∉| trail_atms Γ"
          unfolding fset_trail_atms .
        hence "¬ trail_defined_lit Γ (Pos A ⋅l Var)"
          by (simp add: trail_defined_lit_iff_trail_defined_atm)
        thus "¬ SCL_FOL.trail_defined_lit Γ (Pos A ⋅l Var)"
          by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)
      next
        have "AG |∈| atms_of_clss (NG |∪| UG)"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        hence "AG |∈| atms_of_clss NG"
          unfolding atms_of_clss (NG |∪| UG) = atms_of_clss NG .
        hence "AG t βG"
          using ball_le_βG by metis
        moreover have "gterm_of_term A = AG"
          by (simp add: A_def)
        moreover have "gterm_of_term β = βG"
          by (simp add: β_def)
        ultimately have "gterm_of_term A t gterm_of_term β"
          by argo
        thus "less_B== (atm_of (Pos A) ⋅a Var) β"
          using inj_term_of_gterm[THEN injD]
          by (auto simp: less_B_def A_def β_def)
      next
        show "Γ' = trail_decide Γ (Pos A ⋅l Var)"
          using Γ' = decide_lit (Pos A) # Γ
          unfolding subst_lit_id_subst .
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.conflict N β S' S''"
      
      moreover have "S''. scl_fol.conflict N β S' S''"
      proof -
        have "¬ Ex (scl_fol.conflict N β ((lit_of_glit (Pos AG), None) # Γ, U, 𝒞))"
        proof (rule nex_conflict_if_no_clause_could_propagate_comp)
          show "¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)"
            using step_hyps by argo
        next
          show "x|∈| trail_atms ΓG. x t atm_of (Pos AG)"
            unfolding literal.sel
            using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        next
          have "clause_could_propagate ΓG C (Neg AG)  trail_false_cls ΓG' C" for C
            unfolding ΓG' = (Pos AG, None) # ΓG
            using trail_false_if_could_have_propagatated by fastforce

          thus "¬ (C |∈| iefac  |`| (NG |∪| UG). clause_could_propagate ΓG C (- Pos AG))"
            unfolding uminus_Pos
            using step_hyps by metis
        qed
        moreover have "lit_of_glit (Pos AG) = Pos A"
          unfolding A_def lit_of_glit_def literal.map ..
        ultimately show ?thesis
          unfolding S'_def Γ' = decide_lit (Pos A) # Γ decide_lit_def
          using 𝒞'_def 𝒞_def U' = U step_hyps(1) step_hyps(3) by argo
      qed

      ultimately show False by metis
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (propagate AG CG)

    have "CG |∈| iefac  |`| (NG |∪| UG)" and CG_prop: "clause_could_propagate ΓG CG (Pos AG)"
      using step_hyps linorder_cls.is_least_in_fset_iff by simp_all

    define A :: "('f, 'v) term" where
      "A = term_of_gterm AG"

    define C :: "('f, 'v) term clause" where
      "C = cls_of_gcls CG"

    have "ord_res.is_maximal_lit (Pos AG) CG" and "trail_false_cls ΓG {#K ∈# CG. K  Pos AG#}"
      using clause_could_propagate ΓG CG (Pos AG)
      unfolding clause_could_propagate_def by metis+

    then obtain CG' where "CG = add_mset (Pos AG) CG'"
      by (metis linorder_lit.is_maximal_in_mset_iff mset_add)

    define C' :: "('f, 'v) term clause" where
      "C' = cls_of_gcls CG'"

    let ?f = "gtrailelem_of_trailelem"
    have "Γ' = map ?f ΓG'"
      unfolding Γ'_def ..
    also have " = map ?f ((Pos AG, Some (efac CG)) # ΓG)"
      unfolding ΓG' = (Pos AG, Some _) # ΓG ..
    also have " = ?f (Pos AG, Some (efac CG)) # map ?f ΓG"
      unfolding list.map ..
    also have " = ?f (Pos AG, Some (efac CG)) # Γ"
      unfolding Γ_def ..
    also have " = (lit_of_glit (Pos AG),
      Some (cls_of_gcls {#K ∈# efac CG. K  Pos AG#}, lit_of_glit (Pos AG), Var)) # Γ"
      unfolding gtrailelem_of_trailelem_def prod.case option.map ..
    also have " = (lit_of_glit (Pos AG),
      Some (cls_of_gcls {#K ∈# add_mset (Pos AG) {#K ∈# CG. K  Pos AG#}. K  Pos AG#},
        lit_of_glit (Pos AG), Var)) # Γ"
    proof -
      have "is_pos (Pos AG)"
        by simp

      moreover have "linorder_lit.is_maximal_in_mset CG (Pos AG)"
        using CG_prop unfolding clause_could_propagate_def by argo

      ultimately show ?thesis
        using efac_spec_if_pos_lit_is_maximal by metis
    qed
    also have " = (lit_of_glit (Pos AG),
      Some (cls_of_gcls {#K ∈# CG. K  Pos AG#}, lit_of_glit (Pos AG), Var)) # Γ"
      by (simp add: filter_filter_mset)
    also have " = (Pos A, Some (cls_of_gcls {#K ∈# CG. K  Pos AG#}, Pos A, Var)) # Γ"
      by (simp add: A_def lit_of_glit_def)
    also have " = (Pos A, Some (cls_of_gcls {#L ∈# CG. lit_of_glit L  Pos A#}, Pos A, Var)) # Γ"
      by (metis A_def glit_of_lit_lit_of_glit lit_of_glit_def literal.simps(9))
    also have " = (Pos A, Some ({#L ∈# cls_of_gcls CG. L  Pos A#}, Pos A, Var)) # Γ"
      unfolding cls_of_gcls_def
      unfolding image_mset_filter_mset_swap[of "lit_of_glit" "λL. L  Pos A" CG]
      unfolding cls_of_gcls_def[symmetric] ..
    also have " = (Pos A ⋅l Var, Some ({#L ∈# cls_of_gcls CG. L  Pos A#}, Pos A, Var)) # Γ"
      by simp
    also have " = (Pos A ⋅l Var, Some ({#L ∈# C. L  Pos A#}, Pos A, Var)) # Γ"
      unfolding C_def ..
    finally have "Γ' = propagate_lit (Pos A) {#L ∈# C. L  Pos A#} Var # Γ"
      unfolding propagate_lit_def .

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    obtain CG0 where "CG0 |∈| NG |∪| UG" and "CG = iefac  CG0"
      using CG |∈| iefac  |`| (NG |∪| UG) by blast

    define C0 :: "('f, 'v) term clause" where
      "C0 = cls_of_gcls CG0"

    have "ord_res.is_maximal_lit (Pos AG) CG0"
      using ord_res.is_maximal_lit (Pos AG) CG CG = iefac  CG0
      by (metis iefac_def linorder_lit.is_maximal_in_mset_iff set_mset_efac)

    have "¬ Ex (scl_fol.conflict N β S)"
      using step_hyps nex_conflict_if_nbex_trail_false by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.propagate N β S S'"
        unfolding S_def S'_def U' = U 𝒞_def 𝒞'_def 𝒞G = None 𝒞G' = None option.map
      proof (rule scl_fol.propagateI')
        show "C0 |∈| N |∪| U"
          unfolding C0_def N_def U_def
          using CG0 |∈| NG |∪| UG
          by blast
      next
        show "is_ground_cls (C0  Var)"
          by (simp add: C0_def)

        have "Pos A ∈# C0"
          unfolding A_def C0_def
          by (metis (no_types, lifting) CG = iefac  CG0 ord_res.is_maximal_lit (Pos AG) CG
              cls_of_gcls_def iefac_def image_eqI linorder_lit.is_maximal_in_mset_iff
              lit_of_glit_def literal.map(1) multiset.set_map set_mset_efac)

        then show "C0 = add_mset (Pos A) (C0 - {#Pos A#})"
          by simp

        have "AG |∈| atms_of_clss (NG |∪| UG)"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        hence "AG |∈| atms_of_clss NG"
          unfolding atms_of_clss (NG |∪| UG) = atms_of_clss NG .
        hence "AG t βG"
          using ball_le_βG by metis
        moreover have "gterm_of_term A = AG"
          by (simp add: A_def)
        moreover have "gterm_of_term β = βG"
          by (simp add: β_def)
        ultimately have "gterm_of_term A t gterm_of_term β"
          by argo
        hence "less_B== A β"
          by (auto simp: less_B_def A_def β_def)

        show "K ∈# C0  Var. less_B== (atm_of K) β"
          unfolding subst_cls_id_subst
        proof (intro ballI)
          fix K :: "('f, 'v) Term.term literal"
          assume "K ∈# C0"
          then obtain KG where "KG ∈# CG0" and K_def: "K = lit_of_glit KG"
            unfolding C0_def cls_of_gcls_def by blast

          have "KG l Pos AG"
            using ord_res.is_maximal_lit (Pos AG) CG0 KG ∈# CG0
            unfolding linorder_lit.is_maximal_in_mset_iff by fastforce

          hence "atm_of KG t AG"
            by (metis literal.collapse(1) literal.collapse(2) literal.sel(1)
                ord_res.less_lit_simps(1) ord_res.less_lit_simps(4) reflclp_iff)

          hence "less_B== (atm_of K) A"
            by (auto simp: less_B_def K_def A_def atm_of_lit_of_glit_conv)

          then show "less_B== (atm_of K) β"
            using less_B== A β by order
        qed

        show "{#K ∈# C0. K  Pos A#} = {#K ∈# remove1_mset (Pos A) C0. K ⋅l Var  Pos A ⋅l Var#}"
          by simp

        show "{#K ∈# remove1_mset (Pos A) C0. K = Pos A#} =
          {#K ∈# remove1_mset (Pos A) C0. K ⋅l Var = Pos A ⋅l Var#}"
          by simp

        have "trail_false_cls ΓG ({#KG ∈# CG0. KG  Pos AG#})"
          by (smt (verit, ccfv_threshold) CG = iefac  CG0
              trail_false_cls ΓG {#K ∈# CG. K  Pos AG#} iefac_def mem_Collect_eq set_mset_efac
              set_mset_filter trail_false_cls_def)

        moreover have "(cls_of_gcls {#KG ∈# CG0. KG  Pos AG#} :: ('f, 'v) term clause) =
          {#K ∈# C0. K  Pos A#}"
          by (smt (verit) A_def C0_def cls_of_gcls_def filter_mset_cong0 glit_of_lit_lit_of_glit
              image_mset_filter_mset_swap lit_of_glit_def literal.map(1))

        ultimately have "trail_false_cls Γ ({#K ∈# C0. K  Pos A#}  Var)"
          unfolding subst_cls_id_subst
          using trail_false_cls_nonground_iff_trail_false_cls_ground[THEN iffD2]
          by (metis Γ_def)

        thus "SCL_FOL.trail_false_cls Γ ({#K ∈# C0. K  Pos A#}  Var)"
          unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
          unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
          by argo

        have "x |∈| trail_atms ΓG. x t AG"
          using step_hyps linorder_trm.is_least_in_ffilter_iff by simp
        hence "AG |∉| trail_atms ΓG"
          by blast
        hence "AG  atm_of ` fst ` set ΓG"
          unfolding fset_trail_atms .
        hence "term_of_gterm AG  term_of_gterm ` atm_of ` fst ` set ΓG"
          using inj_image_mem_iff inj_term_of_gterm by metis
        hence "term_of_gterm AG  set (map (λx. term_of_gterm (atm_of (fst x))) ΓG)"
          unfolding image_set list.map_comp comp_def .
        hence "A  set (map (λx. atm_of (lit_of_glit (fst x))) ΓG)"
          unfolding A_def atm_of_lit_of_glit_conv .
        hence "A  atm_of ` fst ` set Γ"
          unfolding image_set list.map_comp comp_def Γ_def gtrailelem_of_trailelem_def
            fst_case_prod_simp .
        hence "A |∉| trail_atms Γ"
          unfolding fset_trail_atms .
        hence "¬ trail_defined_lit Γ (Pos A ⋅l Var)"
          by (simp add: trail_defined_lit_iff_trail_defined_atm)
        thus "¬ SCL_FOL.trail_defined_lit Γ (Pos A ⋅l Var)"
          by (simp add: SCL_FOL.trail_defined_lit_def trail_defined_lit_def)

        have "set_mset (add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C0. K = Pos A#}) =
          {Pos A}"
          by (smt (verit) Collect_cong insert_compr mem_Collect_eq set_mset_add_mset_insert
              set_mset_filter singletonD)
        hence "is_unifier Var (atm_of ` set_mset
          (add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C0. K = Pos A#}))"
          by (metis (no_types, lifting) finite_imageI finite_set_mset image_empty image_insert
              is_unifier_alt singletonD)
        hence "is_unifiers Var {atm_of ` set_mset
          (add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C0. K = Pos A#})}"
          unfolding SCL_FOL.is_unifiers_def by simp
        thus "SCL_FOL.is_imgu Var {atm_of ` set_mset
          (add_mset (Pos A) {#K ∈# remove1_mset (Pos A) C0. K = Pos A#})}"
          unfolding SCL_FOL.is_imgu_def by simp

        have "{#K ∈# CG0. K  Pos AG#} = {#K ∈# CG. K  Pos AG#}"
          using ord_res.is_maximal_lit (Pos AG) CG
          using ord_res.is_maximal_lit (Pos AG) CG0
          unfolding CG = iefac  CG0
          by (metis add_mset_remove_trivial efac_spec_if_pos_lit_is_maximal
              ex1_efac_eq_factoring_chain factorizable_if_neq_efac iefac_def literal.disc(1))

        hence "{#K ∈# C0. K  Pos A#} = {#K ∈# C. K  Pos A#}"
          unfolding C0_def C_def
          by (smt (verit, ccfv_SIG) A_def cls_of_gcls_def filter_mset_cong0 glit_of_lit_lit_of_glit
              image_mset_filter_mset_swap lit_of_glit_def literal.map(1))

        thus "Γ' = trail_propagate Γ (Pos A ⋅l Var) ({#K ∈# C0. K  Pos A#}  Var) Var"
          unfolding subst_cls_id_subst subst_lit_id_subst
          using Γ' = propagate_lit (Pos A) {#L ∈# C. L  Pos A#} Var # Γ
          by argo
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.decide N β S S'"
      thus False
        unfolding S_def S'_def
      proof (cases N β "(Γ, U, 𝒞)" "(Γ', U', 𝒞')" rule: scl_fol.decide.cases)
        case (decideI L γ)
        show False
          using Γ' = decide_lit (L ⋅l γ) # Γ
          using Γ' = propagate_lit (Pos A) {#L ∈# C. L  Pos A#} Var # Γ
          unfolding decide_lit_def propagate_lit_def
          by blast
      qed
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (conflict CG)

    have "Γ' = Γ"
      unfolding Γ'_def ΓG' = ΓG Γ_def ..

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    have
      CG_in: "CG |∈| iefac  |`| (NG |∪| UG)" and
      CG_false: "trail_false_cls ΓG CG" and
      CG_lt: "EG |∈| iefac  |`| (NG |∪| UG). EG  CG  trail_false_cls ΓG EG  CG c EG"
      using linorder_cls.is_least_in_fset _ CG
      unfolding atomize_conj linorder_cls.is_least_in_ffilter_iff by argo

    have "L. is_pos L  ord_res.is_maximal_lit L CG"
    proof (rule notI , elim exE conjE)
      fix L :: "'f gliteral"
      assume "is_pos L" and CG_max_lit: "ord_res.is_maximal_lit L CG"

      have "{#} |∉| iefac  |`| (NG |∪| UG)"
        using CG_lt
        by (metis (full_types) CG_max_lit bot_fset.rep_eq fBex_fempty linorder_cls.leD
            linorder_lit.is_maximal_in_mset_iff mempty_lesseq_cls set_mset_empty
            trail_false_cls_mempty)

      have "trail_false_lit ΓG L"
        using CG_max_lit CG_false
        unfolding linorder_lit.is_maximal_in_mset_iff trail_false_cls_def
        by metis

      then obtain Ln ΓG0 where "ΓG = Ln # ΓG0"
        unfolding trail_false_lit_def
        by (metis (no_types) List.insert_def image_iff insert_Nil neq_Nil_conv)

      moreover have
        AAA: "x xs. ΓG = x # xs 
          ((snd x  None)  fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls (x # xs))) 
          (snd x  None  is_pos (fst x)) 
          (x  set xs. snd x = None)" and
        BBB: "(Γ1 Ln Γ0. ΓG = Γ1 @ Ln # Γ0  snd Ln = None 
          ¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls (Ln # Γ0)))" and
        ΓG_sorted: "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) ΓG"
        using ord_res_11_invars NG SG[unfolded SG_def ord_res_11_invars.simps
            ord_res_10_invars.simps]
        by simp_all

      moreover have "fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG)"
        using CG_in CG_false by metis

      ultimately have "snd Ln  None" and "is_pos (fst Ln)" and "x  set ΓG0. snd x = None"
        unfolding atomize_conj by metis

      have "¬ fBex (iefac  |`| (NG |∪| UG)) (trail_false_cls ΓG0)"
      proof (cases ΓG0)
        case Nil
        then show ?thesis
          using {#} |∉| iefac  |`| (NG |∪| UG)
          unfolding trail_false_cls_def trail_false_lit_def
          by simp
      next
        case (Cons x xs)
        then show ?thesis
          using ΓG = Ln # ΓG0
          using x  set ΓG0. snd x = None
          using BBB[rule_format, of "[Ln]", unfolded append_Cons append_Nil]
          by simp
      qed

      hence "¬ trail_false_cls ΓG0 CG"
        using CG_in by metis

      hence "fst Ln = - L"
        using CG_false CG_max_lit ΓG_sorted[unfolded ΓG = Ln # ΓG0 sorted_wrt.simps]
        by (smt (verit, ccfv_SIG) Neg_atm_of_iff ΓG = Ln # ΓG0 atm_of_uminus
            ord_res.less_lit_simps(4) imageE image_insert insertE
            linorder_lit.dual_order.strict_trans linorder_lit.is_maximal_in_mset_iff
            linorder_lit.neq_iff linorder_trm.order.irrefl list.simps(15) literal.collapse(1)
            ord_res.ground_ordered_resolution_calculus_axioms trail_false_cls_def
            trail_false_lit_def)

      hence "¬ is_pos L"
        using is_pos (fst Ln)
        by (simp add: is_pos_neg_not_is_pos)

      thus False
        using is_pos L by contradiction
    qed

    hence "CG |∈| NG |∪| UG"
    proof -
      obtain C where "C |∈| NG |∪| UG" and "CG = iefac  C"
        using CG_in by blast

      hence "CG = C"
        using L. is_pos L  ord_res.is_maximal_lit L CG
        by (metis clauses_in_ℱ_have_pos_max_lit ex1_efac_eq_factoring_chain iefac_def
            ord_res.ground_factorings_preserves_maximal_literal)

      thus ?thesis
        using C |∈| NG |∪| UG by simp
    qed

    have "scl_fol.conflict N β S S'"
      unfolding S_def S'_def Γ' = Γ U' = U 𝒞_def 𝒞'_def 𝒞G = None 𝒞G' = Some CG option.map
    proof (rule scl_fol.conflictI)
      show "cls_of_gcls CG |∈| N |∪| U"
        unfolding N_def U_def
        using CG |∈| NG |∪| UG by auto
    next
      show "is_ground_cls (cls_of_gcls CG  (Var::'v  ('f, _) Term.term))"
        by simp
    next
      have "trail_false_cls ΓG CG"
        using trail_false_cls ΓG CG .

      hence "trail_false_cls Γ (cls_of_gcls CG  Var)"
        unfolding Γ_def subst_cls_id_subst
        using trail_false_cls_nonground_iff_trail_false_cls_ground by metis

      thus "SCL_FOL.trail_false_cls Γ (cls_of_gcls CG  Var)"
        unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
        unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
        by argo
    qed

    thus ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (skip LG CG nG)

    have "Γ = gtrailelem_of_trailelem (LG, nG) # Γ'"
      unfolding Γ_def Γ'_def ΓG = (LG, nG) # ΓG' by simp

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    have "¬ Ex (scl_fol.conflict N β S)"
      using 𝒞G = Some CG nex_conflict_if_alread_in_conflict by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.skip N β S S'"
        unfolding S_def S'_def U' = U 𝒞_def 𝒞'_def 𝒞G = Some CG 𝒞G' = Some CG option.map
        unfolding Γ = _ # Γ' gtrailelem_of_trailelem_def prod.case
      proof (rule scl_fol.skipI)
        have "- lit_of_glit LG = lit_of_glit (- LG)"
          by (cases LG) (simp_all add: lit_of_glit_def)
        show "- lit_of_glit LG ∉# cls_of_gcls CG  Var"
          unfolding subst_cls_id_subst
          unfolding - lit_of_glit LG = lit_of_glit (- LG)
          unfolding cls_of_gcls_def
          using - LG ∉# CG inj_image_mset_mem_iff[OF inj_lit_of_glit]
          by metis
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.conflict N β S' S''"

      moreover have "S''. scl_fol.conflict N β S' S''"
        unfolding S'_def 𝒞'_def 𝒞G' = Some CG by (simp add: scl_fol.conflict.simps)

      ultimately show False
        by metis
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (resolution LG CG ΓG'' DG)

    have "Γ' = Γ"
      unfolding Γ'_def ΓG' = ΓG Γ_def ..

    have "U' = U"
      unfolding U'_def UG' = UG U_def ..

    have "𝒞 = Some (cls_of_gcls DG, Var)"
      unfolding 𝒞_def 𝒞G = Some DG option.map ..

    hence 𝒞_eq: "𝒞 = Some
      (add_mset (- lit_of_glit LG) (remove1_mset (- lit_of_glit LG) (cls_of_gcls DG)), Var)"
      by (smt (verit, best) add_mset_remove_trivial atm_of_eq_atm_of atm_of_lit_of_glit_conv
          cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset insert_DiffM step_hyps(7)
          uminus_not_id')

    have "𝒞' = Some (
      remove1_mset (- (lit_of_glit LG)) (cls_of_gcls DG) +
      remove1_mset (lit_of_glit LG) (cls_of_gcls CG), Var)"
      unfolding 𝒞'_def 𝒞G' = Some _ option.map
      apply (simp add: cls_of_gcls_def)
      by (smt (verit, ccfv_threshold) add_diff_cancel_right' add_mset_add_single atm_of_eq_atm_of
          atm_of_lit_of_glit_conv diff_single_trivial glit_of_lit_lit_of_glit
          image_mset_remove1_mset_if insert_DiffM is_pos_neg_not_is_pos msed_map_invR)
    hence 𝒞'_eq: "𝒞' = Some (
      (remove1_mset (- (lit_of_glit LG)) (cls_of_gcls DG)  Var +
      remove1_mset (lit_of_glit LG) (cls_of_gcls CG)  Var)  Var, Var)"
      by simp

    have "linorder_lit.is_greatest_in_mset CG LG"
      using ord_res_11_invars NG SG[unfolded SG_def ΓG = (LG, Some CG) # ΓG'']
      unfolding ord_res_11_invars.simps ord_res_10_invars.simps
      by simp

    hence "add_mset LG {#y ∈# CG. y  LG#} = CG"
      using linorder_lit.explode_greatest_in_mset by metis

    hence "CG - {#LG#} = {#K ∈# CG. K  LG#}"
      by (metis add_mset_remove_trivial)

    hence "cls_of_gcls (CG - {#LG#}) = cls_of_gcls {#K ∈# CG. K  LG#}"
      by argo

    moreover have "cls_of_gcls (CG - {#LG#}) = cls_of_gcls CG - cls_of_gcls {#LG#}"
      unfolding cls_of_gcls_def
    proof (rule image_mset_Diff)
      show "{#LG#} ⊆# CG"
        by (metis ord_res.is_strictly_maximal_lit LG CG linorder_lit.is_greatest_in_mset_iff
            single_subset_iff)
    qed

    ultimately have "cls_of_gcls CG - {#lit_of_glit LG#} = cls_of_gcls {#K ∈# CG. K  LG#}"
      by (metis add_mset LG {#y ∈# CG. y  LG#} = CG cls_of_gcls_def image_mset_remove1_mset_if
          union_single_eq_member)

    have "¬ Ex (scl_fol.conflict N β S)"
      using 𝒞G = Some _ nex_conflict_if_alread_in_conflict by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.resolve N β S S'"
        unfolding S_def S'_def Γ' = Γ U' = U
        unfolding 𝒞_eq 𝒞'_eq
      proof (rule scl_fol.resolveI)
        show "Γ = trail_propagate (map gtrailelem_of_trailelem ΓG'')
          (lit_of_glit LG) (remove1_mset (lit_of_glit LG) (cls_of_gcls CG)) Var"
          unfolding Γ_def ΓG = (LG, Some CG) # ΓG'' gtrailelem_of_trailelem_def list.map prod.case
          unfolding propagate_lit_def subst_lit_id_subst option.map
          unfolding remove1_mset (lit_of_glit LG) (cls_of_gcls CG) = cls_of_gcls {#K ∈# CG. K  LG#}
          by argo
      next
        show "lit_of_glit LG ⋅l Var = - (- lit_of_glit LG ⋅l Var)"
          by simp
      next
        show "SCL_FOL.is_renaming Var"
          by simp
      next
        show "SCL_FOL.is_renaming Var"
          by simp
      next
        show "
          vars_cls (add_mset (- lit_of_glit LG)
            (remove1_mset (- lit_of_glit LG) (cls_of_gcls DG))  Var) 
          vars_cls (add_mset (lit_of_glit LG)
            (remove1_mset (lit_of_glit LG) (cls_of_gcls CG))  Var) =
          {}"
          by (metis (no_types, lifting) boolean_algebra.conj_zero_right cls_of_gcls_def
              diff_single_trivial image_mset_add_mset insert_DiffM subst_cls_id_subst
              vars_cls_cls_of_gcls)
      next
        show "SCL_FOL.is_imgu Var
          {{atm_of (- lit_of_glit LG) ⋅a Var, atm_of (lit_of_glit LG) ⋅a Var}}"
          by (simp add: SCL_FOL.is_imgu_def SCL_FOL.is_unifiers_def SCL_FOL.is_unifier_def)
      next
        show "is_grounding_merge Var
          (vars_cls
            (add_mset (- lit_of_glit LG) (remove1_mset (- lit_of_glit LG) (cls_of_gcls DG))  Var))
          (rename_subst_domain Var Var)
          (vars_cls
            (add_mset (lit_of_glit LG) (remove1_mset (lit_of_glit LG) (cls_of_gcls CG))  Var))
          (rename_subst_domain Var Var)"
          by (simp add: is_grounding_merge_def)
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.conflict N β S' S''"

      moreover have "S''. scl_fol.conflict N β S' S''"
        unfolding S'_def 𝒞'_def 𝒞G' = Some _ by (simp add: scl_fol.conflict.simps)

      ultimately show False
        by metis
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  next
    case step_hyps: (backtrack LG CG)

    define K :: "('f, 'v) term literal" where
      "K = - lit_of_glit LG"

    define D :: "('f, 'v) term clause" where
      "D = cls_of_gcls CG - {#K#}"

    have "add_mset K D = cls_of_gcls CG"
      by (smt (verit, best) D_def K_def add_mset_remove_trivial atm_of_eq_atm_of
          atm_of_lit_of_glit_conv cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset
          insert_DiffM step_hyps(6) uminus_not_id')

    have "U' = finsert (add_mset K D) U"
      unfolding U_def U'_def UG' = finsert CG UG
      by (smt (verit, ccfv_SIG) D_def K_def add_mset_remove_trivial atm_of_eq_atm_of
          atm_of_lit_of_glit_conv cls_of_gcls_def fimage_finsert glit_of_lit_lit_of_glit
          image_mset_add_mset insert_DiffM step_hyps(6) uminus_not_id')

    have "𝒞 = Some (add_mset K D, Var)"
      by (smt (verit) D_def K_def 𝒞_def add_mset_remove_trivial atm_of_eq_atm_of
          atm_of_lit_of_glit_conv cls_of_gcls_def glit_of_lit_lit_of_glit image_mset_add_mset
          insert_DiffM option.map(2) step_hyps(1,6) uminus_not_id')

    have "𝒞' = None"
      unfolding 𝒞'_def 𝒞G' = None option.map ..

    have "¬ Ex (scl_fol.conflict N β S)"
      using 𝒞G = Some _ nex_conflict_if_alread_in_conflict by metis

    moreover have "scl_fol.reasonable_scl N β S S'"
      unfolding scl_fol.reasonable_scl_def
    proof (intro conjI impI notI ; (elim exE) ?)
      have "scl_fol.backtrack N β S S'"
        unfolding S_def S'_def
        unfolding U' = finsert (add_mset K D) U 𝒞 = Some (add_mset K D, Var) 𝒞' = None
      proof (rule scl_fol.backtrackI)
        show "Γ = trail_decide ([] @ Γ') (lit_of_glit LG)"
          unfolding append_Nil
          unfolding decide_lit_def
          unfolding Γ_def ΓG = _ # _ list.map Γ'_def[symmetric]
          unfolding gtrailelem_of_trailelem_def prod.case option.map
          ..
      next
        show "lit_of_glit LG = - (K ⋅l Var)"
          unfolding K_def by simp
      next
        have "sorted_wrt (λx y. atm_of (fst y) t atm_of (fst x)) ΓG"
          using ord_res_11_invars NG SG[unfolded SG_def]
          unfolding ord_res_11_invars.simps ord_res_10_invars.simps
          by fast

        hence "trail_consistent ΓG"
          using trail_consistent_if_sorted_wrt_atoms by metis

        hence "¬ trail_defined_lit ΓG' (- LG)"
          by (metis trail_consistent.cases atm_of_eq_atm_of list.distinct(1) list.inject
              prod.sel(1) step_hyps(5) trail_defined_lit_iff_trail_defined_atm)

        hence "¬ trail_false_lit ΓG' (- LG)"
          using trail_defined_lit_iff_true_or_false by metis

        hence "¬ trail_false_cls ΓG' CG"
          using - LG ∈# CG
          unfolding trail_false_cls_def by metis

        hence "¬ trail_false_cls Γ' (add_mset K D)"
          unfolding Γ'_def add_mset K D = cls_of_gcls CG
          unfolding trail_false_cls_nonground_iff_trail_false_cls_ground .

        moreover have "is_ground_cls (add_mset K D)"
          using 𝒞_def 𝒞 = Some (add_mset K D, Var) step_hyps(1) by auto

        ultimately have "γ. is_ground_cls (add_mset K D  γ)  trail_false_cls Γ' (add_mset K D  γ)"
          by simp

        thus "γ. is_ground_cls (add_mset K D  γ)  SCL_FOL.trail_false_cls Γ' (add_mset K D  γ)"
          unfolding SCL_FOL.trail_false_cls_def trail_false_cls_def
          unfolding SCL_FOL.trail_false_lit_def trail_false_lit_def
          by argo
      qed

      thus "scl_fol.scl N β S S'"
        unfolding scl_fol.scl_def by argo
    next
      fix S'' :: "('f, 'v) SCL_FOL.state"
      assume "scl_fol.decide N β S S'"
      thus False
        unfolding S_def 𝒞 = Some _
        by (auto elim!: scl_fol.decide.cases)
    qed

    ultimately show ?thesis
      unfolding scl_fol.regular_scl_def by argo
  qed
qed

end

lemma wfp_on_antimono_stronger:
  fixes
    A :: "'a set" and B :: "'b set" and
    f :: "'a  'b" and
    R :: "'b  'b  bool" and Q :: "'a  'a  bool"
  assumes
    wf: "wfp_on B R" and
    sub: "f ` A  B" and
    mono: "x y. x  A  y  A  Q x y  R (f x) (f y)"
  shows "wfp_on A Q"
  unfolding wfp_on_iff_ex_minimal
proof (intro allI impI)
  fix A' :: "'a set"
  assume "A'  A" and "A'  {}"
  have "f ` A'  B"
    using A'  A sub by blast
  moreover have "f ` A'  {}"
    using A'  {} by blast
  ultimately have "zf ` A'. y. R y z  y  f ` A'"
    using wf wfp_on_iff_ex_minimal by blast
  hence "zA'. y. R (f y) (f z)  y  A'"
    by blast
  thus "zA'. y. Q y z  y  A'"
    using A'  A mono by blast
qed

text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.wfp_on_antimono_stronger}
as it is available in theoryHOL.Wellfounded.›

corollary (in scl_fol_calculus) termination_projectable_strategy:
  fixes
    N :: "('f, 'v) Term.term clause fset" and
    β :: "('f, 'v) Term.term" and
    strategy and strategy_init and proj
  assumes strategy_restricts_regular_scl:
    "S S'. strategy** strategy_init S  strategy S S'  regular_scl N β (proj S) (proj S')" and
    initial_state: "proj strategy_init = initial_state"
  shows "wfp_on {S. strategy** strategy_init S} strategy¯¯"
proof (rule wfp_on_antimono_stronger)
  show "wfp_on {proj S | S. strategy** strategy_init S} (regular_scl N β)¯¯"
  proof (rule wfp_on_subset)
    show "wfp_on {S. (regular_scl N β)** initial_state S} (regular_scl N β)¯¯"
      using termination_regular_scl by metis
  next
    show "{proj S | S. strategy** strategy_init S}  {S. (regular_scl N β)** initial_state S}"
    proof (intro Collect_mono impI, elim exE conjE)
      fix s S assume "s = proj S" and "strategy** strategy_init S"
      show "(regular_scl N β)** initial_state s"
        unfolding s = proj S
        using strategy** strategy_init S
      proof (induction S rule: rtranclp_induct)
        case base
        thus ?case
          unfolding initial_state by simp
      next
        case (step y z)
        thus ?case
          using strategy_restricts_regular_scl
          by (meson rtranclp.simps)
      qed
    qed
  qed
next
  show "proj ` {S. strategy** strategy_init S}  {proj S |S. strategy** strategy_init S}"
    by blast
next
  show "S' S. S  {S. strategy** strategy_init S}  strategy¯¯ S' S 
    (regular_scl N β)¯¯ (proj S') (proj S)"
    using strategy_restricts_regular_scl by simp
qed

text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.scl_fol_calculus.termination_projectable_strategy}
as it is available in theorySimple_Clause_Learning.Termination.›

corollary (in simulation_SCLFOL_ground_ordered_resolution) ord_res_11_termination:
  fixes N :: "'f gclause fset"
  shows "wfp_on {S. (ord_res_11 N)** ({||}, {||}, [], None) S} (ord_res_11 N)¯¯"
proof (rule scl_fol.termination_projectable_strategy)
  fix S S'
  assume run: "(ord_res_11 N)** ({||}, {||}, [], None) S" and step: "ord_res_11 N S S'"

  define β :: "'f gterm" where
    "β = (THE A. linorder_trm.is_greatest_in_fset (atms_of_clss N) A)"

  show "scl_fol.regular_scl (cls_of_gcls |`| N) (term_of_gterm β) (state_of_gstate S) (state_of_gstate S')"
  proof (rule ord_res_11_is_strategy_for_regular_scl)
    show "AG |∈| atms_of_clss N. AG t β"
    proof (cases "atms_of_clss N = {||}")
      case True
      thus ?thesis
        by simp
    next
      case False
      then show ?thesis
        unfolding β_def
        by (metis (full_types) linorder_trm.Uniq_is_greatest_in_fset
            linorder_trm.ex_greatest_in_fset linorder_trm.is_greatest_in_fset_iff sup2CI
            the1_equality')
    qed
  next
    show "(ord_res_11 N)** ({||}, {||}, [], None) S"
      using run .
  next
    show "ord_res_11 N S S'"
      using step .
  qed
next
  show "state_of_gstate ({||}, {||}, [], None) = SCL_FOL.initial_state"
    by simp
qed

corollary (in scl_fol_calculus) static_non_subsumption_projectable_strategy:
  fixes strategy and strategy_init and proj
  assumes
    run: "strategy** strategy_init S" and
    step: "backtrack N β (proj S) S'" and
    strategy_restricts_regular_scl:
      "S S'. strategy** strategy_init S  strategy S S'  regular_scl N β (proj S) (proj S')" and
    initial_state: "proj strategy_init = initial_state"
  defines
    "U  state_learned (proj S)"
  shows "C γ. state_conflict (proj S) = Some (C, γ)  ¬ (D |∈| N |∪| U. subsumes D C)"
proof -
  have "(regular_scl N β)** initial_state (proj S)"
    using run
  proof (induction S rule: rtranclp_induct)
    case base
    thus ?case
      unfolding initial_state by simp
  next
    case (step y z)
    thus ?case
      using strategy_restricts_regular_scl
      by (meson rtranclp.simps)
  qed

  moreover have "backtrack N β (proj S) S'"
    using step by simp

  ultimately show ?thesis
    unfolding U_def
    using static_non_subsumption_regular_scl
    by simp
qed

text ‹For AFP-devel, delete
@{thm Simulation_SCLFOL_ORDRES.scl_fol_calculus.static_non_subsumption_projectable_strategy}
as it is available in theorySimple_Clause_Learning.Non_Redundancy.›

corollary (in simulation_SCLFOL_ground_ordered_resolution) ord_res_11_non_subsumption:
  fixes NG :: "'f gclause fset" and s :: "_ × _ × _ × _"
  defines
    "β  (THE A. linorder_trm.is_greatest_in_fset (atms_of_clss NG) A)"
  assumes
    run: "(ord_res_11 NG)** ({||}, {||}, [], None) s" and
    step: "scl_fol.backtrack (cls_of_gcls |`| NG) (term_of_gterm β) (state_of_gstate s) s'"
  shows "Uer  Γ D. s = (Uer, , Γ, Some D)  ¬ (C |∈| NG |∪| Uer. C ⊆# D)"
proof -
  have "C γ. state_conflict (state_of_gstate s) = Some (C, γ) 
        ¬ (D|∈| cls_of_gcls |`| NG |∪| state_learned (state_of_gstate s). subsumes D C)"
  proof (rule scl_fol.static_non_subsumption_projectable_strategy[
        of "ord_res_11 NG" _ _ _ _ state_of_gstate , OF run step])
    fix S S'
    assume run: "(ord_res_11 NG)** ({||}, {||}, [], None) S" and step: "ord_res_11 NG S S'"
    show "scl_fol.regular_scl (cls_of_gcls |`| NG) (term_of_gterm β)
      (state_of_gstate S) (state_of_gstate S')"
    proof (intro ord_res_11_is_strategy_for_regular_scl ballI)
      fix AG :: "'f gterm"
      assume "AG |∈| atms_of_clss NG"
      show "AG t β"
      proof (cases "atms_of_clss NG = {||}")
        case True
        thus ?thesis
          using AG |∈| atms_of_clss NG
          by simp
      next
        case False
        then show ?thesis
          using AG |∈| atms_of_clss NG
          unfolding β_def
          by (metis (full_types) linorder_trm.Uniq_is_greatest_in_fset
              linorder_trm.ex_greatest_in_fset linorder_trm.is_greatest_in_fset_iff sup2CI
              the1_equality')
      qed
    next
      show "(ord_res_11 NG)** ({||}, {||}, [], None) S"
        using run .
    next
      show "ord_res_11 NG S S'"
        using step .
    qed
  next
    show "state_of_gstate ({||}, {||}, [], None) = initial_state"
      by simp
  qed

  moreover obtain UG  Γ D where "s = (UG, , Γ, Some D)"
  proof atomize_elim
    obtain UG  Γ 𝒞 where "s = (UG, , Γ, 𝒞)"
      by (metis prod.exhaust)

    moreover obtain D where "𝒞 = Some D"
      using step
      by (auto simp: s = (UG, , Γ, 𝒞) elim: scl_fol.backtrack.cases)

    ultimately show "Uer  Γ D. s = (Uer, , Γ, Some D)"
      by metis
  qed

  ultimately have "¬ (C |∈| NG |∪|UG.
    subsumes (cls_of_gcls C :: ('f, 'v) term clause) (cls_of_gcls D))"
    by auto

  hence "¬ (C |∈| NG |∪|UG. (cls_of_gcls C :: ('f, 'v) term clause) ⊆# (cls_of_gcls D))"
    by (simp add: subsumes_def)

  hence "¬ (C |∈| NG |∪|UG. C ⊆# D)"
    by (metis cls_of_gcls_def image_mset_subseteq_mono)

  thus ?thesis
    unfolding s = (UG, , Γ, Some D) by metis
qed

end