Theory Ordered_Completion

theory Ordered_Completion
imports Ordered_Rewriting Normalization_Equivalence KBO_More Abstract_Completion Prime_Critical_Pairs Q_Restricted_Rewriting Exact_Tree_Automata_Completion
(*
Author:  Christian Sternagel <c.sternagel@gmail.com> (2016, 2017)
Author:  Sarah Winkler <sarah.winkler@uibk.ac.at> (2017, 2018)
License: LGPL (see file COPYING.LESSER)
*)

section ‹Ordered Completion›

theory Ordered_Completion
  imports
    "../Rewriting/Ordered_Rewriting"
    QTRS.Encompassment
    "../Normalization_Equivalence/Normalization_Equivalence"
    QTRS.Lexicographic_Extension
    "../Orderings/KBO_More"
    "../Auxiliaries/Multiset_Extension2"
    "../Abstract_Completion/Abstract_Completion"
    "../Abstract_Completion/Prime_Critical_Pairs"
    QTRS.Q_Restricted_Rewriting
    "../Tree_Automata/Exact_Tree_Automata_Completion" (* for stuff on linearity *)
begin

locale ordered_completion = reduction_order
begin

definition "ostep E R = rstep R ∪ ordstep {≻} (E)"

lemma ostep_iff_ordstep:
  assumes "R ⊆ {≻}"
  shows "ostep E R = ordstep {≻} (R ∪ E)"
  using subst_closed_less and assms by (simp add: ordstep_Un ostep_def ordstep_rstep_conv)

lemma ostep_cases [consumes 1, cases pred: ostep]:
  assumes "(s, t) ∈ ostep E R"
  obtains (rstep) l r C σ where "(l, r) ∈ R" and "s = C⟨l ⋅ σ⟩" and "t = C⟨r ⋅ σ⟩"
  | (ordstep) l r C σ where "(l, r) ∈ E" and "l ⋅ σ ≻ r ⋅ σ" and "s = C⟨l ⋅ σ⟩" and "t = C⟨r ⋅ σ⟩"
  using assms by (fastforce simp: ostep_def elim: ordstep.cases)

lemma ostep_mono:
  "E ⊆ E' ⟹ R ⊆ R'⟹ ostep E R ⊆ ostep E' R'"
  by (auto simp: ostep_def rstep_mono [THEN subsetD]
      intro: ordstep_mono [THEN subsetD, OF _ subset_refl, of "E" "E'"])

lemma rstep_imp_ostep: "(s, t) ∈ rstep R ⟹ (s, t) ∈ ostep E R" by (auto simp: ostep_def)

lemma ostep_imp_rstep_sym:
  "(s, t) ∈ ostep E R ⟹ (s, t) ∈ rstep (R ∪ E)"
  unfolding ostep_def using ordstep_imp_rstep by fast

lemma E_imp_ostep:
  "l ⋅ σ ≻ r ⋅ σ ⟹ (l, r) ∈ E ⟹ s = C⟨l ⋅ σ⟩ ⟹ t = C⟨r ⋅ σ⟩ ⟹ (s, t) ∈ ostep E R"
  by (auto simp: ostep_def intro: ordstep.intros)

lemma ostep_subst:
  "(s, t) ∈ ostep E R ⟹ (s ⋅ σ, t ⋅ σ) ∈ ostep E R"
  using ordstep_subst [of "{≻}", OF subst_closed_less] and rstep.subst
  by (auto simp: ostep_def)

lemma ostep_ctxt:
  "(s, t) ∈ ostep E R ⟹ (C⟨s⟩, C⟨t⟩) ∈ ostep E R"
  using ordstep_ctxt [where ord = "{≻}"] and rstep.ctxt
  by (auto simp: ostep_def)

lemma ostep_imp_less:
  "R ⊆ {≻} ⟹ (s, t) ∈ ostep E R ⟹ s ≻ t"
  by (force simp: ostep_def subst ctxt elim: ordstep.cases)

inductive
  oKB :: "('a, 'b) trs × ('a, 'b) trs ⇒ ('a, 'b) trs × ('a, 'b) trs ⇒ bool" (infix "⊢oKB" 55)
  where
    deduce: "(s, t) ∈ rstep (R ∪ E) ⟹ (s, u) ∈ rstep (R ∪ E) ⟹ (E, R) ⊢oKB (E ∪ {(t, u)}, R)" |
    orientl: "s ≻ t ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, t)}, R ∪ {(s, t)})" |
    orientr: "t ≻ s ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, t)}, R ∪ {(t, s)})" |
    delete: "(s, s) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, s)}, R)" |
    compose: "(t, u) ∈ ostep E (R - {(s, t)}) ⟹ (s, t) ∈ R ⟹ (E, R) ⊢oKB (E, (R - {(s, t)}) ∪ {(s, u)})" |
    simplifyl: "(s, u) ∈ ostep (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB ((E - {(s, t)}) ∪ {(u, t)}, R)" |
    simplifyr: "(t, u) ∈ ostep (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB ((E - {(s, t)}) ∪ {(s, u)}, R)" |
    collapse: "(t, u) ∈ ostep E (R - {(t, s)}) ⟹ (t, s) ∈ R ⟹ (E, R) ⊢oKB (E ∪ {(u, s)}, R - {(t, s)})"

lemma oKB_less:
  assumes "(E, R) ⊢oKB (E', R')" and "R ⊆ {≻}"
  shows "R' ⊆ {≻}"
proof -
  have [dest]: "(x, y) ∈ R ⟹ x ≻ y" for x y using ‹R ⊆ {≻}› by auto
  note assms(1)
  moreover
  { fix s t u assume "(t, u) ∈ rstep (R - {(s, t)})" and "(s, t) ∈ R"
    then have "s ≻ u" by (induct t u) (blast intro: subst ctxt dest: trans) }
  moreover
  { fix s t u assume "(t, u) ∈ ordstep {≻} (E)" and "(s, t) ∈ R"
    then have "s ≻ u" by (cases) (auto dest: trans ctxt) }
  ultimately show ?thesis by (cases) (fastforce simp: ostep_def)+
qed

lemma oKB_rtrancl_less:
  assumes "oKB** (E, R) (E', R')" and "R ⊆ {≻}"
  shows "R' ⊆ {≻}"
  using assms by (induct "(E, R)" "(E', R')" arbitrary: E' R') (auto dest!: oKB_less)

lemma oKB_E:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "E ⊆ Id ∪ E' ∪ R' ∪ (ostep E' R') O E' ∪ E' O (ostep E' R')¯"
  using assms
proof (cases)
  case (simplifyl s u t)
  then have "(s, u) ∈ ostep E' R'"
    by (intro ostep_mono [THEN subsetD, OF _ _ simplifyl(3)]) auto
  then have "(s, t) ∈ (ostep E' R') O E'" using simplifyl by auto
  then show ?thesis using simplifyl by auto
next
  case (simplifyr t u s)
  then have "(t, u) ∈ ostep E' R'"
    by (intro ostep_mono [THEN subsetD, OF _ _ simplifyr(3)]) auto
  then have "(s, t) ∈ E' O (ostep E' R')¯" using simplifyr by blast
  then show ?thesis using simplifyr by blast
qed auto

lemma oKB_E_rev:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "E' ⊆ E ∪ (rstep (R ∪ E))¯ O (rstep (R ∪ E)) ∪ (ostep E R)¯ O E ∪ E O (ostep E R)"
  using assms
proof (cases)
  case (simplifyl s u t)
  have "(s, u) ∈ ostep E R"  by (intro ostep_mono [THEN subsetD, OF _ _ simplifyl(3)]) auto
  with simplifyl(4) have "(u, t) ∈ (ostep E R)¯ O E" by auto
  then show ?thesis using simplifyl by auto
next
  case (simplifyr t u s)
  have "(t, u) ∈ ostep E R"
    by (intro ostep_mono [THEN subsetD, OF _ _ simplifyr(3)]) auto
  with simplifyr(4) have "(s, u) ∈ E O (ostep E R)" by auto
  then show ?thesis using simplifyr by auto
next
  case (deduce u s t)
  then have "(s, t) ∈ (rstep (R ∪ E))¯ O (rstep (R ∪ E))" by blast
  with deduce show ?thesis by auto
next
  case (collapse t u s)
  have "(t, u) ∈ ostep E R"
    by (intro ostep_mono [THEN subsetD, OF _ _ collapse(3)]) auto
  from ostep_imp_rstep_sym [OF this] subsetD [OF rstep_mono, OF Un_upper1 rstep.rule [OF collapse(4)]]
  have "(u, s) ∈ (rstep (R ∪ E))¯ O (rstep (R ∪ E))" by blast
  with collapse show ?thesis by auto
qed auto

lemma oKB_R:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "R ⊆ R' ∪ R' O (ostep E' R')¯ ∪ ostep E' R' O E'"
  using assms
proof (cases)
  case (compose t u s)
  then have "(t, u) ∈ ostep E' R'"
    by (intro ostep_mono [THEN subsetD, OF _ _ compose(3)]) auto
  then have "(s, t) ∈ R' O (ostep E' R')¯" using compose by blast
  then show ?thesis using compose by blast
next
  case (collapse t u s)
  then have "(t, u) ∈ ostep E' R'"
    by (intro ostep_mono [THEN subsetD, OF _ _ collapse(3)]) auto
  then have "(t, s) ∈ ostep E' R' O E'" using collapse by auto
  then show ?thesis using collapse by blast
qed auto

lemma oKB_R_rev:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "R' ⊆ R ∪ E ∪ R O (ostep E R)"
  using assms
proof (cases)
  case (compose s t u)
  have "(s, t) ∈ ostep E R"
    by (intro ostep_mono [THEN subsetD, OF _ _ compose(3)]) auto
  with compose show ?thesis by fast
qed auto

lemma conversion2:
  "(x, y) ∈ r ⟹ (y, z) ∈ r ⟹ (x, z) ∈ r*"
  by (auto simp: conversion_def intro: rtrancl_trans)

lemma oKB_rstep_subset:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "rstep (E ∪ R) ⊆ (rstep (E' ∪ R'))*"
  using assms
  apply (auto elim!: rstepE dest!: oKB_E [OF oKB, THEN subsetD] oKB_R [OF oKB, THEN subsetD])
     apply (auto  dest!: ostep_imp_rstep_sym)
     apply (rule_tac y = "C⟨y ⋅ σ⟩" in conversion2; fast)+
  done

lemma oKB_conversion_subset:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "(rstep (E ∪ R))* ⊆ (rstep (E' ∪ R'))*"
  using conversion_mono [OF oKB_rstep_subset [OF assms]] by simp

lemma oKB_rtrancl_conversion_subset:
  assumes "oKB** (E, R) (E', R')"
  shows "(rstep (E ∪ R))* ⊆ (rstep (E' ∪ R'))*"
  using assms
  by (induct "(E, R)" "(E', R')" arbitrary: E' R')
    (force dest!: oKB_conversion_subset)+

lemma oKB_rstep_subset_rev:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "rstep (E' ∪ R') ⊆ (rstep (E ∪ R))*"
  using assms
  apply (auto elim!: rstepE)
   apply (auto dest!: oKB_R_rev [OF oKB, THEN subsetD] oKB_E_rev [OF oKB, THEN subsetD])
     apply (auto dest!: ostep_imp_rstep_sym)
     apply (rule conversion2; fast)+
  done

lemma oKB_conversion_supset:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "(rstep (E' ∪ R'))* ⊆ (rstep (E ∪ R))*"
  using conversion_mono [OF oKB_rstep_subset_rev [OF assms]] by simp


section ‹Abstract ordered completion according to proof for FSCD'17›

inductive_set encstep1 for E :: "('a, 'b) trs" and R :: "('a, 'b) trs"
  where
    estep: "(l, r) ∈ E ⟹ s = C⟨l ⋅ σ⟩ ⟹ t = C⟨r ⋅ σ⟩ ⟹ l ⋅ σ ≻ r ⋅ σ ⟹ l ⊲⋅ s ⟹ (s, t) ∈ encstep1 E R"
  | rstep: "(s, t) ∈ rstep R ⟹ (s, t) ∈ encstep1 E R"

inductive_set encstep2 for E :: "('a, 'b) trs" and R :: "('a, 'b) trs"
  where
    estep: "(l, r) ∈ E ⟹ s = C⟨l ⋅ σ⟩ ⟹ t = C⟨r ⋅ σ⟩ ⟹ l ⋅ σ ≻ r ⋅ σ ⟹ l ⊲⋅ s ⟹ (s, t) ∈ encstep2 E R"
  | rstep: "(l, r) ∈ R ⟹ s = C⟨l ⋅ σ⟩ ⟹ t = C⟨r ⋅ σ⟩ ⟹ l ⊲⋅ s ⟹ (s, t) ∈ encstep2 E R"

lemma encstep1_less:
  assumes "(s, t) ∈ encstep1 E R" and "R ⊆ {≻}"
  shows "s ≻ t"
  using assms(1)
by (cases, insert compatible_rstep_imp_less [OF assms(2)] ctxt, blast+)

inductive
  oKBi :: "('a, 'b) trs × ('a, 'b) trs ⇒ ('a, 'b) trs × ('a, 'b) trs ⇒ bool" (infix "⊢oKB" 55)
  where
    deduce: "(s, t) ∈ rstep (R ∪ E) ⟹ (s, u) ∈ rstep (R ∪ E) ⟹ (E, R) ⊢oKB (E ∪ {(t, u)}, R)" |
    orientl: "s ≻ t ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, t)}, R ∪ {(s, t)})" |
    orientr: "t ≻ s ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, t)}, R ∪ {(t, s)})" |
    delete: "(s, s) ∈ E ⟹ (E, R) ⊢oKB (E - {(s, s)}, R)" |
    compose: "(t, u) ∈ ostep E (R - {(s, t)}) ⟹ (s, t) ∈ R ⟹ (E, R) ⊢oKB (E, (R - {(s, t)}) ∪ {(s, u)})" |
    simplifyl: "(s, u) ∈ encstep1 (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB ((E - {(s, t)}) ∪ {(u, t)}, R)" |
    simplifyr: "(t, u) ∈ encstep1 (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKB ((E - {(s, t)}) ∪ {(s, u)}, R)" |
    collapse: "(t, u) ∈ encstep2 E (R - {(t, s)}) ⟹ (t, s) ∈ R ⟹ (E, R) ⊢oKB (E ∪ {(u, s)}, R - {(t, s)})"

lemma encstep2_encstep1:
  assumes "(s, t) ∈ encstep2 E R"
  shows "(s, t) ∈ encstep1 E R"
  unfolding encstep1.simps using assms by (cases) auto

lemma encstep1_ostep:
  assumes "(s, t) ∈ encstep1 E R"
  shows "(s, t) ∈ ostep E R"
  using assms
proof (cases)
  case (estep l r C σ)
  then show ?thesis unfolding ostep_def using ordstep.simps by fast
next
  case rstep
  then show ?thesis unfolding ostep_def rstep_def by fast
qed

lemma encstep2_ostep: "(s, t) ∈ encstep2 E R ⟹ (s, t) ∈ ostep E R"
  using encstep2_encstep1 encstep1_ostep by blast

lemma encstep1_mono:
  "E ⊆ E' ⟹ R ⊆ R'⟹ encstep1 E R ⊆ encstep1 E' R'"
  using rstep_mono [of R R']
  by (auto elim!: encstep1.cases intro: encstep1.intros)

lemma encstep2_mono:
  "E ⊆ E' ⟹ R ⊆ R'⟹ encstep2 E R ⊆ encstep2 E' R'"
  by (auto elim!: encstep2.cases intro: encstep2.intros)

lemma encstep1_imp_rstep_sym:
  "(s, t) ∈ encstep1 E R ⟹ (s, t) ∈ rstep (R ∪ E)"
  using encstep1_ostep ostep_imp_rstep_sym by force

lemma encstep2_imp_rstep_sym:
  "(s, t) ∈ encstep2 E R ⟹ (s, t) ∈ rstep (R ∪ E)"
  using encstep2_ostep ostep_imp_rstep_sym by force

lemma oKBi_oKB:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "(E, R) ⊢oKB (E', R')"
  using assms
proof (cases)
  case (simplifyl s u t)
  then show ?thesis using oKB.simplifyl and encstep1_ostep by blast
next
  case (simplifyr t u s)
  then show ?thesis using oKB.simplifyr and encstep1_ostep by blast
next
  case (collapse t u s)
  then show ?thesis using oKB.collapse encstep2_ostep [OF collapse(3)] by blast
qed (blast intro: oKB.intros)+

lemma oKBi_subset:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "E' ∪ R' ⊆ (rstep (E ∪ R))*"
  using oKBi_oKB [OF assms] and oKB_rstep_subset_rev subset_rstep [of "E' ∪ R'"] by blast

lemma oKBi_rstep_subset:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "rstep(E' ∪ R') ⊆ (rstep (E ∪ R))*"
  using oKBi_oKB [OF assms] and oKB_rstep_subset_rev subset_rstep [of "E' ∪ R'"] by blast

lemma oKBi_conversion_subset:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "(rstep (E' ∪ R'))* ⊆ (rstep (E ∪ R))*"
  using oKBi_rstep_subset [OF assms] and conversion_mono and conversion_conversion_idemp by metis

lemma oKBi_E_supset:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "E - E' ⊆ (encstep1 E' R' O E') ∪ R' ∪ Id"
  using assms by (cases) (use encstep1_mono [of "E - {(s, t)}" E' R R' for s t] in blast)+

lemma oKBi_E_supset':
  assumes "(E, R) ⊢oKB (E', R')"
  shows "E - E' ⊆ (encstep1 E' R' O E') ∪ (E' O (encstep1 E' R')¯) ∪ R' ∪ Id"
  using assms by (cases) (use encstep1_mono [of "E - {(s, t)}" E' R R' for s t] in blast)+

lemma oKBi_R_supset:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "R - R' ⊆ encstep2 E' R' O E' ∪ R' O (ostep E' R')¯"
  using assms
proof (cases)
  case (compose s t u)
  then have "(s, t) ∈ ostep E' R'"
    using ostep_mono [of E' E "R - {(u, s)}" R'] by auto
  with compose show ?thesis by fast
next
  case (collapse s t u)
  with encstep2_mono [of E "E ∪ {(t, u)}"] show ?thesis by fast
qed auto

lemma oKBi_supset:
  assumes step: "(E, R) ⊢oKB (E', R')"
  shows "E ∪ R ⊆ (rstep (E' ∪ R'))*"
  using assms subset_rstep encstep1_imp_rstep_sym encstep2_imp_rstep_sym
  by (smt Diff_partition oKB_rstep_subset oKBi_oKB ordered_completion_axioms sup.boundedE)

lemma oKBi_rstep_supset:
  assumes step: "(E, R) ⊢oKB (E', R')"
  shows "rstep(E ∪ R) ⊆ (rstep (E' ∪ R'))*"
  using oKBi_supset [OF assms]
  by (simp add: step oKB_rstep_subset oKBi_oKB ordered_completion_axioms)

lemma oKBi_conversion_supset:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "(rstep (E ∪ R))* ⊆ (rstep (E' ∪ R'))*"
  using oKBi_rstep_supset [OF assms] and conversion_mono and conversion_conversion_idemp by metis

lemma oKBi_conversion:
  assumes oKB: "(E, R) ⊢oKB (E', R')"
  shows "(rstep (E ∪ R))* = (rstep (E' ∪ R'))*"
  using assms and oKBi_conversion_supset and oKBi_conversion_subset by blast

lemma oKBi_less:
  assumes "R ⊆ {≻}" and "(E, R) ⊢oKB (E', R')"
  shows "R' ⊆ {≻}"
  using assms(1) and oKBi_oKB [OF assms(2)] and oKB_less by auto

lemma oKBi_rtrancl_less:
  assumes "oKBi** (E, R) (E', R')" and "R ⊆ {≻}"
  shows "R' ⊆ {≻}"
  using assms oKB_less
proof (induct "(E, R)" "(E', R')" arbitrary: E' R')
  case rtrancl_refl
  then show ?case by fast
next
  case (rtrancl_into_rtrancl S)
  obtain E2 R2 where S: "S = (E2, R2)" by force
  from rtrancl_into_rtrancl show ?case using oKBi_less unfolding S by metis
qed

end

locale okb_irun = ordered_completion +
  fixes R E
  assumes R0: "R 0 = {}"
    and irun: "⋀i. (E i, R i) ⊢oKB (E (Suc i), R (Suc i))"
begin

sublocale kb less enc ..

abbreviation Rinf ("R") where "R ≡ (⋃i. R i)"
abbreviation Einf ("E") where "E ≡ (⋃i. E i)"

definition "Rω = (⋃i. ⋂j∈{j. j≥i}. R j)"
definition "Eω = (⋃i. ⋂j∈{j. j≥i}. E j)"

lemma Rw_subset_Rinf: "Rω ⊆ R"
  by (auto simp: Rω_def)

lemma Ew_subset_Einf: "Eω ⊆ E"
  by (auto simp: Eω_def)

lemma oKBi_rtrancl_i: "oKBi** (E 0, R 0) (E i, R i)"
  using irun by (induct i) (auto intro: rtranclp.rtrancl_into_rtrancl)

lemma oKBi_conversion_ERi: "(rstep (E 0))* = (rstep (E i ∪ R i))*"
  using irun and oKBi_conversion and R0 by (induct i) auto 

lemma oKBi_conversion_ERinf: "(rstep (E 0))* = (rstep (E ∪ R))*"
proof
  show "(rstep (E 0))* ⊆ (rstep (E ∪ R))*"
    using oKBi_conversion_ERi conversion_mono unfolding rstep_union
    by (metis (no_types, lifting) Sup_upper range_eqI rstep_mono subset_trans sup.cobounded1)
next
  have "rstep (E ∪ R) ⊆ (rstep (E 0))*"
  proof
    fix s t
    assume "(s,t) ∈ rstep (E ∪ R)"
    then obtain i where "(s,t) ∈ rstep (E i ∪ R i)" by fast
    with oKBi_conversion_ERi show "(s,t) ∈ (rstep (E 0))*" by auto
  qed
  thus "(rstep (E ∪ R))* ⊆ (rstep (E 0))*" unfolding conversion_def
    by (metis converse_mono conversion_converse conversion_def le_supI rtrancl_subset_rtrancl)
qed

lemma Ri_less: "R i ⊆ {≻}"
  using oKBi_rtrancl_less [OF oKBi_rtrancl_i] and R0 by auto

lemma Rinf_less: "R ⊆ {≻}"
  using Ri_less by auto

lemma Rw_less: "Rω ⊆ {≻}"
  unfolding Rω_def by (use Ri_less in auto)

lemma encstep2D:
  assumes "(s, t) ∈ encstep2 E R"
  shows "∃C τ l r. (l, r) ∈ R ∪ E ∧ s = C⟨l ⋅ τ⟩ ∧ t = C⟨r ⋅ τ⟩ ∧ l ⋅ τ ≻ r ⋅ τ ∧ l ⊲⋅ s"
  using assms
proof (cases)
  case (estep l r C σ)
  then show ?thesis by blast
next
  case (rstep l r C σ)
  from rstep(1) and Ri_less and subst [of l r] have "l ⋅ σ ≻ r ⋅ σ" by auto
  with rstep show ?thesis by blast
qed

definition lessencp (infix "⋅≻" 50) where "lessencp s t ⟷ (s, t) ∈ (relto ({⋅⊳} ∪ {≻}) {⋅⊵})+"
abbreviation lessenc ("{⋅≻}") where "lessenc ≡ {(s, t). s ⋅≻ t}"

definition leqencp (infix "⋅≽" 50) where "leqencp s t ⟷ (s, t) ∈ (relto ({⋅⊳} ∪ {≻}) {⋅⊵})+ ∪ {⋅⊵}*"
abbreviation leqenc ("{⋅≽}") where "leqenc ≡ {(s, t). s ⋅≽ t}"

lemma SN_lessenc: "SN {⋅≻}"
  using reduction_order.SN_encomp_Un_less_relto_encompeq
    SN_encomp_Un_less_relto_encompeq SN_trancl_SN_conv unfolding lessencp_def by auto

lemma lessenc_compat1: "{⋅≻} O {⋅⊵} ⊆ {⋅≻}"
proof
  fix s t
  let ?R = "{⋅⊵}* O ({⋅⊳} ∪ {≻}) O {⋅⊵}*"
  assume "(s, t) ∈ {⋅≻} O {⋅⊵}"
  then obtain u where u: "(u, t) ∈ {⋅⊵}" "s ⋅≻ u" by auto
  then obtain v where v: "(v, u) ∈ ?R" "(s, v) ∈ ?R*" using tranclD2 unfolding lessencp_def by metis
  from u(1) v(1) have "(v, t) ∈ ?R O {⋅⊵}" by blast
  with u(1) v(1) have "(v, t) ∈ ?R" using O_assoc O_mono2 [of "{⋅⊵} O {⋅⊵}*" "{⋅⊵}*"]
    by (smt r_into_rtrancl relcomp.relcompI rtrancl_idemp_self_comp u(1) v(1))
  with v(2) have "s ⋅≻ t" unfolding trancl_unfold_right lessencp_def by blast
  then show "(s, t) ∈ {⋅≻}" by auto
qed

lemma lessleqenc_compat1: "{⋅≻} O {⋅≽} ⊆ {⋅≻}"
proof
  fix x z
  assume "(x, z) ∈ {⋅≻} O {⋅≽}"
  then obtain y where y: "(x, y) ∈ {⋅≻}" "(y, z) ∈ {⋅≽}" by auto
  then consider "y ⋅≻ z" | "(y, z) ∈ {⋅⊵}*"
    unfolding leqencp_def lessencp_def Un_iff by blast
  then show "(x, z) ∈ {⋅≻}"
  proof (cases)
    case 1
    with y(1) show ?thesis using trancl_trans unfolding lessencp_def by auto
  next
    case 2
    from encompeq_trans trans_inv_image transI have t: "trans {⋅⊵}"
      by (smt Pair_inject case_prodE case_prodI2 mem_Collect_eq)
    have "{⋅⊵}* = {⋅⊵}" by(rule trans_refl_imp_rtrancl_id, insert t encompeq_refl refl_O_iff, auto)
    with 2 [unfolded this] y(1) lessenc_compat1 show ?thesis by auto
  qed
qed

lemma lessenc_compat2: "{⋅⊵} O {⋅≻} ⊆ {⋅≻}"
proof
  fix s t
  let ?R = "{⋅⊵}* O ({⋅⊳} ∪ {≻}) O {⋅⊵}*"
  assume "(s, t) ∈ {⋅⊵} O {⋅≻}"
  then obtain u where u: "(s, u) ∈ {⋅⊵}" "u ⋅≻ t" by auto
  then obtain v where v: "(u, v) ∈ ?R" "(v, t) ∈ ?R*" using tranclD unfolding lessencp_def by metis
  from u(1) v(1) have "(s, v) ∈ {⋅⊵} O ?R" by blast
  with u(1) v(1) have "(s, v) ∈ ?R" using O_assoc O_mono2 [of "{⋅⊵} O {⋅⊵}*" "{⋅⊵}*"]
    by (smt r_into_rtrancl relcomp.relcompI rtrancl_idemp_self_comp u(1) v(1))
  with v(2) have "s ⋅≻ t" unfolding trancl_unfold_left lessencp_def by blast
  then show "(s, t) ∈ {⋅≻}" by auto
qed

lemma encstep1_lessenc:
  assumes "(s, t) ∈ encstep1 (E i) (R i)"
  shows "s ⋅≻ t"
  using assms encstep1_less[OF assms Ri_less] unfolding lessencp_def by fast

lemma trans_less_enc: "trans {⋅≻}"
proof -
  from commutes_rewrel_encomp [of "{≻}"] and ctxt and subst have "{⋅⊵} O {≻} ⊆  {≻} O {⋅⊵}" by simp
  with trans_relto [of "{≻}" "{⋅⊵}"] and trans show ?thesis unfolding lessencp_def by auto
qed

lemma refl_encompeq: "refl {⋅⊵}"
  using encompeq_refl reflpI reflp_refl_eq by fastforce

abbreviation "mulless ≡ s_mul_ext {⋅⊵} {⋅≻}"
abbreviation "mulleq ≡ ns_mul_ext {⋅⊵} {⋅≻}"

sublocale lessenc_op: SN_order_pair mulless "ns_mul_ext {⋅⊵} {⋅≻}"
proof (intro SN_order_pair.mul_ext_SN_order_pair)
  show "SN_order_pair {⋅≻} {⋅⊵}"
  proof
    show "trans {⋅≻}" using trans_less_enc by auto
  next
    show "refl {⋅⊵}" using refl_encompeq .
  next
    show "trans {⋅⊵}" using encomp_order.order_trans transpI transp_trans by metis
  next
    from SN_lessenc show "SN {⋅≻}" by auto
  next
    from lessenc_compat2 show "{⋅⊵} O {⋅≻} ⊆ {⋅≻}" by auto
  next
    from lessenc_compat1 show "{⋅≻} O {⋅⊵} ⊆ {⋅≻}" by auto
  qed
qed

lemma mset_two:
  assumes "s ⋅≻ u"
  shows "({#s, t#}, {#u, t#}) ∈ mulless"
proof -
  from assms s_mul_ext_singleton have su: "({#s#}, {#u#}) ∈ mulless" by auto
  from ns_mul_ext_refl [OF refl_encompeq] have "({#t#}, {#t#}) ∈ mulleq" by auto
  from s_ns_mul_ext_union_compat [OF su this] show ?thesis
    using add.commute add_mset_add_single by metis
qed

lemma mset_two2:
  assumes "s ⋅≻ u"
  shows "({#t, s#}, {#t, u#}) ∈ mulless"
  using mset_two [OF assms] add_mset_commute by (metis (no_types, lifting))

definition Sw ("Sω") where "Sω ≡ Rω ∪ {(s ⋅ σ, t ⋅ σ) | s t σ. (s, t) ∈ Eω ∧ s ⋅ σ ≻ t ⋅ σ}"
  
lemma Sw_less: "Sω ⊆ {≻}"
  using Ri_less unfolding Sw_def Rω_def by fast
  
lemma Sw_step_less: "rstep Sω ⊆ {≻}"
  using compatible_rstep_imp_less[OF Sw_less] by auto
    
lemma Sw_implies_REw: "rstep Sω ⊆ (rstep (Rω ∪ Eω))"
  unfolding Sw_def rstep_mono by fastforce

lemma SN_Sw_step: "SN (rstep Sω)"
  using Sw_step_less SN_subset [OF SN_less] by fastforce
    
lemma rstep_imp_mstep:
  assumes "(t, u) ∈ rstep RR" and "∃t' ∈# S. t' ≽ t" and "∃u' ∈# S. u' ≽ u"
  shows "(t, u) ∈ mstep S RR"
  using assms by (auto simp: mstep_iff dest: trans)
    
lemma mstep_succeq_mono:
  assumes "∀m ∈# M. ∃n ∈# N. n ≽ m"
  shows "mstep M RR ⊆ mstep N RR"
proof
  fix s t
  assume mstep: "(s, t) ∈ mstep M RR"
  with assms trans show "(s, t) ∈ mstep N RR" unfolding mstep_def mem_Collect_eq split
    by (smt sup2E sup2I1)
qed

lemma msteps_succeq_mono:
  assumes "∀m ∈# M. ∃n ∈# N. n ≽ m"
  shows "(mstep M RR)* ⊆ (mstep N RR)*"
  by (rule mstep_succeq_mono [OF assms, THEN conversion_mono])

lemma mstep_subst_ctxt:
  assumes "(s, t) ∈ mstep S RR"
  shows "(C⟨s ⋅ σ⟩, C⟨t ⋅ σ⟩) ∈ mstep {#C⟨u ⋅ σ⟩. u ∈# S #} RR"
proof -
  let ?S = "{#C⟨u ⋅ σ⟩. u ∈# S #}"
  from assms have step: "(C⟨s ⋅ σ⟩, C⟨t ⋅ σ⟩) ∈ rstep RR" unfolding mstep_def by auto
  from assms have gt:"∃s' ∈# S. s' ≽ s" "∃t' ∈# S. t' ≽ t" unfolding mstep_def by auto
  with subst ctxt have s':"∃s' ∈# ?S. s' ≽ C⟨s ⋅ σ⟩" by fast
  from gt(2) subst ctxt have "∃t' ∈# ?S. t' ≽ C⟨t ⋅ σ⟩" by fast
  with step s' show ?thesis unfolding mstep_def by blast
qed

lemma msteps_subst_ctxt:
  assumes "(s, t) ∈ (mstep S RR)*"
  shows "(C⟨s ⋅ σ⟩, C⟨t ⋅ σ⟩) ∈ (mstep {#C⟨u ⋅ σ⟩. u ∈# S #} RR)*"
proof -
  let ?S = "{#C⟨u ⋅ σ⟩. u ∈# S #}"
  from mstep_subst_ctxt have x: "⋀ s t. (s, t) ∈ (mstep S RR) ⟹ (C⟨s ⋅ σ⟩, C⟨t ⋅ σ⟩) ∈ (mstep ?S RR)" by auto
  from assms rtrancl_imp_relpow obtain n where "(s, t) ∈ (mstep S RR) ^^ n" by fast
  then have "(C⟨s ⋅ σ⟩, C⟨t ⋅ σ⟩) ∈ (mstep ?S RR) ^^ n"
  proof (induct n arbitrary: s)
    case (Suc n)
    then obtain s' where s': "(s, s') ∈ (mstep S RR)" "(s', t) ∈ (mstep S RR) ^^ n"
      unfolding relpow_Suc by blast
    from x [OF s'(1)] Suc(1) [OF s'(2)] show ?case unfolding relpow_Suc by blast
  qed simp
  with relpow_imp_rtrancl show ?thesis by blast
qed

lemma Einf_without_Ew:
  assumes "(l,r) ∈ E" and "(l,r) ∉  Eω"
  shows "∃j. (l, r) ∈ (E j) ∧ (l, r) ∉ (E (Suc j))"
proof-
  from assms obtain j where "(l, r) ∈ E j" and "∃i>j. (l, r) ∉ E i"
    unfolding Eω_def le_eq_less_or_eq by blast
  moreover define i where "i = (LEAST i. i > j ∧ (l, r) ∉ E i)"
  ultimately have "i > j" and not_i: "(l, r) ∉ E i" by (metis (lifting) LeastI)+
  then have "i - 1 < i" and [simp]: "Suc (i - Suc 0) = i" by auto
  from not_less_Least [OF this(1) [unfolded i_def], folded i_def] and ‹i > j› and ‹(l, r) ∈ E j›
  have pred_i: "(l, r) ∈ E (i - 1)" by (cases "i = Suc j") auto
  have i_simp: "Suc (i - 1) = i" by auto
  show ?thesis by (rule exI[of _ "i - 1"], insert not_i pred_i i_simp, auto)
qed

lemma Rinf_without_Rw:
  assumes "(l,r) ∈ R" and "(l,r) ∉  Rω"
  shows "∃j. (l, r) ∈ (R j) ∧ (l, r) ∉ (R (Suc j))"
proof-
  from assms obtain j where "(l, r) ∈ R j" and "∃i>j. (l, r) ∉ R i"
    unfolding Rω_def le_eq_less_or_eq by blast
  moreover define i where "i = (LEAST i. i > j ∧ (l, r) ∉ R i)"
  ultimately have "i > j" and not_i: "(l, r) ∉ R i" by (metis (lifting) LeastI)+
  then have "i - 1 < i" and [simp]: "Suc (i - Suc 0) = i" by auto
  from not_less_Least [OF this(1) [unfolded i_def], folded i_def] and ‹i > j› and ‹(l, r) ∈ R j›
  have pred_i: "(l, r) ∈ R (i - 1)" by (cases "i = Suc j") auto
  have i_simp: "Suc (i - 1) = i" by auto
  show ?thesis by (rule exI[of _ "i - 1"], insert not_i pred_i i_simp, auto)
qed
  
lemma Ei_subset_EwRi: "mstep S E ⊆ (mstep S (R ∪ Eω))*"
proof -
  { fix t0 u0
    assume "(t0, u0) ∈ rstep E" and St: "∃s ∈# S. s ≽ t0" and Su: "∃s ∈# S. s ≽ u0"
    with rstep_imp_C_s_r [OF this(1)] obtain C σ t u
      where Cs: "t0 = C⟨t ⋅ σ⟩" "u0 = C⟨u ⋅ σ⟩" and tu: "(t, u) ∈ E" by blast
    from tu St Su have "(C⟨t ⋅ σ⟩, C⟨u ⋅ σ⟩) ∈ (mstep S (R ∪ Eω))*" unfolding Cs
    proof (induct "{#t, u#}" arbitrary: t u S C σ rule: SN_induct [OF lessenc_op.SN])
      case 1
      note sc = this
      let ?t = "C⟨t ⋅ σ⟩" and ?u = "C⟨u ⋅ σ⟩"
      have comm: "⋀t u. {#t, u#} =  {#u, t#}" using add_mset_commute by auto
      note r_imp_msteps = rstep_imp_mstep [OF _ sc(4) sc(3)] rstep_imp_mstep [OF _ sc(3) sc(4)]
      show ?case
      proof (cases "(?t, ?u) ∈ rstep Eω")
        case True
        with r_imp_msteps mstep_Un [of _ "R" "Eω"] show ?thesis by auto
      next
        case False
        with rstep_imp_C_s_r have "(t, u) ∉ Eω" by fast
        from Einf_without_Ew[OF sc(2) this]  obtain i where
          pred_i:"(t, u) ∈ E i" and not_i: "(t, u) ∉ E (Suc i)" by auto
        let ?enc = "encstep1 (E (Suc i)) (R (Suc i)) O (E (Suc i))"
        from pred_i oKBi_E_supset [OF irun [of i]] and not_i
        consider "(t, u) ∈ (R (Suc i))" | "u = t" | "(t, u) ∈ ?enc" by auto blast
        then show ?thesis
        proof (cases)
          case 1
          then have "(?t, ?u) ∈ rstep (R)" by auto
          with r_imp_msteps [of "R"] converse_iff have "(?t, ?u) ∈ (mstep S R) ∨ (?u, ?t) ∈ (mstep S R)"
            by fastforce
          with mstep_Un conversion_inv show ?thesis by blast
        next
          case 3
          then have "(t, u) ∈ ?enc ∨ (u, t) ∈ ?enc" by auto
          with sc obtain t' u' where tu': "(t', u') ∈ ?enc" "(t=t' ∧ u=u') ∨ (t=u' ∧ u=t')" by blast
          then have tu_tu': "{# t, u #} = {# t', u' #}" by force
          from sc(3) sc(4) tu'(2) trans have Stu: "∃s ∈# S. s ≽ C⟨u' ⋅ σ⟩ ∧ (∃s ∈# S. s ≽ C⟨t' ⋅ σ⟩)" by meson
          from tu'(1) obtain v where
            v: "(t', v) ∈ encstep1 (E (Suc i)) (R (Suc i))" "(v, u') ∈ (E (Suc i))" by auto
          from tu_tu' mset_two [OF encstep1_lessenc [OF v(1)]]
          have dec: "({# t, u #}, {# v, u' #}) ∈ mulless" by auto
          then have dec': "({# t, u #}, {# u', v #}) ∈ mulless" using add_mset_commute by metis
          from encstep1_ostep [OF v(1), THEN ostep_imp_less [OF Ri_less]] have tv: "t' ≽ v" by blast
          have vu: "(C⟨v⋅ σ⟩, C⟨u'⋅ σ⟩) ∈ (mstep S (R ∪ Eω))*"
          proof -
            from trans subst ctxt tv Stu have Sv: "∃s ∈# S. s ≽ C⟨v⋅ σ⟩" by fast
            from v(2) have "(v, u') ∈ E" by auto
            with sc(1) [OF dec] sc(1) [OF dec'] Sv Stu show ?thesis
              unfolding Un_iff conversion_inv by blast
          qed
          from v(1) have "(C⟨t' ⋅ σ⟩, C⟨v ⋅ σ⟩) ∈ (mstep S (R ∪ Eω))*"
          proof (cases)
            case (estep l r D τ)
            then have lr: "(l, r) ∈ E" by auto
            with encompeq.intros [of t' D "l ⋅ τ" Var] estep(2) have tl: "l ⋅ τ ⊴⋅ t'" by auto
            with encompeq.intros [of "r ⋅ τ" ] estep(2) have "r ⊴⋅ r ⋅ τ" by auto
            with estep(4) tl have tr: "t' ⋅≻ r" unfolding lessencp_def by blast
            from estep(5) have "t' ⋅≻ l" unfolding lessencp_def by blast
            with tr ns_s_mul_ext_union_multiset_l [OF ns_mul_ext_bottom [of "{#u'#}"], of _ "{#l, r#}"]
            have "({#u', t'#}, {#l, r#}) ∈ mulless" using multi_member_last by auto
            with tu_tu' have dec: "({#t, u#}, {#l, r#}) ∈ mulless" using comm by metis
            with comm [of l] have dec': "({#t, u#}, {#r, l#}) ∈ mulless" by presburger
            note ih = sc(1) [OF dec, of "{#l, r#}"  Var] sc(1) [OF dec', of "{#l, r#}"  Var]
            from lr [unfolded Un_iff] ih have "(l, r) ∈ (mstep {#l, r#} (R ∪ Eω))*"
              unfolding conversion_inv [of r] ctxt.cop_nil subst_apply_term_empty by force
            from msteps_subst_ctxt [OF this] estep(2) estep(3)
            have mstep: "(t', v) ∈ (mstep {#t', v#} (R ∪ Eω))*" by auto
            from Stu tv trans ctxt [OF subst] have "∃s∈#S. s ≽ C⟨v ⋅ σ⟩" by fast
            with Stu tv have S:"∀m∈#{#C⟨t' ⋅ σ⟩, C⟨v ⋅ σ⟩#}. ∃n∈#S. n ≽ m" by auto
            have mset:"{#C⟨u ⋅ σ⟩. u ∈# {#t', v#}#} = {#C⟨t' ⋅ σ⟩, C⟨v ⋅ σ⟩#}" by simp
            with msteps_subst_ctxt [OF mstep, of C σ] conversion_mono [OF mstep_succeq_mono [OF S]]
              show ?thesis unfolding mset by blast
          next
            case rstep
            with rstep_subset_less [OF Ri_less] have tv': "t' ≻ v" by blast
            from rstep have "(t', v) ∈ rstep (R ∪ Eω)" by fast
            from rstep_imp_mstep [OF this] tv'  have "(t', v) ∈ mstep {#t', t'#} (R ∪ Eω)" by auto
            from mstep_subst_ctxt [OF this, of C σ] mstep_succeq_mono [of "{#C⟨t' ⋅ σ⟩, C⟨t' ⋅ σ⟩#}" S] Stu
            show ?thesis by fastforce
          qed
          from transD [OF conversion_trans, OF this vu] tu'(2) symD [OF conversion_sym]
          show ?thesis by metis
        qed auto
      qed
    qed
    with Cs have "(t0, u0) ∈ (mstep S (R ∪ Eω))*" by auto
  }
  then show ?thesis unfolding mstep_def by blast
qed

lemma Einf_sym_mstep: "mstep S (E) ⊆ (mstep S (R ∪ Eω))*"
proof -
  have inv: "mstep S (E¯) = (mstep S E)¯"
    unfolding mstep_def reverseTrs [symmetric] converse_iff by auto
  from Ei_subset_EwRi show ?thesis unfolding mstep_Un inv Un_subset_iff
    by (metis converse_mono conversion_converse)
qed

abbreviation "lexless ≡ lex_two {⋅≻} Id {⋅≻}"

lemma SN_lexless: "SN lexless"
  using SN_lessenc
  by (intro lex_two) (auto simp only: SN_trancl_SN_conv intro: SN_encomp_Un_less_relto_encompeq)

lemma Ri_subset_ERw: "mstep S R ⊆ (mstep S (Rω ∪ Eω))*"
proof -
  { fix t u
    assume "(t, u) ∈ rstep R" and St: "∃s ∈# S. s ≽ t" and Su: "∃s ∈# S. s ≽ u"
    with rstep_imp_C_s_r [OF this(1)] obtain C σ l r
      where Cs: "t = C⟨l ⋅ σ⟩" "u = C⟨r ⋅ σ⟩" and tu: "(l, r) ∈ R" by blast
    from tu St Su have "(C⟨l ⋅ σ⟩, C⟨r ⋅ σ⟩) ∈ (mstep S (Rω ∪ Eω))*" unfolding Cs
    proof (induct "(l, r)" arbitrary: l r S C σ rule: SN_induct [OF SN_lexless])
      case 1
      note sc = this
      with Ri_less have lr: "l ≻ r" by auto
      let ?t = "C⟨l ⋅ σ⟩" and ?u = "C⟨r ⋅ σ⟩"
      show ?case
      proof (cases "(l, r) ∈ Rω")
        case True
          (* rule l → r is persistent *)
        then have "(?t, ?u) ∈ rstep Rω" by auto
        with rstep_imp_mstep [OF this sc(3) sc(4)] mstep_Un show ?thesis by auto
      next
        case False
        (* rule l → r was removed in step (E i, R i) ⊢oKB (E (i+1), R (i+1)) *)
        from Rinf_without_Rw[OF sc(2) False] obtain i where
          pred_i: "(l, r) ∈ R i" and not_i: "(l, r) ∉ R (Suc i)" by auto
        (* first verify that (R ∪ Eω) conversion below l corresponds to a (Rω ∪ Eω) conversion *)
        { fix S'
          assume below_l: "∀ s' ∈# S'. l ⋅≻ s'"
          { fix s t
            assume "(s, t) ∈ mstep S' R"
            then have mstep: "(s, t) ∈ rstep R" "∃s' ∈# S'. s' ≽ s" "∃s' ∈# S'. s' ≽ t"
              unfolding mstep_def by auto
            then obtain C σ l1 r1 where l1r1: "s = C⟨l1 ⋅ σ⟩" "t = C⟨r1 ⋅ σ⟩" "(l1, r1) ∈ R" by fast
            from mstep obtain s' where s': "s' ∈# S'" "s' ⋅≽ s" unfolding leqencp_def by blast
            from l1r1 encompeq.intros have sl1: "s ⋅⊵ l1" by auto
            from s' below_l have ls': "l ⋅≻ s'" by auto
            with s' lessleqenc_compat1 have "l ⋅≻ s" by auto
            with sl1 lessenc_compat1 have "l ⋅≻ l1" by auto
            then have "((l, r), (l1, r1)) ∈ lexless" by force
            from sc(1) [OF this l1r1(3)] mstep have "(s, t) ∈ (mstep S' (Rω ∪ Eω))*"
              unfolding l1r1 by meson
          } note step = this
          from r_into_rtrancl have "mstep S' Eω ⊆ (mstep S' Eω)*" by auto
          with step mstep_Un have "mstep S' (R ∪ Eω) ⊆ (mstep S' (Rω ∪ Eω))*" by auto
          from conversion_mono [OF this] have R: "(mstep S' (R ∪ Eω))* ⊆ (mstep S' (Rω ∪ Eω))*"
            by simp
        } note mstep_Rinf_msteps_ERw = this
        (* derive equation (1): (R ∪ E) steps below l correspond to a (Rω ∪ Eω) conversion *)
        { fix S'
          assume below_l: "∀ s' ∈# S'. l ⋅≻ s'"
          note R = mstep_Rinf_msteps_ERw [OF this]
          with Einf_sym_mstep have "mstep S' (E) ⊆ (mstep S' (Rω ∪ Eω))*" by auto
          with R mstep_Un have "mstep S' (R ∪ E) ⊆ (mstep S' (Rω ∪ Eω))*" by auto
          from conversion_mono [OF this] have "(mstep S' (R ∪ E))* ⊆ (mstep S' (Rω ∪ Eω))*"
            by simp
        } note msteps_ERinf_mstep_Rw = this
        (* case distinction according to Lemma 31 *)
        let ?Ei = "E (Suc i)" and ?Ri = "R (Suc i)"
        from pred_i oKBi_R_supset [OF irun [of i]] and not_i
          consider "(l, r) ∈ encstep2 ?Ei ?Ri O ?Ei" | "(l, r) ∈ ?Ri O (ostep ?Ei ?Ri)¯" by blast
        then have "(l, r) ∈ (mstep {#l, r#} (Rω ∪ Eω))*"
        proof (cases)
          (* collapse step *)
          case 1
          then obtain u where u: "(l, u) ∈ encstep2 ?Ei ?Ri" "(u, r) ∈ ?Ei" by auto
          (* l ⟷Rω  Eω{l, r}* u *)
          with encstep2_mono [of ?Ei "E" ?Ri "R"] have lu: "(l, u) ∈ encstep2 E R" by auto
          from encstep2D [OF this] obtain D τ l0 r0 where
            D: "(l0, r0) ∈ R ∪ E" "l = D⟨l0 ⋅ τ⟩" "u = D⟨r0 ⋅ τ⟩" " l0 ⋅ τ ≻ r0 ⋅ τ" "l0 ⊲⋅ l" by blast
          then have l_l0: "l ⋅≻ l0" unfolding lessencp_def by auto
          from D encompeq.intros [of l D _ Var] have l_l0τ: "l ⋅⊵ l0 ⋅ τ" by auto
          from encompeq.intros [of _ ] have "r0 ⋅ τ ⋅⊵ r0" by auto
          with l_l0τ D(4) have l_r0: "l ⋅≻ r0" unfolding lessencp_def by blast
          have "(l0, r0) ∈ (mstep {#l0, r0#} (Rω ∪ Eω))*"
          proof (cases "(l0, r0) ∈ R")
            case True
            with Ri_less have l_gt_r: "l0 ≻ r0" by auto
            from l_l0 have "((l, r), (l0, r0)) ∈ lexless" by auto
            from sc(1) [OF this True, of "{#l0, r0#}"  Var, simplified] l_gt_r show ?thesis by auto
          next
            case False
            with D(1) have "(l0, r0) ∈ E" by auto
            then have "(l0, r0) ∈ mstep {#l0, r0#} (E)" unfolding mstep_def by auto
            with in_mono [OF Einf_sym_mstep] have l0r0: "(l0, r0) ∈ (mstep {#l0, r0#} (R ∪ Eω))*" by auto
            from l_l0 l_r0 have "∀ s' ∈# {#l0, r0#}. l ⋅≻ s'" by auto
            from l0r0 mstep_Un mstep_Rinf_msteps_ERw [OF this] show ?thesis by auto
          qed
          from D msteps_subst_ctxt [OF this] have lu: "(l, u) ∈ (mstep {#l, u#} (Rω ∪ Eω))*" by auto
          from u(1) encstep1_less [OF encstep2_encstep1 Ri_less] have l_succeq_u: "l ≽ u" by auto
          then have "∀m∈#{#l, u#}. ∃n∈#{#l, r#}. n ≽ m" by auto
          from lu msteps_succeq_mono [OF this] have lu: "(l, u) ∈ (mstep {#l, r#} (Rω ∪ Eω))*" by blast
              (* u ⟷Rω  Eω{l, r}* r *)
          from u(1) encstep1_lessenc [OF encstep2_encstep1] have "l ⋅≻ u" by auto
          with lr have lt_l: "∀ s' ∈# {#u, r#}. l ⋅≻ s'" unfolding lessencp_def by force
          from u(2) have "(u, r) ∈ mstep {#u, r#} (R ∪ E)" unfolding mstep_def by auto
          from r_into_rtrancl [OF this] msteps_ERinf_mstep_Rw [OF lt_l]
          have ur: "(u, r) ∈ (mstep {#u, r#} (Rω ∪ Eω))*" by auto
          from l_succeq_u have "∀m∈#{#u, r#}. ∃n∈#{#l, r#}. n ≽ m" by auto
          from ur msteps_succeq_mono [OF this] have "(u, r) ∈ (mstep {#l, r#} (Rω ∪ Eω))*" by blast
          with lu transD [OF conversion_trans] show ?thesis by auto
        next
          (* compose step *)
          case 2
          then obtain u where u: "(l, u) ∈ ?Ri" "(u, r) ∈ (ostep ?Ei ?Ri)¯" by auto
              (* l ⟷Rω  Eω{l, r}* u*)
          with ostep_imp_less [OF Ri_less] have ru: "r ≻ u" by auto
          then have "r ⋅≻ u" unfolding lessencp_def by blast
          then have "((l, r), (l, u)) ∈ lexless" unfolding lex_two.simps mem_Collect_eq split by auto
          from trans [OF lr ru] sc(1) [OF this, of "{#l, r#}"  Var, simplified] u(1)
          have lu: "(l, u) ∈ (mstep {#l, r#} (Rω ∪ Eω))*" by auto
              (* u ⟷Rω  Eω{l, r}* r*)
          from ostep_imp_rstep_sym u(2) have "(r, u) ∈ rstep (?Ri ∪ ?Ei)" by auto
          with rstep_UN rstep_union have rstep: "(r, u) ∈ rstep (R ∪ E)" by fast
          from rstep have mstep: "(r, u) ∈ mstep {#u, r#} (R ∪ E)"
            unfolding mstep_def mem_Collect_eq split by auto
          from lr trans [OF lr ru] have lt_l: "∀ s' ∈# {#u, r#}. l ⋅≻ s'"
            unfolding lessencp_def by force
          from msteps_ERinf_mstep_Rw [OF this] r_into_rtrancl [OF mstep]
          have ur: "(u, r) ∈ (mstep {#u, r#} (Rω ∪ Eω))*" unfolding conversion_inv [of u] by auto
          have "(u, r) ∈ (mstep {#l, r#} (Rω ∪ Eω))*"
            by(rule subsetD [OF msteps_succeq_mono], insert trans [OF lr ru] ur, auto)
          with lu transD [OF conversion_trans] show ?thesis by auto
        qed
        from msteps_subst_ctxt [OF this] have tu: "(?t, ?u) ∈ (mstep {#?t, ?u#} (Rω ∪ Eω))*" by auto
        show ?thesis by(rule subsetD [OF msteps_succeq_mono], insert sc(3) sc(4) tu, auto)
      qed
    qed
    with Cs have "(t, u) ∈ (mstep S (Rω ∪ Eω))*" by auto
  }
  then show ?thesis using mstep_def by auto
qed

lemma RiEw_subset_ERw: "(mstep S (R ∪ Eω))* ⊆ (mstep S (Rω ∪ Eω))*"
proof -
  from r_into_rtrancl have "mstep S Eω ⊆ (mstep S Eω)*" by auto
  with Ri_subset_ERw mstep_Un have "mstep S (R ∪ Eω) ⊆ (mstep S (Rω ∪ Eω))*" by auto
  from conversion_mono [OF this] show ?thesis by auto
qed

lemma ERi_subset_ERw: "mstep S (E ∪ R) ⊆ (mstep S (Rω ∪ Eω))*"
  using Ei_subset_EwRi Ri_subset_ERw RiEw_subset_ERw unfolding mstep_Un [of _ "E"] by blast

lemma oKBi_conversion_ERw: "(rstep (E 0))* = (rstep (Rω ∪ Eω))*"
proof
  have "rstep (E 0) ⊆ (rstep (Rω ∪ Eω))*"
  proof
    fix s t
    assume "(s, t) ∈ rstep (E 0)"
    with ERi_subset_ERw [of "{#s, t#}"] mstep_def have st: "(s, t) ∈ (mstep {#s, t#} (Rω ∪ Eω))*" by force
    have "mstep {#s, t#} (Rω ∪ Eω) ⊆ rstep (Rω ∪ Eω)" using mstep_def by force
    from conversion_mono [OF this] st show "(s, t) ∈ (rstep (Rω ∪ Eω))*" by auto
  qed
  from conversion_mono [OF this] show "(rstep (E 0))* ⊆ (rstep (Rω ∪ Eω))*" by simp
next
  have "rstep (Rω ∪ Eω) ⊆ (rstep (E 0))*"
  proof
    fix s t
    assume "(s, t) ∈ rstep (Rω ∪ Eω)"
    then obtain k where "(s, t) ∈ rstep (R k ∪ E k)" unfolding Eω_def Rω_def by blast
    with oKBi_conversion_ERi show "(s, t) ∈ (rstep (E 0))*" by blast
  qed
  from conversion_mono [OF this] show "(rstep (Rω ∪ Eω))* ⊆ (rstep (E 0))*" by simp
qed
end

locale okb_irun_nonfailing = okb_irun less
  for less :: "('a, 'b::infinite) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50) +
  assumes Ew_empty:"Eω = {}"
  and fair:"PCP Rω ⊆ (rstep E)"
begin

lemma [simp]:"Sω = Rω" unfolding Ew_empty Sw_def by auto

lemma Ew_empty_implies_Rw_peak_decreasing:
  assumes "(s, t) ∈ slab Rω s" and "(s, u) ∈ slab Rω s"
  shows "(t, u) ∈ (⋃ v ∈ {v. s ≻ v}. slab Rω v)*"
proof-
  from assms have rsteps: "(s, t) ∈ rstep Rω" "(s, u) ∈ rstep Rω" unfolding source_step_def by auto
  note vc = SN_imp_variable_condition [OF SN_Sw_step]
  from vc peak_imp_nabla2 [OF _ rsteps] have tu: "(t, u) ∈ nabla Rω s ^^ 2" by auto
  from compatible_rstep_imp_less [OF Rw_less] have Sw_step_less: "rstep Rω ⊆ {≻}" by blast
  let ?less_s_conv = "(⋃z∈{z. s ≻ z}. slab Rω z)*"
  let ?leq_conv = "λ v. (⋃z∈{z. v ≽ z}. slab Rω z)*"
  { fix v w
    assume vw: "(v, w) ∈ nabla Rω s" and "s ≻ v" and "s ≻ w"
    then have vw: "(v, w) ∈ (rstep Rω) ∨ (v, w) ∈ (rstep (PCP Rω))" unfolding nabla_def by auto
    { assume "(v, w) ∈ (rstep Rω)"
      then obtain x where vx: "(v, x) ∈ (rstep Rω)*" and wx: "(w, x) ∈ (rstep Rω)*"
        unfolding join_def rtrancl_converse by blast
      from rsteps_slabI [OF vx _ Sw_step_less, of v] have "(v, x) ∈ ?leq_conv v" by auto
      from slab_conv_less_label [OF this] ‹s ≻ v› have vx: "(v, x) ∈ ?less_s_conv" by auto
      from rsteps_slabI [OF wx _ Sw_step_less, of w] have "(w, x) ∈ ?leq_conv w" by auto
      from slab_conv_less_label [OF this] ‹s ≻ w› have "(x, w) ∈ ?less_s_conv"
        unfolding conversion_inv by auto
      from conversion_trans' [OF vx this] have "(v, w) ∈ ?less_s_conv" by auto
    } note if_joinable = this
    { assume vw:"(v, w) ∈ (rstep (PCP Rω))"
      from rstep_mono[OF fair] have "rstep (PCP Rω) ⊆ (rstep E)"
        unfolding rstep_union reverseTrs rstep_rstep by fastforce
      hence *:"(rstep (PCP Rω)) ⊆ (rstep E)" by blast
      have "(rstep E) ⊆ (rstep (E ∪ R))" using rstep_union by fast
      hence "(rstep (PCP Rω)) ⊆ (rstep (E ∪ R))" using * by auto
      from in_mono[OF this, of "(v,w)"] vw have "(v, w) ∈ (rstep (E ∪ R))" by fast
      hence vw:"(v, w) ∈ (mstep {#v, w#} (E ∪ R))"
        unfolding mstep_def mem_Collect_eq Un_iff converse_iff by auto
      from ERi_subset_ERw conversion_converse have "⋀S. (mstep S (E ∪ R)) ⊆ (mstep S (Rω ∪ Eω))*" by blast
      from in_mono[OF this, rule_format, OF vw] have vw: "(v, w) ∈ (mstep {#v, w#} Rω)*"
        unfolding Ew_empty by auto
      from ‹s ≻ v› ‹s ≻ w› have "∀ t ∈# {#v, w#}. s ≻ t" by auto
      from vw msteps_imp_source_steps [OF this] have "(v, w) ∈ ?less_s_conv" by auto
    }
    with vw if_joinable have "(v, w) ∈ ?less_s_conv" by auto
  } note nabla_slab = this
  from tu obtain v where v: "(t, v) ∈ nabla Rω s" "(v, u) ∈ nabla Rω s" by auto
  from v [unfolded nabla_def] have sv: "(s, v) ∈ (rstep Rω)+" by auto
  from sv rsteps_subset_less [OF Sw_less] have sv: "s ≻ v" by auto
  from Sw_step_less rsteps have st: "s ≻ t" and su: "s ≻ u" by auto
  note tv = nabla_slab [OF v(1) st sv]
  note vu = nabla_slab [OF v(2) sv su]
  with conversion_trans' [OF tv vu] show ?thesis by auto
qed

lemma Ew_empty_implies_CR_Rw:"CR (rstep Rω)"
proof -
  interpret ars_peak_decreasing "slab Rω" UNIV "op ≻"
    by (unfold_locales, insert Ew_empty_implies_Rw_peak_decreasing SN_less, auto)
  from CR show ?thesis unfolding UN_source_step .
qed

end

locale gtotal_okb_irun = okb_irun + gtotal_reduction_order less
begin

lemma gmstep_ERw_Sw: "GROUND ((mstep S (Rω ∪ Eω))) ⊆ GROUND (((mstep S Sω))=)"
proof
  fix s t
  assume "(s, t) ∈ GROUND ((mstep S (Rω ∪ Eω)))"
  then have ground: "ground s" "ground t" and mstep: "(s, t) ∈ (mstep S (Rω ∪ Eω))" unfolding GROUND_def by auto
  then have dom: "∃s' t'. s' ∈# S ∧ t' ∈# S ∧ s' ≽ s ∧ t' ≽ t" unfolding mstep_def by auto
  from mstep reverseTrs have mstep: "(s, t) ∈ mstep S (Rω ∪ Eω)"
    unfolding mstep_def Un_iff converse_iff mem_Collect_eq split rstep_union by blast
  from mstep mstep_Un consider "(s, t) ∈ mstep S (Eω)" | "(s, t) ∈ mstep S (Rω)" by auto
  then show "(s, t) ∈ GROUND (((mstep S Sω))=)"
  proof (cases)
    case 1
    from ground gtotal have "s ≻ t ∨ t ≻ s ∨ s = t" by auto
    from 1 rstepE obtain C σ l r where C: "(l, r) ∈ Eω" "s = C⟨l ⋅ σ⟩" "t = C⟨r ⋅ σ⟩"
      unfolding mstep_def by blast
    from ground [unfolded C ground_ctxt_apply] have "ground (l ⋅ σ)" and "ground (r ⋅ σ)" by auto
    from gtotal [OF this] have oriented: "l ⋅ σ ≻ r ⋅ σ ∨ r ⋅ σ ≻ l ⋅ σ ∨ l ⋅ σ = r ⋅ σ" by auto
    from C(1) have rl: "(r, l) ∈ Eω" by auto
    have eq: "l ⋅ σ = r ⋅ σ ⟹ s = t" unfolding C by auto
    have lr: "l ⋅ σ ≻ r ⋅ σ ⟹ (l ⋅ σ, r ⋅ σ) ∈ Sω" unfolding Sw_def
      by(rule UnI2, unfold mem_Collect_eq split, insert 1 C(1), auto)
    have rl: "r ⋅ σ ≻ l ⋅ σ ⟹ (r ⋅ σ, l ⋅ σ) ∈ Sω" unfolding Sw_def
      by(rule UnI2, unfold mem_Collect_eq split, insert 1 C(1), auto)
    { assume "l ⋅ σ ≻ r ⋅ σ ∨ r ⋅ σ ≻ l ⋅ σ"
      with lr rl have "(l ⋅ σ, r ⋅ σ) ∈ Sω" by auto
      with rstepI have "(s, t) ∈ (rstep Sω)" unfolding C by auto
      with dom have "(s, t) ∈ (mstep S Sω)" unfolding mstep_def mem_Collect_eq split by auto
      with ground have ?thesis unfolding GROUND_def by auto
    }
    with oriented eq ground show ?thesis using GROUND_def by blast
  next
    case 2
    with mstep [unfolded mstep_def] have "(s, t) ∈ (mstep S Sω)"
      unfolding Sw_def mstep_Un mstep_def Un_iff converse_iff mem_Collect_eq split by auto
    with ground show ?thesis unfolding GROUND_def by auto
  qed
qed
  
lemma ground_Einf_Sw:
  assumes "(s, t) ∈ (rstep E)" and gs: "ground s" and gt: "ground t"
  shows "(s, t) ∈ (mstep {#s, t#} Sω)*"
proof -
  { fix s t
    assume a: "(s, t) ∈ rstep E" and gs: "ground s" and gt: "ground t"
    define S where "S = {#s, t#}"
    with a have "(s, t) ∈ rstep (E ∪ R)" unfolding rstep_union by auto
    then have "(s, t) ∈ mstep {#s, t#} (E ∪ R)" unfolding mstep_def by auto
    with ERi_subset_ERw have "(s, t) ∈ (mstep S (Rω ∪ Eω))*" unfolding S_def by auto
    from this [unfolded conversion_def] have "(s, t) ∈ (mstep S Sω)*"
    proof (induct)
      case base
      then show ?case by auto
    next
      case (step u v)
      note rsteps = step(2) [unfolded mstep_def Un_iff converse_iff mem_Collect_eq split]
      then have st_uv: "(s ≽ u ∨ t ≽ u) ∧ (s ≽ v ∨ t ≽ v)" unfolding S_def by force
      from ground_less have ground_leq: "⋀ s t. ground s ⟹ s ≽ t ⟹ ground t" by blast
      from st_uv ground_leq [OF gs] ground_leq [OF gt] have "ground u ∧ ground v" by metis
      with step(2) have "(u, v) ∈ GROUND ((mstep S (Rω ∪ Eω)))" unfolding GROUND_def by auto
      with gmstep_ERw_Sw have uv: "(u, v) ∈ ((mstep S Sω))=" unfolding GROUND_def by auto
      with step(3) rtrancl_trans [OF step(3) [unfolded conversion_def] r_into_rtrancl] show ?case
        unfolding conversion_def by fast
    qed
    then have "(s, t) ∈ (mstep {#s, t#} Sω)*" unfolding S_def by auto
  } note estep = this
  from assms(1) estep [OF _ gs gt] estep [OF _ gt gs] show ?thesis
    unfolding Un_iff add_mset_commute [of t s "{#}"] converse_iff conversion_inv by auto
qed

end

locale ordered_completion_inf =
  ordered_completion less for less :: "('a, 'b::infinite) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50)
begin

lemma less_set_permute:
  "π ∙ {≻} = {≻}"
  apply (auto simp: )
   apply (metis inv_rule_mem_trs_simps(1) local.subst mem_Collect_eq old.prod.case term_apply_subst_Var_Rep_perm term_pt.permute_minus_cancel(1))
  by (metis case_prodI inv_rule_mem_trs_simps(1) mem_Collect_eq reduction_order.subst reduction_order_axioms term_apply_subst_Var_Rep_perm)

lemma ostep_permute:
  "(π ∙ s, π ∙ t) ∈ ostep (π ∙ E) (π ∙ R) ⟷ (s, t) ∈ ostep E R"
  using ordstep_permute [of π s t "{≻}" "E"]
  unfolding less_set_permute
  by (simp add: eqvt [symmetric] ostep_def)

lemma ordstep_permute_litsim:
  assumes "(s, t) ∈ ordstep {≻} R" and "R ≐ R'"
  shows "(π ∙ s, π ∙ t) ∈ ordstep {≻} R'"
  using assms(1)
proof(cases)
  case (1 l r C σ)
  from litsim_mem[OF assms(2) 1(1)] obtain ψ where rule:"(ψ ∙ l, ψ ∙ r) ∈ R'" by (auto simp:eqvt)
  define D and τ where "D = π ∙ C" and "τ = sop (-ψ) ∘s σ ∘s sop π"
  with 1 have s:"π ∙ s = D⟨(ψ ∙ l) ⋅ τ⟩" and t:"π ∙ t = D⟨(ψ ∙ r) ⋅ τ⟩"
    by (auto simp: eqvt [symmetric] term_pt.permute_flip)
  from 1(4) subst less_set_permute have "ψ ∙ l ⋅ τ ≻ ψ ∙ r ⋅ τ" unfolding τ_def by fastforce
  with ordstep.intros[OF rule s t] show ?thesis by auto
qed

lemma encompeq_permute': "s ⊴⋅ t ⟹ (π ∙ s ⊴⋅ ψ ∙ t)"
proof-
  assume "s ⊴⋅ t"
  then obtain C σ where t:"t = C⟨s⋅σ⟩" by auto
  define D and τ where "D = ψ ∙ C" and "τ = sop (-π) ∘s σ ∘s sop ψ"
  with t have t:"ψ ∙ t = D⟨(π ∙ s) ⋅ τ⟩" by (auto simp: eqvt [symmetric] term_pt.permute_flip)
  thus "π ∙ s ⊴⋅ ψ ∙ t" by blast
qed

lemma encompeq_permute: "s ⊴⋅ t = (π ∙ s ⊴⋅ ψ ∙ t)"
  by (metis encompeq_permute' term_pt.permute_minus_cancel(2))

lemma encomp_permute: "s ⊲⋅ t = (π ∙ s ⊲⋅ ψ ∙ t)"
  using encompeq_permute unfolding encomp_def by auto

lemma encstep1_permute_litsim:
  assumes "(s, t) ∈ encstep1 E R" and "E ≐ E'" and "R ≐ R'"
  shows "(π ∙ s, π ∙ t) ∈ encstep1 E' R'"
  using assms(1)
proof(cases)
  case (estep l r C σ)
  from litsim_mem[OF litsim_symcl[OF assms(2)] estep(1)] obtain ψ where eq:"(ψ ∙ l, ψ ∙ r) ∈ E'" by (auto simp:eqvt)
  define D and τ where "D = π ∙ C" and "τ = sop (-ψ) ∘s σ ∘s sop π"
  with estep have s:"π ∙ s = D⟨(ψ ∙ l) ⋅ τ⟩" and t:"π ∙ t = D⟨(ψ ∙ r) ⋅ τ⟩"
    by (auto simp: eqvt [symmetric] term_pt.permute_flip)
  from estep(4) subst less_set_permute have less:"ψ ∙ l ⋅ τ ≻ ψ ∙ r ⋅ τ" unfolding τ_def by fastforce
  from estep(5) encomp_permute have "ψ ∙ l ⊲⋅ π ∙ s" by auto
  with encstep1.intros(1)[OF eq s t less] show ?thesis .
next
  case rstep
  with litsim_rstep_eq[OF assms(3)] perm_rstep_conv encstep1.intros show ?thesis by auto
qed

lemma encstep2_permute_litsim:
  assumes "(s, t) ∈ encstep2 E R" and "E ≐ E'" and "R ≐ R'"
  shows "(π ∙ s, π ∙ t) ∈ encstep2 E' R'"
  using assms(1)
proof(cases)
  case (estep l r C σ)
  from litsim_mem[OF litsim_symcl[OF assms(2)] estep(1)] obtain ψ where eq:"(ψ ∙ l, ψ ∙ r) ∈ E'" by (auto simp:eqvt)
  define D and τ where "D = π ∙ C" and "τ = sop (-ψ) ∘s σ ∘s sop π"
  with estep have s:"π ∙ s = D⟨(ψ ∙ l) ⋅ τ⟩" and t:"π ∙ t = D⟨(ψ ∙ r) ⋅ τ⟩"
    by (auto simp: eqvt [symmetric] term_pt.permute_flip)
  from estep(4) subst less_set_permute have less:"ψ ∙ l ⋅ τ ≻ ψ ∙ r ⋅ τ" unfolding τ_def by fastforce
  from estep(5) encomp_permute have "ψ ∙ l ⊲⋅ π ∙ s" by auto
  with encstep2.intros(1)[OF eq s t less] show ?thesis .
next
  case (rstep l r C σ)
  from litsim_mem[OF assms(3) rstep(1)] obtain ψ where rule:"(ψ ∙ l, ψ ∙ r) ∈ R'" by (auto simp:eqvt)
  define D and τ where "D = π ∙ C" and "τ = sop (-ψ) ∘s σ ∘s sop π"
  with rstep have s:"π ∙ s = D⟨(ψ ∙ l) ⋅ τ⟩" and t:"π ∙ t = D⟨(ψ ∙ r) ⋅ τ⟩"
    by (auto simp: eqvt [symmetric] term_pt.permute_flip)
  from rstep(4) encomp_permute have "ψ ∙ l ⊲⋅ π ∙ s" by auto
  with encstep2.intros(2)[OF rule s t] show ?thesis .
qed

lemma ostep_permute_litsim:
  assumes "(s, t) ∈ ostep E R" and "E ≐ E'" and "R ≐ R'"
  shows "(π ∙ s, π ∙ t) ∈ ostep E' R'"
  using assms ordstep_permute_litsim [of s t, OF _ litsim_symcl[OF assms(2)]]
  unfolding ostep_def
  by (metis UnE UnI1 UnI2 litsim_rstep_eq rstep_permute_iff)

lemma oKB_permute:
  assumes "(E, R) ⊢oKB (E', R')"
  shows "(π ∙ E, π ∙ R) ⊢oKB (π ∙ E', π ∙ R')"
  using assms
proof (cases)
  case (deduce s t u)
  with oKB.deduce [of "π ∙ s" "π ∙ t" "π ∙ R" "π ∙ E" "π ∙ u"] show ?thesis
    by (simp add: eqvt) (metis (no_types, lifting) rstep_permute rstep_simps(5) rstep_union)
next
  case (orientl s t)
  with oKB.orientl [of "π ∙ s" "π ∙ t" "π ∙ E" "π ∙ R"] show ?thesis
    by (simp add: eqvt) (metis local.subst term_apply_subst_Var_Rep_perm)
next
  case (orientr t s)
  with oKB.orientr [of "π ∙ t" "π ∙ s" "π ∙ E" "π ∙ R"] show ?thesis
    by (simp add: eqvt) (metis local.subst term_apply_subst_Var_Rep_perm)
next
  case (delete s)
  with oKB.delete [of "π ∙ s" "π ∙ E" "π ∙ R"] show ?thesis by (simp add: eqvt)
next
  case (compose t u s)
  with oKB.compose [of "π ∙ t" "π ∙ u" "π ∙ E" "π ∙ R" "π ∙ s"] show ?thesis
    using ostep_permute [of π t u E "R - {(s, t)}"] by (simp add: eqvt)
next
  case (simplifyl s u t)
  with oKB.simplifyl [of "π ∙ s" "π ∙ u" "π ∙ E" "π ∙ t" "π ∙ R"] show ?thesis
    using ostep_permute [of π s u "E - {(s, t)}" R] by (simp add: eqvt)
next
  case (simplifyr t u s)
  with oKB.simplifyr [of "π ∙ t" "π ∙ u" "π ∙ E" "π ∙ s" "π ∙ R"] show ?thesis
    using ostep_permute [of π t u "E - {(s, t)}" R] by (simp add: eqvt)
next
  case (collapse t u s)
  with oKB.collapse [of "π ∙ t" "π ∙ u" "π ∙ E" "π ∙ R" "π ∙ s"] show ?thesis
    using ostep_permute [of π t u "E" "R - {(t, s)}"] by (simp add: eqvt)
qed

lemma oKB_permute':
  assumes "(π ∙ E, π ∙ R) ⊢oKB (π ∙ E', π ∙ R')"
  shows "(E, R) ⊢oKB (E', R')"
  using oKB_permute[OF assms, of "- π"] by auto

inductive oKB' :: "('a, 'b) trs × ('a, 'b) trs ⇒ ('a, 'b) trs × ('a, 'b) trs ⇒ bool"  (infix "⊢oKBπ" 55)
  where
    deduce: "(s, t) ∈ rstep (R ∪ E) ⟹ (s, u) ∈ rstep (R ∪ E) ⟹ oKB' (E, R) (E ∪ {(p ∙ t, p ∙ u)}, R)" |
    orientl: "s ≻ t ⟹ (s, t) ∈ E ⟹ oKB' (E, R) (E - {(s, t)}, R ∪ {(p ∙ s, p ∙ t)})" |
    orientr: "t ≻ s ⟹ (s, t) ∈ E ⟹ oKB' (E, R) (E - {(s, t)}, R ∪ {(p ∙ t, p ∙ s)})" |
    delete: "(s, s) ∈ E ⟹ oKB' (E, R) (E - {(s, s)}, R)" |
    compose: "(t, u) ∈ ostep E (R - {(s, t)}) ⟹ (s, t) ∈ R ⟹ oKB' (E, R) (E, (R - {(s, t)}) ∪ {(p ∙ s, p ∙ u)})" |
    simplifyl: "(s, u) ∈ ostep (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ oKB' (E, R) ((E - {(s, t)}) ∪ {(p ∙ u, p ∙ t)}, R)" |
    simplifyr: "(t, u) ∈ ostep (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ oKB' (E, R) ((E - {(s, t)}) ∪ {(p ∙ s, p ∙ u)}, R)" |
    collapse: "(t, u) ∈ ostep E (R - {(t, s)}) ⟹ (t, s) ∈ R ⟹ oKB' (E, R) (E ∪ {(p ∙ u, p ∙ s)}, R - {(t, s)})"

inductive
  oKBi' :: "('a, 'b) trs × ('a, 'b) trs ⇒ ('a, 'b) trs × ('a, 'b) trs ⇒ bool"  (infix "⊢oKBπ" 55)
  where
    deduce: "(s, t) ∈ rstep (R ∪ E) ⟹ (s, u) ∈ rstep (R ∪ E) ⟹ oKBi' (E, R) (E ∪ {(p ∙ t, p ∙ u)}, R)" |
    orientl: "s ≻ t ⟹ (s, t) ∈ E ⟹ oKBi' (E, R) (E - {(s, t)}, R ∪ {(p ∙ s, p ∙ t)})" |
    orientr: "t ≻ s ⟹ (s, t) ∈ E ⟹ oKBi' (E, R) (E - {(s, t)}, R ∪ {(p ∙ t, p ∙ s)})" |
    delete: "(s, s) ∈ E ⟹ oKBi' (E, R) (E - {(s, s)}, R)" |
    compose: "(t, u) ∈ ostep E (R - {(s, t)}) ⟹ (s, t) ∈ R ⟹ oKBi' (E, R) (E, (R - {(s, t)}) ∪ {(p ∙ s, p ∙ u)})" |
    simplifyl: "(s, u) ∈ encstep1 (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ oKBi' (E, R) ((E - {(s, t)}) ∪ {(p ∙ u, p ∙ t)}, R)" |
    simplifyr: "(t, u) ∈ encstep1 (E - {(s, t)}) R ⟹ (s, t) ∈ E ⟹ oKBi' (E, R) ((E - {(s, t)}) ∪ {(p ∙ s, p ∙ u)}, R)" |
    collapse: "(t, u) ∈ encstep2 E (R - {(t, s)}) ⟹ (t, s) ∈ R ⟹ oKBi' (E, R) (E ∪ {(p ∙ u, p ∙ s)}, R - {(t, s)})"

lemma oKB_step_variant_free':
  assumes "(E0, R0) ⊢oKB (E1, R1)"
    "variant_free_trs E0" and "variant_free_trs R0" and "¬ (variant_free_trs E1 ∧ variant_free_trs R1)"
  shows "∃E0' R0' E1' R1'. (E0', R0') ⊢oKB (E1', R1') ∧ E0' ≐ E0 ∧ R0' ≐ R0 ∧ E1' ≐ E1 ∧ R1' ≐ R1 ∧
         variant_free_trs E0' ∧ variant_free_trs R0' ∧ variant_free_trs E1' ∧ variant_free_trs R1'"
proof-
  have insert:"⋀S S' π s t. S' ≐ S ⟹ S ∪ {(π ∙ s, π ∙ t)} ≐ S' ∪ {(s, t)}"
    using litsim_insert[OF _ rule_pt.permute_prod_eqvt[symmetric]] subsumable_trs.litsim_sym by fastforce
  { fix l r and S :: "('a, 'b) trs"
    assume v:"variant_free_trs S" and lr:"(l,r) ∈ S"
    from lr v[unfolded variant_free_trs_def, rule_format, of l r] have "∀p. (p ∙ l, p ∙ r) ∉ S - {(l,r)}" 
      by fastforce
    hence "∀p p'. (p ∙ (p' ∙ l), p ∙ (p' ∙ r)) ∉ S - {(l,r)}" by (metis term_pt.permute_plus)
    with variant_free_trs_insert[OF variant_free_trs_diff, OF v]
      have "⋀π. variant_free_trs (S - {(l,r)} ∪ {(π ∙ l, π ∙ r)})" by auto
  }
  note vf = this
  note vfE = assms(2)[unfolded variant_free_trs_def, rule_format]
  note vfR = assms(3)[unfolded variant_free_trs_def, rule_format]
  { fix l r and S :: "('a, 'b) trs"
    assume lr:"(l,r) ∈ S"
    from litsim_insert[OF subsumable_trs.litsim_refl, of _ _ "(l,r)" "S - {(l, r)}"] insert_Diff[OF lr]
      have "⋀π. S ≐ S - {(l,r)} ∪ {(π ∙ l, π ∙ r)}"  by (auto simp:eqvt)
    hence "⋀π. S - {(l,r)} ∪ {(π ∙ l, π ∙ r)} ≐ S" using subsumable_trs.litsim_sym by fast
  }
  note permute_replace = this
  note eqrel = subsumable_trs.litsim_trans subsumable_trs.litsim_sym subsumable_trs.litsim_refl
  show ?thesis using assms proof(cases)
    case (deduce s t u)
    with assms(3) assms(2) assms(4) have "¬ (variant_free_trs E1)" by auto
    with assms(2) variant_free_trs_insert have "∃ρ. (ρ ∙ t, ρ ∙ u) ∈ E0" unfolding deduce by fastforce
    then obtain ρ where ρ:"(ρ ∙ t, ρ ∙ u) ∈ E0" by auto
    from insert_absorb[OF ρ] insert[OF eqrel(3), of E0 ρ t u] have litsim:"E0 ≐ E1" unfolding deduce(1) by force
    from deduce rstep_permute have "(ρ ∙ s, ρ ∙ t) ∈ rstep (R0 ∪ E0)" "(ρ ∙ s, ρ ∙ u) ∈ rstep (R0 ∪ E0)" by auto
    from oKBi.deduce[OF this] insert_absorb[OF ρ] have step:"(E0, R0) ⊢oKB (E0, R0)" by auto
    show ?thesis by (rule exI[of _ E0], rule exI[of _ R0], rule exI[of _ E0], rule exI[of _ R0],
                     insert step assms eqrel litsim deduce(2), auto)
next
  case (orientl s t)
    with assms(3) assms(2) assms(4) variant_free_trs_diff have "¬ (variant_free_trs R1)" by auto
    with assms(3) variant_free_trs_insert have "∃ρ. (ρ ∙ s, ρ ∙ t) ∈ R0" unfolding orientl by fastforce
    then obtain ρ where ρ:"(ρ ∙ s, ρ ∙ t) ∈ R0" by auto
    from insert_absorb[OF ρ] insert[OF eqrel(3), of R0 ρ s t] have litsimR:"R0 ≐ R1" unfolding orientl by force
    let ?E0 = "E0 - {(s,t)} ∪ {(ρ ∙ s, ρ ∙ t)}"
    from vfE[of s t] orientl(4) have "(ρ ∙ s, ρ ∙ t) ∉ E0 ∨ (ρ ∙ s, ρ ∙ t) = (s,t)" by metis
    hence E:"?E0 - {(ρ ∙ s, ρ ∙ t)} = E1" unfolding orientl(1) by blast
    have mem:"(ρ ∙ s, ρ ∙ t) ∈ ?E" by auto
    from orientl(3) less_set_permute have "ρ ∙ s ≻ ρ ∙ t" by auto
    from oKBi.orientl[OF this mem, of R0] insert_absorb[OF ρ] have step:"(?E, R0) ⊢oKB (E1, R0)" unfolding E by auto
    have litsimE:"?E ≐ E0" by (meson orientl(4) permute_replace eqrel)
    from vf[OF assms(2) orientl(4)] have vf0:"variant_free_trs ?E0" by auto
    from variant_free_trs_diff[OF assms(2)] orientl(1) have vf1:"variant_free_trs E1" by auto
    show ?thesis by (rule exI[of _ ?E0], rule exI[of _ R0], rule exI[of _ E1], rule exI[of _ R0],
                     insert step litsimE litsimR eqrel assms(2) assms(3) vf0 vf1, auto)
next
  case (orientr t s)
    with assms(3) assms(2) assms(4) variant_free_trs_diff have "¬ (variant_free_trs R1)" by auto
    with assms(3) variant_free_trs_insert have "∃ρ. (ρ ∙ t, ρ ∙ s) ∈ R0" unfolding orientr by fastforce
    then obtain ρ where ρ:"(ρ ∙ t, ρ ∙ s) ∈ R0" by auto
    from insert_absorb[OF ρ] insert[OF eqrel(3), of R0 ρ t s] have litsimR:"R0 ≐ R1" unfolding orientr by force
    let ?E0 = "E0 - {(s,t)} ∪ {(ρ ∙ s, ρ ∙ t)}"
    from vfE[of s t] orientr(4) have "(ρ ∙ s, ρ ∙ t) ∉ E0 ∨ (ρ ∙ s, ρ ∙ t) = (s,t)" by metis
    hence E:"?E0 - {(ρ ∙ s, ρ ∙ t)} = E1" unfolding orientr(1) by blast
    have mem:"(ρ ∙ s, ρ ∙ t) ∈ ?E" by auto
    from orientr(3) less_set_permute have "ρ ∙ t ≻ ρ ∙ s" by auto
    from oKBi.orientr[OF this mem, of R0] insert_absorb[OF ρ] have step:"(?E, R0) ⊢oKB (E1, R0)" unfolding E by auto
    have litsimE:"?E ≐ E0" by (meson orientr(4) permute_replace eqrel)
    from vf[OF assms(2) orientr(4)] have vf0:"variant_free_trs ?E0" by auto
    from variant_free_trs_diff[OF assms(2)] orientr(1) have vf1:"variant_free_trs E1" by auto
    show ?thesis by (rule exI[of _ ?E0], rule exI[of _ R0], rule exI[of _ E1], rule exI[of _ R0],
                     insert step litsimE litsimR eqrel assms(2) assms(3) vf0 vf1, auto)
next
case (delete s)
  with variant_free_trs_diff[OF assms(2), of "{(s,s)}"] assms(3) assms(4) show ?thesis by fast
next
  case (compose t u s)
    with assms(3) assms(2) assms(4) variant_free_trs_diff have "¬ (variant_free_trs R1)" by auto
    with assms(3) variant_free_trs_insert[OF variant_free_trs_diff[OF assms(3)], of s u "{(s, t)}"]
      have "∃ρ. (ρ ∙ s, ρ ∙ u) ∈ R0 - {(s,t)}" unfolding compose by auto
    then obtain ρ where ρ:"(ρ ∙ s, ρ ∙ u) ∈ R0 - {(s,t)}" by auto
    define R0' where "R0' = R0 - {(s, t)} ∪ {(ρ ∙ s,ρ ∙ t)}"
    hence mem':"(ρ ∙ s, ρ ∙ t) ∈ R0'" by auto
    from vf[OF assms(3)] compose(4) have vf0:"variant_free_trs R0'" unfolding R0'_def by auto
    define R1' where "R1' = R0 - {(s, t)}"
    from vfR[of s t] compose(4) have "(ρ ∙ s, ρ ∙ t) ∉ R0 ∨ (ρ ∙ s, ρ ∙ t) = (s,t)" by metis
    hence R:"R0' - {(ρ ∙ s, ρ ∙ t)} = R1'" unfolding compose(2) R0'_def R1'_def by blast
    from R eqrel(3) R1'_def have R':"R0 - {(s, t)} ≐ R0' - {(ρ ∙ s, ρ ∙ t)}" by auto
    note ostep = ostep_permute_litsim[OF compose(3) eqrel(3) this] 
    from oKBi.compose[OF ostep] insert_absorb[OF ρ] R0'_def have step:"(E0, R0') ⊢oKB (E1, R1')"
      unfolding compose(1) R1'_def R by auto
    from vf[OF assms(3) compose(4)] have vf0:"variant_free_trs R0'" unfolding R0'_def by blast
    from variant_free_trs_diff[OF assms(3)] have vf1:"variant_free_trs R1'" unfolding R1'_def by blast
    from permute_replace[OF compose(4)] have litsim0:"R0' ≐ R0" unfolding R0'_def by blast
    from insert_absorb[OF ρ] insert[OF eqrel(3), of "R0 - {(s, t)}" ρ s u] have litsim1:"R1' ≐ R1"
      unfolding R1'_def compose(2) by force
    show ?thesis by (rule exI[of _ E0], rule exI[of _ R0'], rule exI[of _ E0], rule exI[of _ R1'],
          insert step compose(1) vf0 vf1 litsim0 litsim1 assms(2), auto)
next
  case (simplifyl s u t)
    with assms(3) assms(2) assms(4) have nvf:"¬ (variant_free_trs E1)" by auto
    note vf' = variant_free_trs_insert[OF variant_free_trs_diff[OF assms(2)], of u t "{(s, t)}"]
    with nvf have "∃ρ. (ρ ∙ u, ρ ∙ t) ∈ E0 - {(s, t)}" unfolding simplifyl by fastforce
    then obtain ρ where ρ:"(ρ ∙ u, ρ ∙ t) ∈ E0 - {(s, t)}" by auto
    define E0' where "E0' = E0 - {(s, t)} ∪ {(ρ ∙ s, ρ ∙ t)}"
    hence mem:"(ρ ∙ s, ρ ∙ t) ∈ E0'" by auto
    define E1' where "E1' = E0 - {(s, t)}"
    from vfE[of s t] simplifyl(4) have "(ρ ∙ s, ρ ∙ t) ∉ E0 ∨ (ρ ∙ s, ρ ∙ t) = (s,t)" by metis
    hence E:"E0' - {(ρ ∙ s, ρ ∙ t)} = E1'" unfolding simplifyl(1) E0'_def E1'_def by blast
    from E eqrel(3) E1'_def have E':"E0 - {(s, t)} ≐ E0' - {(ρ ∙ s, ρ ∙ t)}" by auto
    note encstep = encstep1_permute_litsim[OF simplifyl(3) this eqrel(3)]
    from oKBi.simplifyl[OF encstep mem] E insert_absorb[OF ρ] have step:"(E0', R0) ⊢oKB (E1', R1)"
      unfolding simplifyl(2) E1'_def by auto
    from vf[OF assms(2) simplifyl(4)] have vf0:"variant_free_trs E0'" unfolding E0'_def by blast
    from variant_free_trs_diff[OF assms(2)] have vf1:"variant_free_trs E1'" unfolding E1'_def by blast
    from permute_replace[OF simplifyl(4)] have litsim0:"E0' ≐ E0" unfolding E0'_def by blast
    from insert_absorb[OF ρ] insert[OF eqrel(3), of E1' ρ u t] have litsim1:"E1' ≐ E1"
      unfolding E1'_def simplifyl(1) by auto
    show ?thesis by (rule exI[of _ E0'], rule exI[of _ R0], rule exI[of _ E1'], rule exI[of _ R1],
          insert step vf0 vf1 litsim0 litsim1 assms(3), unfold simplifyl(2), auto)
next
  case (simplifyr t u s)
    with assms(3) assms(2) assms(4) have nvf:"¬ (variant_free_trs E1)" by auto
    note vf' = variant_free_trs_insert[OF variant_free_trs_diff[OF assms(2)], of s u "{(s, t)}"]
    with nvf have "∃ρ. (ρ ∙ s, ρ ∙ u) ∈ E0 - {(s, t)}" unfolding simplifyr by fastforce
    then obtain ρ where ρ:"(ρ ∙ s, ρ ∙ u) ∈ E0 - {(s, t)}" by auto
    define E0' where "E0' = E0 - {(s, t)} ∪ {(ρ ∙ s, ρ ∙ t)}"
    hence mem:"(ρ ∙ s, ρ ∙ t) ∈ E0'" by auto
    define E1' where "E1' = E0 - {(s, t)}"
    from vfE[of s t] simplifyr(4) have "(ρ ∙ s, ρ ∙ t) ∉ E0 ∨ (ρ ∙ s, ρ ∙ t) = (s,t)" by metis
    hence E:"E0' - {(ρ ∙ s, ρ ∙ t)} = E1'" unfolding simplifyl(1) E0'_def E1'_def by blast
    from E eqrel(3) E1'_def have E':"E0 - {(s, t)} ≐ E0' - {(ρ ∙ s, ρ ∙ t)}" by auto
    note encstep = encstep1_permute_litsim[OF simplifyr(3) this eqrel(3)] 
    from oKBi.simplifyr[OF encstep mem] E insert_absorb[OF ρ] have step:"(E0', R0) ⊢oKB (E1', R1)"
      unfolding simplifyr(2) E1'_def by auto
    from vf[OF assms(2) simplifyr(4)] have vf0:"variant_free_trs E0'" unfolding E0'_def by blast
    from variant_free_trs_diff[OF assms(2)] have vf1:"variant_free_trs E1'" unfolding E1'_def by blast
    from permute_replace[OF simplifyr(4)] have litsim0:"E0' ≐ E0" unfolding E0'_def by blast
    from insert_absorb[OF ρ] insert[OF eqrel(3), of E1' ρ s u] have litsim1:"E1' ≐ E1"
      unfolding E1'_def simplifyr(1) by auto
    show ?thesis by (rule exI[of _ E0'], rule exI[of _ R0], rule exI[of _ E1'], rule exI[of _ R1],
          insert step vf0 vf1 litsim0 litsim1 assms(3), unfold simplifyr(2), auto)
next
  case (collapse t u s)
    with assms(3) assms(2) assms(4) variant_free_trs_diff have "¬ (variant_free_trs E1)" by auto
    with assms(2) variant_free_trs_insert have "∃ρ. (ρ ∙ u, ρ ∙ s) ∈ E0" unfolding collapse by fastforce
    then obtain ρ where ρ:"(ρ ∙ u, ρ ∙ s) ∈ E0" by auto
    define R0' where "R0' = R0 - {(t, s)} ∪ {(ρ ∙ t, ρ ∙ s)}"
    hence mem:"(ρ ∙ t, ρ ∙ s) ∈ R0'" by auto
    from vfR[of t s] collapse(4) have R:"(ρ ∙ t, ρ ∙ s) ∉ R0 ∨ (ρ ∙ t, ρ ∙ s) = (t,s)" by metis
    hence R':"R0' - {(ρ ∙ t, ρ ∙ s)} = R1" unfolding collapse(2) R0'_def by blast
    from R have "R0 - {(t, s)} ≐ R0' - {(ρ ∙ t, ρ ∙ s)}" unfolding R0'_def by fastforce
    note encstep = encstep2_permute_litsim[OF collapse(3) eqrel(3) this] 
    from oKBi.collapse[OF encstep mem] insert_absorb[OF ρ] have step:"(E0, R0') ⊢oKB (E0, R1)"
      unfolding R' by auto
    from vf[OF assms(3) collapse(4)] have vf0:"variant_free_trs R0'" unfolding R0'_def by auto
    with R' variant_free_trs_diff have vf1:"variant_free_trs R1" by blast
    from collapse(4) permute_replace have litsimR:"R0' ≐ R0" unfolding R0'_def by auto
    from insert_absorb[OF ρ] insert[OF eqrel(3), of E0 ρ u s] have litsimE:"E0 ≐ E1" unfolding collapse(1) by auto
    show ?thesis by (rule exI[of _ E0], rule exI[of _ R0'], rule exI[of _ E0], rule exI[of _ R1],
          insert step vf0 vf1 litsimR litsimE assms(2) eqrel(3), auto)
  qed
qed

lemma oKB_step_variant_free:
  assumes "(E0, R0) ⊢oKB (E1, R1)"
    "variant_free_trs E0" and "variant_free_trs R0"
  shows "∃E0' R0' E1' R1'. (E0', R0') ⊢oKB (E1', R1') ∧ E0' ≐ E0 ∧ R0' ≐ R0 ∧ E1' ≐ E1 ∧ R1' ≐ R1 ∧
         variant_free_trs E0' ∧ variant_free_trs R0' ∧ variant_free_trs E1' ∧ variant_free_trs R1'"
proof(cases "¬(variant_free_trs E1 ∧ variant_free_trs R1)")
  case True
  then show ?thesis using oKB_step_variant_free'[OF assms] by auto
next
  case False            
  show ?thesis by (rule exI[of _ E0], rule exI[of _ R0], rule exI[of _ E1], rule exI[of _ R1], insert assms False, auto)
qed

lemma step_implies_litsim_step:
  assumes "(E0, R0) ⊢oKB (E, R)" and "R' ≐ R" and "E' ≐ E" and
    "variant_free_trs E'" and "variant_free_trs R'" and "variant_free_trs E" and "variant_free_trs R"
and "variant_free_trs E0" and "variant_free_trs R0" and "R0 ⊆ {≻}"
  shows "∃E0' R0'. (E0', R0') ⊢oKB (E', R') ∧ R0 ≐ R0' ∧ E0 ≐ E0' ∧ variant_free_trs E0' ∧ variant_free_trs R0'"
proof-
  have replace:"⋀S S' π s t. S' ≐ S ⟹ S ∪ {(π ∙ s, π ∙ t)} ≐ S' ∪ {(s, t)}"
    using litsim_insert[OF _ rule_pt.permute_prod_eqvt[symmetric]] subsumable_trs.litsim_sym by fastforce
  { fix l r and S :: "('a, 'b) trs"
    assume lr:"(l,r) ∈ S"
    from litsim_insert[OF subsumable_trs.litsim_refl, of _ _ "(l,r)" "S - {(l, r)}"] insert_Diff[OF lr]
      have "⋀π. S ≐ S - {(l,r)} ∪ {(π ∙ l, π ∙ r)}"  by (auto simp:eqvt)
    hence "⋀π. S - {(l,r)} ∪ {(π ∙ l, π ∙ r)} ≐ S" using subsumable_trs.litsim_sym by fast
  }
  note permute_replace = this
  note sym = subsumable_trs.litsim_sym
  { fix ρ t u E0 E0' E'
    assume ρ:"(ρ ∙ t, ρ ∙ u) ∈ E'" and E:"E0' = (if (t, u) ∈ E0 then E' else E' - {(ρ ∙ t, ρ ∙ u)})" and litsim:"E' ≐ E0 ∪ {(t, u)}"
    and vf:"variant_free_trs (E0 ∪ {(t, u)})" and vf':"variant_free_trs E'"
    have litsim:"E0 ≐ E0'" proof(cases "(t, u) ∈ E0")
      case True
      from sym True insert_absorb[OF True] insert_absorb[OF ρ] replace[OF litsim[THEN sym], of ρ t u]
      show "E0 ≐ E0'" unfolding E by force
    next
      case False
      with litsim_diff1[OF _ vf vf' _ _ ρ, of ρ "(t,u)"] litsim[THEN sym] show "E0 ≐ E0'"
        unfolding E using litsim by (auto simp:eqvt)
    qed
  } note Eaux = this
  note litsim_defs = subsumable_trs.litsim_def subsumeseq_trs_def
  { fix ψ s t E0 E E' G
    assume E:"E = E0 - {(s,t)} ∪ G" and G:"∀π.(π ∙ s,π ∙ t) ∉ G" and mem:"(s,t) ∈ E0" and litsim:"E' ≐ E" and
      vf:"variant_free_trs E0" and mem':"(ψ ∙ s, ψ ∙ t) ∈ E'" (is "(?s,?t) ∈ _")
    from mem' litsim obtain π where π:"(π ∙ ?s, π ∙ ?t) ∈ E" unfolding litsim_defs by (auto simp:eqvt)
    with E G[rule_format, of "π + ψ"] have "(π ∙ ?s, π ∙ ?t) ∈ E0" by force
    with vf[unfolded variant_free_trs_def] mem have "(π ∙ ?s, π ∙ ?t) = (s,t)" unfolding term_pt.permute_plus[symmetric] by metis
    with π E G[rule_format, of "π + ψ"] have False by force
  } note no_variant = this
  note no_variant' = no_variant[of _ _ _ _ "{}", simplified]
  { fix π u s
    fix z :: 'b
    assume less:"u ≻ s" and "u = π ∙ s"
    with subst[OF less[unfolded this], of "λ_. Var z"] irrefl have False
      by (simp add: comp_def permute_term_subst_apply_term)
  } note less_perm = this
  from assms show ?thesis proof(cases)
case (deduce s t u)
  with assms(3) obtain ρ where ρ:"(ρ ∙ t, ρ ∙ u) ∈ E'" unfolding litsim_defs by (auto simp:eqvt)
  define E0' where "E0' ≡ if (t, u) ∈ E0 then E' else E' - {(ρ ∙ t, ρ ∙ u)}"
  from Eaux[OF ρ _ assms(3)[unfolded deduce(1)] assms(6)[unfolded deduce(1)]] E0'_def assms(4)
    have litsim:"E0 ≐ E0'" by auto
  from litsim_rstep_eq[OF this] assms(2)[THEN litsim_rstep_eq] have rstep:"rstep (R0 ∪ E0) = rstep (R' ∪ E0')"
    unfolding deduce(2) rstep_union reverseTrs[symmetric] by argo
  with deduce rstep_permute_iff have "(ρ ∙ s, ρ ∙ t) ∈ rstep (R' ∪ E0')" "(ρ ∙ s, ρ ∙ u) ∈ rstep (R' ∪ E0')" by auto
  from oKBi.deduce[OF this] insert_absorb[OF ρ] have step:"(E0', R') ⊢oKB (E', R')"
    unfolding E0'_def by (cases "(t, u) ∈ E0", auto)
  from assms variant_free_trs_diff have vf:"variant_free_trs E0'" unfolding E0'_def by force
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R'], insert assms step litsim sym vf, unfold deduce(2), auto)
next
  case (orientl s t)
  note o = this
  with assms(2) obtain ρ where ρ:"(ρ ∙ s, ρ ∙ t) ∈ R'" unfolding litsim_defs by (auto simp:eqvt)
  define R0' where "R0' ≡ if (s, t) ∈ R0 then R' else R' - {(ρ ∙ s, ρ ∙ t)}"
  define E0' where "E0' ≡ E' ∪ {(ρ ∙ s, ρ ∙ t)}"
  from Eaux[OF ρ _ assms(2)[unfolded o(2)] assms(7)[unfolded o(2)] assms(5)] R0'_def have litsimR:"R0 ≐ R0'" by auto
  from less_set_permute o(3) have less:"ρ ∙ s ≻ ρ ∙ t" by auto
  note E'_variant = no_variant'[OF o(1) o(4) assms(3) assms(8)]
  hence E:"E' = E0' - {(ρ ∙ s, ρ ∙ t)}" unfolding E0'_def by force
  have R:"R' = R0' ∪ {(ρ ∙ s, ρ ∙ t)}" unfolding R0'_def using ρ by auto
  from oKBi.orientl[OF less, of E0' R0'] E0'_def have step:"(E0', R0') ⊢oKB (E',R')" unfolding E R by auto
  from replace[OF assms(3)[THEN sym], of ρ s t] o sym insert_Diff have litsimE:"E0 ≐ E0'"
    unfolding E0'_def by fastforce
  from E'_variant variant_free_trs_insert[OF assms(4), of "ρ ∙ s" "ρ ∙ t"] term_pt.permute_plus
    have vfE:"variant_free_trs E0'" unfolding E0'_def by (metis Un_insert_right sup_bot.right_neutral)
  from assms variant_free_trs_diff have vfR:"variant_free_trs R0'" unfolding R0'_def by force
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R0'], insert step litsimR litsimE vfR vfE, auto)
next
  case (orientr t s)
  note o = this
  with assms(2) obtain ρ where ρ:"(ρ ∙ t, ρ ∙ s) ∈ R'" unfolding litsim_defs by (auto simp:eqvt)
  define R0' where "R0' ≡ if (t, s) ∈ R0 then R' else R' - {(ρ ∙ t, ρ ∙ s)}"
  define E0' where "E0' ≡ E' ∪ {(ρ ∙ s, ρ ∙ t)}"
  from Eaux[OF ρ _ assms(2)[unfolded o(2)] assms(7)[unfolded o(2)] assms(5)] R0'_def have litsimR:"R0 ≐ R0'" by auto
  from less_set_permute o(3) have less:"ρ ∙ t ≻ ρ ∙ s" by auto
  note E'_variant = no_variant'[OF o(1) o(4) assms(3) assms(8)]
  hence E:"E' = E0' - {(ρ ∙ s, ρ ∙ t)}" unfolding E0'_def by force
  have R:"R' = R0' ∪ {(ρ ∙ t, ρ ∙ s)}" unfolding R0'_def using ρ by auto
  from oKBi.orientr[OF less, of E0' R0'] E0'_def have step:"(E0', R0') ⊢oKB (E',R')" unfolding E R by auto
  from replace[OF assms(3)[THEN sym], of ρ s t] o sym insert_Diff have litsimE:"E0 ≐ E0'"
    unfolding E0'_def by fastforce
  from E'_variant variant_free_trs_insert[OF assms(4), of "ρ ∙ s" "ρ ∙ t"] term_pt.permute_plus
    have vfE:"variant_free_trs E0'" unfolding E0'_def by (metis Un_insert_right sup_bot.right_neutral)
  from assms variant_free_trs_diff have vfR:"variant_free_trs R0'" unfolding R0'_def by force
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R0'], insert step litsimR litsimE vfR vfE, auto)
next
  case (delete s)
  define E0' where "E0' ≡ E' ∪ {(s, s)}"
  note E'_variant = no_variant'[OF delete(1) delete(3) assms(3) assms(8)]
  from this[of 0] have "E0' - {(s, s)} = E'" using E0'_def by auto
  with oKBi.delete[of s E0'] have step:"(E0', R') ⊢oKB (E', R')" using E0'_def by auto
  from E'_variant variant_free_trs_insert[OF assms(4)]
  have vfE:"variant_free_trs E0'" unfolding E0'_def by auto
  from replace[OF assms(3)[THEN sym], of 0 s s] sym insert_Diff delete have litsimE:"E0 ≐ E0'"
    unfolding E0'_def by fastforce
  note facts = step assms litsimE vfE sym
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R'], insert facts, unfold delete(2), auto)
next
  case (compose t u s)
  with assms(2) obtain ρ where ρ:"(ρ ∙ s, ρ ∙ u) ∈ R'" unfolding litsim_defs by (auto simp:eqvt)
  define R0' where "R0' ≡ if (s, u) ∈ R0 then R' ∪ {(ρ ∙ s, ρ ∙ t)} else R' - {(ρ ∙ s, ρ ∙ u)} ∪ {(ρ ∙ s, ρ ∙ t)}"
  have mem:"(ρ ∙ s, ρ ∙ t) ∈ R0'" unfolding R0'_def by auto
  have litsim:"R0 ≐ R0'" proof(cases "(s, u) ∈ R0")
    case True
    with sym insert_absorb[OF True] insert_absorb[OF ρ] replace[OF sym[OF assms(2)], of ρ ] compose(4)
    show ?thesis unfolding compose(2) unfolding R0'_def
      by (smt Un_empty_right Un_insert_right insert_Diff insert_commute)
  next
    case False
    from False compose have R0:"R0 = R - {(s,u)} ∪ {(s,t)}" by (cases "t = u", auto)
    from litsim_diff1[OF assms(2)[THEN sym] assms(7) assms(5) _ _ ρ, of ρ "(s,u)"] compose(2)
      have "R - {(s, u)} ≐ R' - {(ρ ∙ s, ρ ∙ u)}" by (auto simp:eqvt)
    from litsim_insert[OF this] show "R0 ≐ R0'"
      unfolding R0'_def if_not_P[OF False] unfolding R0 by (auto simp:eqvt)
  qed
  from ostep_imp_less[OF _ compose(3)] assms(10) have "t ≻ u" by auto
  from less_perm[OF this] have diff:"∀π. (π ∙ s, π ∙ t) ∉ {(s, u)}"
    by (metis Pair_inject singletonD term_pt.permute_minus_cancel(2))
  note no_st_variant =  no_variant[OF compose(2) diff compose(4) assms(2) assms(9)]
  have R':"R' = R0' - {(ρ ∙ s, ρ ∙ t)} ∪ {(ρ ∙ s, ρ ∙ u)}"
    using ρ no_st_variant[of ρ] unfolding R0'_def by force
  have vf0:"variant_free_trs R0'" proof(cases "(s, u) ∈ R0")
    case True
    with no_st_variant variant_free_trs_insert[OF assms(5)] show ?thesis unfolding R0'_def
      by (smt Un_insert_right sup_bot.right_neutral term_pt.permute_plus)
  next
    case False
    with no_st_variant variant_free_trs_insert[OF variant_free_trs_diff[OF assms(5)]] show ?thesis
      unfolding R0'_def by (smt DiffE Un_insert_right sup_bot.comm_neutral term_pt.permute_plus)
  qed
  from litsim_diff1[OF litsim assms(9) vf0 _ compose(4) mem] have "R0 - {(s, t)} ≐ R0' - {(ρ ∙ s, ρ ∙ t)}" by (auto simp:eqvt)
  from ostep_permute_litsim[OF compose(3) _ this] sym[OF assms(3)]
    have step:"(ρ ∙ t, ρ ∙ u) ∈ ostep E' (R0' - {(ρ ∙ s, ρ ∙ t)})" unfolding compose(1) by auto
  from oKBi.compose[OF step mem] R' have "(E', R0') ⊢oKB (E',R')" by auto
  note facts = this litsim assms(3)[unfolded compose(1), THEN sym] vf0 assms(4)
  show ?thesis by (rule exI[of _ E'], rule exI[of _ R0'], insert facts, auto)
next
  case (simplifyl s u t)
  with assms(3) obtain ρ where ρ:"(ρ ∙ u, ρ ∙ t) ∈ E'" unfolding litsim_defs by (auto simp:eqvt)
  define E0' where "E0' ≡ if (u, t) ∈ E0 then E' ∪ {(ρ ∙ s, ρ ∙ t)} else E' - {(ρ ∙ u, ρ ∙ t)} ∪ {(ρ ∙ s, ρ ∙ t)}"
  have mem:"(ρ ∙ s, ρ ∙ t) ∈ E0'" unfolding E0'_def by auto
  have litsimE:"E0 ≐ E0'" proof(cases "(u, t) ∈ E0")
    case True
    with sym insert_absorb[OF True] insert_absorb[OF ρ] replace[OF sym[OF assms(3)], of ρ s t] simplifyl(4)
    show "E0 ≐ E0'" unfolding simplifyl(1) unfolding E0'_def
      by (smt Un_empty_right Un_insert_right insert_Diff insert_commute)
  next
    case False
    from False simplifyl have E0:"E0 = E - {(u,t)} ∪ {(s,t)}" by (cases "s = u", auto)
    from litsim_diff1[OF assms(3)[THEN sym] assms(6) assms(4) _ _ ρ, of ρ "(u,t)"] simplifyl(1)
      have "E - {(u, t)} ≐ E' - {(ρ ∙ u, ρ ∙ t)}" by (auto simp:eqvt)
    from litsim_insert[OF this] show "E0 ≐ E0'"
      unfolding E0'_def if_not_P[OF False] unfolding E0 by (auto simp:eqvt)
  qed
  note less = encstep1_less[OF simplifyl(3) assms(10)]
  from less_perm[OF this] have diff:"∀π. (π ∙ s, π ∙ t) ∉ {(u, t)}"
    by (metis Pair_inject singletonD term_pt.permute_minus_cancel(2))
  note no_st_variant = no_variant[OF simplifyl(1) diff simplifyl(4) assms(3) assms(8)]
  have E':"E' = E0' - {(ρ ∙ s, ρ ∙ t)} ∪ {(ρ ∙ u, ρ ∙ t)}"
    using ρ no_st_variant[of ρ] unfolding E0'_def by force
  have vf0:"variant_free_trs E0'" proof(cases "(u, t) ∈ E0")
    case True
    with no_st_variant variant_free_trs_insert[OF assms(4)] show ?thesis unfolding E0'_def
      by (smt Un_insert_right sup_bot.right_neutral term_pt.permute_plus)
  next
    case False
    with no_st_variant variant_free_trs_insert[OF variant_free_trs_diff[OF assms(4)]] show ?thesis
      unfolding E0'_def by (smt DiffE Un_insert_right sup_bot.comm_neutral term_pt.permute_plus)
  qed
  from litsim_diff1[OF litsimE assms(8) vf0 _ simplifyl(4) mem] have "E0 - {(s, t)} ≐ E0' - {(ρ ∙ s, ρ ∙ t)}" by (auto simp:eqvt)
  from encstep1_permute_litsim[OF simplifyl(3) this, of R'] sym[OF assms(2)]
    have step:"(ρ ∙ s, ρ ∙ u) ∈ encstep1 (E0' - {(ρ ∙ s, ρ ∙ t)}) R'" unfolding simplifyl(2) by auto
  from oKBi.simplifyl[OF step mem] E' have "(E0', R') ⊢oKB (E',R')" by auto
  note facts = this litsimE assms(2)[unfolded simplifyl(2), THEN sym] vf0 assms(5)
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R'], insert facts, auto)
next
  case (simplifyr t u s)
  with assms(3) obtain ρ where ρ:"(ρ ∙ s, ρ ∙ u) ∈ E'" unfolding litsim_defs by (auto simp:eqvt)
  define E0' where "E0' ≡ if (s, u) ∈ E0 then E' ∪ {(ρ ∙ s, ρ ∙ t)} else E' - {(ρ ∙ s, ρ ∙ u)} ∪ {(ρ ∙ s, ρ ∙ t)}"
  have mem:"(ρ ∙ s, ρ ∙ t) ∈ E0'" unfolding E0'_def by auto
  have litsimE:"E0 ≐ E0'" proof(cases "(s, u) ∈ E0")
    case True
    with sym insert_absorb[OF True] insert_absorb[OF ρ] replace[OF sym[OF assms(3)], of ρ s t] simplifyr(4)
    show "E0 ≐ E0'" unfolding simplifyr(1) unfolding E0'_def
      by (smt Un_empty_right Un_insert_right insert_Diff insert_commute)
  next
    case False
    from False simplifyr have E0:"E0 = E - {(s,u)} ∪ {(s,t)}" by (cases "t = u", auto)
    from litsim_diff1[OF assms(3)[THEN sym] assms(6) assms(4) _ _ ρ, of ρ "(s,u)"] simplifyr(1)
      have "E - {(s, u)} ≐ E' - {(ρ ∙ s, ρ ∙ u)}" by (auto simp:eqvt)
    from litsim_insert[OF this] show "E0 ≐ E0'"
      unfolding E0'_def if_not_P[OF False] unfolding E0 by (auto simp:eqvt)
  qed
  note less = encstep1_less[OF simplifyr(3) assms(10)]
  from less_perm[OF this] have diff:"∀π. (π ∙ s, π ∙ t) ∉ {(s, u)}"
    by (metis Pair_inject singletonD term_pt.permute_minus_cancel(2))
  note no_st_variant = no_variant[OF simplifyr(1) diff simplifyr(4) assms(3) assms(8)]
  have E':"E' = E0' - {(ρ ∙ s, ρ ∙ t)} ∪ {(ρ ∙ s, ρ ∙ u)}"
    using ρ no_st_variant[of ρ] unfolding E0'_def by force
  have vf0:"variant_free_trs E0'" proof(cases "(s, u) ∈ E0")
    case True
    with no_st_variant variant_free_trs_insert[OF assms(4)] show ?thesis unfolding E0'_def
      by (smt Un_insert_right sup_bot.right_neutral term_pt.permute_plus)
  next
    case False
    with no_st_variant variant_free_trs_insert[OF variant_free_trs_diff[OF assms(4)]] show ?thesis
      unfolding E0'_def by (smt DiffE Un_insert_right sup_bot.comm_neutral term_pt.permute_plus)
  qed
  from litsim_diff1[OF litsimE assms(8) vf0 _ simplifyr(4) mem] have "E0 - {(s, t)} ≐ E0' - {(ρ ∙ s, ρ ∙ t)}" by (auto simp:eqvt)
  from encstep1_permute_litsim[OF simplifyr(3) this, of R'] sym[OF assms(2)]
    have step:"(ρ ∙ t, ρ ∙ u) ∈ encstep1 (E0' - {(ρ ∙ s, ρ ∙ t)}) R'" unfolding simplifyr(2) by auto
  from oKBi.simplifyr[OF step mem] E' have "(E0', R') ⊢oKB (E',R')" by auto
  note facts = this litsimE assms(2)[unfolded simplifyr(2), THEN sym] vf0 assms(5)
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R'], insert facts, auto)
next
  case (collapse t u s)
  note c = this
  with assms(3) obtain ρ where ρ:"(ρ ∙ u, ρ ∙ s) ∈ E'" unfolding litsim_defs by (auto simp:eqvt)
  define E0' where "E0' ≡ if (u, s) ∈ E0 then E' else E' - {(ρ ∙ u, ρ ∙ s)}"
  define R0' where "R0' ≡ R' ∪ {(ρ ∙ t, ρ ∙ s)}"
  from Eaux[OF ρ _ assms(3)[unfolded c(1)] assms(6)[unfolded c(1)] assms(4)] E0'_def have litsimE:"E0 ≐ E0'" by auto
  from c have R0:"R0 = R ∪ {(t,s)}" by fast
  from litsim_insert[OF sym[OF assms(2)]] have litsimR:"R0 ≐ R0'" unfolding R0'_def R0 by (auto simp:eqvt)
  from variant_free_trs_diff assms(4) have vfE:"variant_free_trs E0'" unfolding E0'_def by auto
  note no_variant' = no_variant'[OF c(2) c(4) assms(2) assms(9)]
  from no_variant' variant_free_trs_insert[OF assms(5), of "ρ ∙ t" "ρ ∙ s"] have vfR:"variant_free_trs R0'"
    unfolding R0'_def using term_pt.permute_plus by (metis Un_insert_right sup_bot.right_neutral)
  from no_variant' have R:"R' = R0' - {(ρ ∙ t, ρ ∙ s)}" unfolding R0'_def by force
  have E:"E' = E0' ∪ {(ρ ∙ u, ρ ∙ s)}" unfolding E0'_def using ρ by auto
  from assms(2)[unfolded R c(2)] sym have "R0 - {(t, s)} ≐ R0' - {(ρ ∙ t, ρ ∙ s)}" by auto
  from encstep2_permute_litsim[OF collapse(3) litsimE this]
    have ostep:"(ρ ∙ t, ρ ∙ u) ∈ encstep2 E0' (R0' - {(ρ ∙ t, ρ ∙ s)})" by auto
  from oKBi.collapse[OF ostep] R0'_def have step:"(E0', R0') ⊢oKB (E',R')" unfolding E R by auto
  show ?thesis by (rule exI[of _ E0'], rule exI[of _ R0'], insert step litsimR litsimE vfR vfE, auto)
qed
qed

lemma oKB_permute_step_oKB_step:
  assumes "(E0, R0) ⊢oKBπ (E1, R1)" and "E0 ≐ E0'" and "R0 ≐ R0'" and
    "variant_free_trs E0" and "variant_free_trs E0'" and
    "variant_free_trs R0" and "variant_free_trs R0'"
  shows "∃E1' R1'. (E0', R0') ⊢oKB (E1', R1') ∧ E1 ≐ E1' ∧ R1 ≐ R1'"
proof-
  note sym = subsumable_trs.litsim_sym
  note refl = subsumable_trs.litsim_refl
  have insert:"⋀S S' π s t. S' ≐ S ⟹ S ∪ {(π ∙ s, π ∙ t)} ≐ S' ∪ {(s, t)}"
    using litsim_insert[OF _ rule_pt.permute_prod_eqvt[symmetric]] sym by fastforce
  note insert_ER = insert[OF sym[OF assms(2)]] insert[OF sym[OF assms(3)]]
  note litsim = assms(2) assms(3) litsim_symcl litsim_union
  note diff_E = litsim_diff1[OF assms(2) assms(4) assms(5)]
  note diff_R = litsim_diff1[OF assms(3) assms(6) assms(7)]
  show ?thesis using assms(1)
  proof (cases)
    case (deduce s t u π)
    have "rstep (R0 ∪ E0) = rstep (R0' ∪ E0')" by (rule litsim_rstep_eq, insert litsim, fast)
    note step = oKBi.deduce[OF deduce(3)[unfolded this] deduce(4)[unfolded this]] insert_ER assms(3)
    show ?thesis by (rule exI[of _ "E0' ∪ {(t, u)}"], rule exI[of _ R0'], insert step, unfold deduce, auto)
  next
    case (orientl s t π)
    from litsim_mem[OF assms(2) this(4)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ t) ∈ E0'" by (auto simp:eqvt)
    from orientl(3) have "ψ ∙ s ≻ ψ ∙ t" by (metis subst term_apply_subst_Var_Rep_perm)
    note step = oKBi.orientl[OF this mem, of R0']
    let ?E = "E0' - {(ψ ∙ s, ψ ∙ t)}" and ?R = "R0' ∪ {(ψ ∙ s, ψ ∙ t)}"
    from diff_E[OF _ orientl(4) mem] have E:"E1 ≐ ?E" unfolding orientl(1)  by (auto simp:eqvt)
    from insert_ER(2)[of "π - ψ" "ψ ∙ s" "ψ ∙ t"] have R:"R1 ≐ ?R" unfolding orientl(2) by auto
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ ?R], insert step E R, auto)
  next
    case (orientr t s π)
    from litsim_mem[OF assms(2) this(4)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ t) ∈ E0'" by (auto simp:eqvt)
    from orientr(3) have "ψ ∙ t ≻ ψ ∙ s" by (metis subst term_apply_subst_Var_Rep_perm)
    note step = oKBi.orientr[OF this mem, of R0']
    let ?E = "E0' - {(ψ ∙ s, ψ ∙ t)}" and ?R = "R0' ∪ {(ψ ∙ t, ψ ∙ s)}"
    from diff_E[OF _ orientr(4) mem] have E:"E1 ≐ ?E" unfolding orientr(1)  by (auto simp:eqvt)
    from insert_ER(2)[of "π - ψ" "ψ ∙ t" "ψ ∙ s"] have R:"R1 ≐ ?R" unfolding orientr(2) by auto
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ ?R], insert step E R, auto)
  next
    case (delete s)
    from litsim_mem[OF assms(2) this(3)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ s) ∈ E0'" by (auto simp:eqvt)
    let ?E = "E0' - {(ψ ∙ s, ψ ∙ s)}"
    from diff_E[OF _ _ mem] have E:"E1 ≐ ?E" using delete by (auto simp:eqvt)
    note step = oKBi.delete[OF mem, of R0']
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ R0'], insert E step delete(2) assms(3), auto)
  next
    case (compose t u s π)
    from litsim_mem[OF assms(3) this(4)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ t) ∈ R0'" by (auto simp:eqvt)
    from diff_R[OF _ compose(4) mem] have sim:"R0 - {(s, t)} ≐ R0' - {(ψ ∙ s, ψ ∙ t)}" by (auto simp:eqvt)
    note ostep = ostep_permute_litsim[OF compose(3) assms(2) this]
    hence "(ψ ∙ t, ψ ∙ u) ∈ ostep E0' (R0' - {(ψ ∙ s, ψ ∙ t)})" by (auto simp:eqvt)
    note step = oKBi.compose[OF this mem]
    let ?R = "R0' - {(ψ ∙ s, ψ ∙ t)} ∪ {(ψ ∙ s, ψ ∙ u)}"
    from insert[OF sim[THEN sym], of "π - ψ" "ψ ∙ s" "ψ ∙ u"] have R:"R1 ≐ ?R" unfolding compose(2) by auto
    show ?thesis by (rule exI[of _ E0'], rule exI[of _ ?R], insert step R compose(1) assms(2), auto)
  next
    case (simplifyl s u t π)
    from litsim_mem[OF assms(2) this(4)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ t) ∈ E0'" by (auto simp:eqvt)
    from diff_E[OF _ simplifyl(4) mem] have sim:"E0 - {(s, t)} ≐ E0' - {(ψ ∙ s, ψ ∙ t)}" by (auto simp:eqvt)
    note ostep = encstep1_permute_litsim[OF] simplifyl(3) this assms(3)
    hence "(ψ ∙ s, ψ ∙ u) ∈ encstep1 (E0' - {(ψ ∙ s, ψ ∙ t)}) R0'" by (auto simp:eqvt)
    note step = oKBi.simplifyl[OF this mem]
    let ?E = "E0' - {(ψ ∙ s, ψ ∙ t)} ∪ {(ψ ∙ u, ψ ∙ t)}"
    from insert[OF sim[THEN sym], of "π - ψ" "ψ ∙ u" "ψ ∙ t"] have E:"E1 ≐ ?E" unfolding simplifyl(1) by auto
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ R0'], insert step E simplifyl(2) assms(3), auto)
  next
    case (simplifyr t u s π)
    from litsim_mem[OF assms(2) this(4)] obtain ψ where mem:"(ψ ∙ s, ψ ∙ t) ∈ E0'" by (auto simp:eqvt)
    from diff_E[OF _ simplifyr(4) mem] have sim:"E0 - {(s, t)} ≐ E0' - {(ψ ∙ s, ψ ∙ t)}" by (auto simp:eqvt)
    note ostep = encstep1_permute_litsim[OF simplifyr(3) this assms(3)]
    hence "(ψ ∙ t, ψ ∙ u) ∈ encstep1 (E0' - {(ψ ∙ s, ψ ∙ t)}) R0'" by (auto simp:eqvt)
    note step = oKBi.simplifyr[OF this mem]
    let ?E = "E0' - {(ψ ∙ s, ψ ∙ t)} ∪ {(ψ ∙ s, ψ ∙ u)}"
    from insert[OF sim[THEN sym], of "π - ψ" "ψ ∙ s" "ψ ∙ u"] have E:"E1 ≐ ?E" unfolding simplifyr(1) by auto
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ R0'], insert step E simplifyr(2) assms(3), auto)
  next
    case (collapse t u s π)
    from litsim_mem[OF assms(3) this(4)] obtain ψ where mem:"(ψ ∙ t, ψ ∙ s) ∈ R0'" by (auto simp:eqvt)
    let ?E = "E0' ∪ {(ψ ∙ u, ψ ∙ s)}"
    let ?R = "R0' - {(ψ ∙ t, ψ ∙ s)}"
    from diff_R[OF _ collapse(4) mem] have R:"R0 - {(t, s)} ≐ ?R" by (auto simp:eqvt)
    note ostep = encstep2_permute_litsim[OF collapse(3) assms(2) this]
    hence "(ψ ∙ t, ψ ∙ u) ∈ encstep2 E0' ?R" by (auto simp:eqvt)
    note step = oKBi.collapse[OF this mem]
    from insert_ER(1)[of "π - ψ" "ψ ∙ u" "ψ ∙ s"] have E:"E1 ≐ ?E" unfolding collapse(1) by auto
    show ?thesis by (rule exI[of _ ?E], rule exI[of _ ?R], insert step E R collapse(2), auto)
  qed
qed

lemma oKB_run_rename:
  assumes "∀i < n. (E i, R i) ⊢oKB (E (Suc i), R (Suc i))" and "∀i ≤ n. variant_free_trs (E i) ∧ variant_free_trs (R i)"
    and "Rf ≐ R n" and "Ef ≐ E n" and "variant_free_trs Ef" and "variant_free_trs Rf" and "R 0 ⊆ {≻}"
  shows "∃E' R'.((∀i < n. (E' i, R' i) ⊢oKB (E' (Suc i), R' (Suc i))) ∧
               (∀i ≤ n. variant_free_trs (E' i) ∧ variant_free_trs (R' i) ∧ E i ≐ E' i ∧ R i ≐ R' i)) ∧
               E' n = Ef ∧ R' n = Rf"
  using assms
proof(induct n arbitrary: Ef Rf)
  case 0
  show ?case by (rule exI[of _ "λi. Ef"], rule exI[of _ "λi. Rf"], insert subsumable_trs.litsim_sym 0, auto)
next
  case (Suc n)
  note eqrel = subsumable_trs.litsim_refl subsumable_trs.litsim_trans subsumable_trs.litsim_sym
  from Suc(2) have step:"(E n, R n) ⊢oKB (E (Suc n), R (Suc n))" by auto
  from Suc(3) have vf:"variant_free_trs (E (Suc n))" "variant_free_trs (R (Suc n))" by auto
  from Suc(3) have vf':"variant_free_trs (E n)" "variant_free_trs (R n)" by auto
  from Suc(2) have "(oKBi ^^ n) (E 0, R 0) (E n, R n)" unfolding relpowp_fun_conv by auto
  from relpowp_imp_rtranclp[OF this] Suc(8) oKBi_rtrancl_less have "R n ⊆ {≻}" by auto
  with step_implies_litsim_step[OF step Suc(4) Suc(5) Suc(6) Suc(7) vf vf'] eqrel obtain En Rn where
    ERn:"(En, Rn) ⊢oKB (Ef, Rf)" "Rn ≐ R n" "En ≐ E n" "variant_free_trs En" "variant_free_trs Rn" by blast
  from Suc have "∀i < n. (E i, R i) ⊢oKB (E (Suc i), R (Suc i))" and "∀i ≤ n. variant_free_trs (E i) ∧ variant_free_trs (R i)" by auto
  from Suc(1)[OF this ERn(2) ERn(3) ERn(4) ERn(5) Suc(8)] obtain E' R' where prefix:
    "En = E' n" "Rn = R' n" "∀i < n. (E' i, R' i) ⊢oKB (E' (Suc i), R' (Suc i))"
    "∀i ≤ n. variant_free_trs (E' i) ∧ variant_free_trs (R' i) ∧ E i ≐ E' i ∧ R i ≐ R' i" by auto
  let ?E = "λi. if i ≤ n then E' i else Ef"
  let ?R = "λi. if i ≤ n then R' i else Rf"
  { fix i
    assume "i ≤ Suc n"
    from this[unfolded le_eq_less_or_eq] prefix(4)[rule_format, of i] Suc subsumable_trs.litsim_sym
    have "variant_free_trs (?E i) ∧ variant_free_trs (?R i) ∧ E i ≐ ?E i ∧ R i ≐ ?R i"
    by (cases "i ≤ n", auto)
  } note vf = this
  { fix i
    assume "i < Suc n"
    from this[unfolded le_eq_less_or_eq] prefix(3) ERn(1) have "(?E i, ?R i) ⊢oKB (?E (Suc i), ?R (Suc i))"
      unfolding prefix by (cases "i = n", auto)
  } note steps = this
  show ?case by (rule exI[of _ ?E], rule exI[of _ ?R], insert steps vf, auto)
qed

lemma R_less_litsim_R_less:
  assumes "R ⊆ {≻}" and "R' ≐ R"
  shows "R' ⊆ {≻}"
proof
  fix l r
  assume "(l,r) ∈ R'"
  with assms(2)[unfolded subsumable_trs.litsim_def subsumeseq_trs_def]
    obtain π where "(π ∙ l,π ∙ r) ∈ R" by (auto simp:eqvt)
  hence "π ∙ l ≻ π ∙ r" using assms(1) by blast
  with less_set_permute show "(l, r) ∈ {≻}" using eqvt by blast
qed

lemma oKB_permute_run:
  assumes "∀i < n. (E i, R i) ⊢oKBπ (E (Suc i), R (Suc i))" and "∀i ≤ n. variant_free_trs (E i) ∧ variant_free_trs (R i)" and "R 0 ⊆ {≻}"
  shows "∃E' R'.((∀i < n. (E' i, R' i) ⊢oKB (E' (Suc i), R' (Suc i))) ∧ (∀i ≤ n. variant_free_trs (E' i) ∧ variant_free_trs (R' i) ∧ E i ≐ E' i ∧ R i ≐ R' i))"
  using assms(1) assms(2)
proof(induct n)
  case 0
  show ?case by (rule exI[of _ "E"], rule exI[of _ "R"], insert subsumable_trs.litsim_refl 0, auto)
next
  case (Suc n)
  note eqrel = subsumable_trs.litsim_refl subsumable_trs.litsim_trans subsumable_trs.litsim_sym
  from Suc have all:"∀i < n. (E i, R i) ⊢oKBπ (E (Suc i), R (Suc i))" "∀i ≤ n. variant_free_trs (E i) ∧ variant_free_trs (R i)" by auto
  from Suc(1)[OF this] obtain E' R' where steps:"∀i < n. (E' i, R' i) ⊢oKB (E' (Suc i), R' (Suc i))"
    "∀i ≤ n. variant_free_trs (E' i) ∧ variant_free_trs (R' i) ∧ E i ≐ E' i ∧ R i ≐ R' i" by blast
  from R_less_litsim_R_less[OF assms(3)] steps(2) eqrel have less:"R' 0 ⊆ {≻}" by blast
  from steps(2) have litsim:"E n ≐ E' n" "R n ≐ R' n" by auto
  from Suc(3) have vf:"variant_free_trs (E n)" "variant_free_trs (R n)" by auto
  from steps(2) have vf':"variant_free_trs (E' n)" "variant_free_trs (R' n)" by auto
  from Suc(2) have "(E n, R n) ⊢oKBπ (E (Suc n), R (Suc n))" by auto
  from oKB_permute_step_oKB_step[OF this litsim vf(1) vf'(1) vf(2) vf'(2)] obtain Ef Rf
    where t:"(E' n, R' n) ⊢oKB (Ef, Rf)" "R (Suc n) ≐ Rf" "E (Suc n) ≐ Ef" by auto
  from oKB_step_variant_free[OF this(1) vf'] obtain En'' Rn'' Ef'' Rf'' where
    s:"(En'', Rn'') ⊢oKB (Ef'', Rf'')" "En'' ≐ E' n" "Rn'' ≐ R' n" "Ef'' ≐ Ef" "Rf'' ≐ Rf"
    "variant_free_trs En''" "variant_free_trs Rn''" "variant_free_trs Ef''" "variant_free_trs Rf''" by auto
  from oKB_run_rename[OF steps(1) _ s(3) s(2) s(6) s(7) less] steps(2) obtain E'' R'' where prefix:
    "∀i < n. (E'' i, R'' i) ⊢oKB (E'' (Suc i), R'' (Suc i))"
    "∀i ≤ n. variant_free_trs (E'' i) ∧ variant_free_trs (R'' i) ∧ E' i ≐ E'' i ∧ R' i ≐ R'' i" "E'' n = En'' ∧ R'' n = Rn''"
    by auto
  let ?E = "λi. if i ≤ n then E'' i else Ef''"
  let ?R = "λi. if i ≤ n then R'' i else Rf''"
  { fix i
    assume i:"i ≤ Suc n"
    have "variant_free_trs (?E i) ∧ variant_free_trs (?R i) ∧ E i ≐ ?E i ∧ R i ≐ ?R i" proof(cases "i = Suc n")
      case True
      from s(4) s(5) t(2) t(3) eqrel have "E (Suc n) ≐ Ef'' ∧ R (Suc n) ≐ Rf''" by fast
      with s show ?thesis unfolding True by simp
    next
      case False
      with i have i:"i ≤ n" by auto
      with prefix(2)[rule_format, of i] steps(2)[rule_format, of i] eqrel have "E i ≐ E'' i ∧ R i ≐ R'' i" by fast
      with i prefix(2)[rule_format, of i] show ?thesis by simp
    qed
  } note vf = this
  { fix i
    assume "i < Suc n"
    from this[unfolded le_eq_less_or_eq] prefix(3) s(1) prefix(1)[rule_format, of i]
      have "(?E i, ?R i) ⊢oKB (?E (Suc i), ?R (Suc i))" by (cases "i = n", auto)
  } note steps = this
  show ?case by (rule exI[of _ ?E], rule exI[of _ ?R], insert steps vf, auto)
qed
end

definition "E_ord ord E = {(s ⋅ σ, t ⋅ σ) | s t σ. (s, t) ∈ E ∧ ord (s ⋅ σ) (t ⋅ σ)}"

lemma E_ord_mono: "E ⊆ E' ⟹ E_ord ord E ⊆ E_ord ord E'" unfolding E_ord_def by auto

context ordered_completion
begin

context
  fixes   :: "('a, 'b) trs"
  assumes rules: "ℛ ⊆ {≻}"
begin

abbreviation "𝒮 ≡ ℛ ∪ E_ord (op ≻) ℰ"

lemma rstep_S_ordstep:
  assumes "(s, t) ∈ rstep 𝒮"
  shows "(s, t) ∈ ordstep {≻} 𝒮"
  using assms
proof (cases rule: rstepE)
  case (rstep C σ l r)
  then show ?thesis
    using rules
    by (intro ordstep.intros [where C = C and l = l and r = r and σ = σ])
      (auto dest: subst simp: E_ord_def)
qed

lemma rstep_S_rstep_EW: "(s, t) ∈ rstep 𝒮 ⟹ (s, t) ∈ rstep (ℛ ∪ ℰ)"
  unfolding E_ord_def by fast

lemma ordstep_subset_S:
  "ordstep {≻} (ℛ ∪ ℰ) ⊆ ordstep {≻} 𝒮"
proof -
  { fix s t assume "(s, t) ∈ ordstep {≻} (ℛ ∪ ℰ)"
    then obtain C l r σ where "(l, r) ∈ ℛ ∪ ℰ"
      and [simp]: "s = C⟨l ⋅ σ⟩" "t = C⟨r ⋅ σ⟩" "l ⋅ σ ≻ r ⋅ σ" by (cases) auto
    then consider "(l, r) ∈ ℛ" | "(l, r) ∈ ℰ" by blast
    then have "(s, t) ∈ ordstep {≻} 𝒮"
    proof (cases)
      case 1
      then show ?thesis by (auto intro: ordstep.intros)
    next
      case 2
      then show ?thesis
        by (intro ordstep.intros [where C = C and σ = Var and l = "l ⋅ σ" and r = "r ⋅ σ"])
          (force simp: E_ord_def)+
    qed }
  then show ?thesis by auto
qed

lemma NF_ordstep_S_subset: "NF (ordstep {≻} 𝒮) ⊆ NF (ordstep {≻} (ℛ ∪ ℰ))"
  using ordstep_subset_S by blast

lemma ordstep_S_conv [simp]:
  "ordstep {≻} 𝒮 = rstep 𝒮"
  using rstep_S_ordstep by (auto elim: ordstep.cases)

lemma S_less:"rstep 𝒮 ⊆ {≻}"
  using ordstep_S_conv ordstep_imp_ord[OF ctxt_closed_less] by blast

end
end

lemma (in gtotal_reduction_order) GROUND_rstep_ordstep:
  "GROUND(rstep (RR ∪ EE)) ⊆ (GROUND (rstep (RR ∪ E_ord (op ≻) EE)))*"
proof
  fix s t
  assume step: "(s, t) ∈ GROUND(rstep (RR ∪ EE))"
  then obtain C l r σ where C: "s = C⟨l ⋅ σ⟩" "t = C⟨r ⋅ σ⟩" "(l, r) ∈ RR ∪ EE" "ground s" "ground t"
    unfolding GROUND_def by auto
  show "(s, t) ∈ (GROUND (rstep (RR ∪ E_ord (op ≻) EE)))*"
  proof (cases "(l, r) ∈ RR")
    case True
    with step have "(s, t) ∈ (rstep (RR ∪ E_ord (op ≻) EE))" unfolding C rstep_union by blast
    with C(4) C(5) show ?thesis unfolding GROUND_def by auto
  next
    case False
    with step rstep_union C(3) have in_E: "(l, r) ∈ EE" unfolding C by auto
    from C(4) C(5) have g': "ground (l ⋅ σ)" "ground (r ⋅ σ)" unfolding C by auto
    from gtotal [OF g'] consider " l ⋅ σ = r ⋅ σ" | "l ⋅ σ ≻ r ⋅ σ" | "r ⋅ σ ≻ l ⋅ σ" by auto
    then show ?thesis
    proof (cases)
      case 1
      then show ?thesis unfolding C by auto
    next
      case 2
      with in_E have "(l ⋅ σ, r ⋅ σ) ∈ E_ord (op ≻) EE" unfolding E_ord_def mem_Collect_eq by auto
      then have "(s, t) ∈ rstep (RR ∪ E_ord op ≻ EE)" unfolding C by auto
      with C(4) C(5) show ?thesis unfolding GROUND_def by auto
    next
      case 3
      with in_E have "(r ⋅ σ, l ⋅ σ) ∈ E_ord (op ≻) EE" unfolding E_ord_def mem_Collect_eq by auto
      then have "(t, s) ∈ rstep (RR ∪ E_ord op ≻ EE)" unfolding C by auto
      with C(4) C(5) show ?thesis unfolding GROUND_def by auto
    qed
  qed
qed

locale gtotal_okb_irun_inf =
  ordered_completion_inf less +
  gtotal_okb_irun less for less :: "('a, 'b::infinite) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50)
begin

sublocale gtotal_reduction_order_inf ..

definition "PCP_ext RR = {(u, v) | r r' p μ u v.
    ooverlap {≻} RR r r' p μ u v ∧
    (∀u ⊲ fst r ⋅ μ. u ∈ NF (ordstep {≻} RR))}"

context
  fixes   :: "('a, 'b) trs"
  assumes rules: "ℛ ⊆ {≻}"
begin

lemma perm_R_less: "π ∙ (l, r) ∈ ℛ ⟹ l ≻ r"
  using rules by (auto simp: eqvt perm_less)


lemma perm_S_ordstep:
  assumes "π ∙ (s, t) ∈ 𝒮 ℛ ℰ"
  shows "(s ⋅ σ, t ⋅ σ) ∈ ordstep {≻} (𝒮 ℛ ℰ)"
  using assms
  by (intro ordstep.intros [where C =  and σ = "sop (-π) ∘s σ" and l = "π ∙ s" and r = "π ∙ t"])
    (auto simp: eqvt perm_R_less E_ord_def subst, (metis perm_less subst)+)

lemma perm_R_ordstep:
  assumes "π ∙ (s, t) ∈ ℛ"
  shows "(s ⋅ σ, t ⋅ σ) ∈ ordstep {≻} (ℛ ∪ ℰ)"
  using assms
  by (intro ordstep.intros [where C =  and σ = "sop (-π) ∘s σ" and l = "π ∙ s" and r = "π ∙ t"])
    (auto simp: eqvt perm_R_less subst)

lemma E_ordstep:
  assumes "(s, t) ∈ ℰ" and "s ⋅ σ ≻ t ⋅ σ"
  shows "(s ⋅ σ ∘s τ, t ⋅ σ ∘s τ) ∈ ordstep {≻} (ℛ ∪ ℰ)"
  using assms
  by (intro ordstep.intros [where C =  and σ = "σ ∘s τ" and l = "s" and r = "t"])
    (auto simp: subst)

lemma PCP_xPCP:
  assumes "ground s" and "ground t"
    and "(s, t) ∈ (rstep (PCP (𝒮 ℛ ℰ)))"
  shows "(s, t) ∈ (rstep (PCP_ext (ℛ ∪ ℰ))) ∪ (rstep (𝒮 ℛ ℰ))"
    (is "_ ∈ ?R")
proof -
  have S: "⋀s t. (s, t) ∈ (𝒮 ℛ ℰ) ⟹ s ≻ t ∨ (t, s) ∈ (𝒮 ℛ ℰ)"
    using rules by (auto simp: E_ord_def)
  have RE: "⋀s t. (s, t) ∈ ℛ ∪ ℰ ⟹ s ≻ t ∨ (t, s) ∈ ℛ ∪ ℰ"
    using rules by (auto)
  let ?S = "ordstep {≻} (𝒮 ℛ ℰ)"
  let ?Eo = "E_ord (op ≻) ℰ"
  let ?RE = "ordstep {≻} (ℛ ∪ ℰ)"
  { fix s t and σ :: "('a, 'b) subst"
    assume "(s, t) ∈ PCP (𝒮 ℛ ℰ)" and ground: "ground (s ⋅ σ)" "ground (t ⋅ σ)"
    then obtain μ l1 r1 p l2 r2 where o: "overlap (𝒮 ℛ ℰ) (𝒮 ℛ ℰ) (l1, r1) p (l2, r2)"
      and "the_mgu l1 (l2 |_ p) = μ"
      and s: "s = replace_at l2 p r1 ⋅ μ"
      and t: "t = r2 ⋅ μ"
      and NF: "∀u ⊲ l1 ⋅ μ. u ∈ NF (rstep (𝒮 ℛ ℰ))"
      by (auto simp: PCP_def the_mgu_def split: option.splits)
    then have mgu: "mgu l1 (l2 |_ p) = Some μ"
      by (auto simp: the_mgu_def unifiers_def overlap_def simp del: mgu.simps split: option.splits dest!: mgu_complete)
    from mgu_sound [OF this]
    have "l1 ⋅ μ = (l2 |_ p) ⋅ μ" by (auto simp: is_imgu_def)
    moreover
    define C and τ where "C = ctxt_of_pos_term p (l2 ⋅ μ) ⋅c σ" and "τ = μ ∘s σ"
    ultimately
    have eq: "□⟨l2 ⋅ τ⟩ = C⟨l1 ⋅ τ⟩"
      using o
      by simp (metis ctxt_of_pos_term_subst ctxt_supt_id fst_conv overlap_source_eq(1) subst_apply_term_ctxt_apply_distrib)

    have s': "s ⋅ σ = C⟨r1 ⋅ τ⟩"
      using o by (simp add: s C_def τ_def overlap_def ctxt_of_pos_term_subst funposs_imp_poss)

    have t': "t ⋅ σ = □⟨r2 ⋅ τ⟩" by (simp add: t τ_def)

    have p: "p ∈ funposs l2" using o by (auto simp: overlap_def)
    then have [simp]: "hole_pos C = p" by (auto simp: C_def funposs_imp_poss)

    have disj: "vars_rule (l2, r2) ∩ vars_rule (l1, r1) = {}" using o by (auto simp: overlap_def)
    note Sconv[simp] = ordstep_S_conv[OF rules]

    { fix π1 and π2 assume *: 1 ∙ (l1, r1) ∈ ℛ" 2 ∙ (l2, r2) ∈ ℛ"
      then have rule1: 1 ∙ (l1, r1) ∈ 𝒮 ℛ ℰ" and rule2: 2 ∙ (l2, r2) ∈ 𝒮 ℛ ℰ" by auto
      have "(s ⋅ σ , t ⋅ σ) ∈ ?R"
      proof (cases "hole_pos C ∈ funposs l2")
        case False
        from non_ooverlap_GROUND_joinable [OF S eq refl False _ _ _ s' t' ground rule2]
          and perm_S_ordstep [OF rule1, of τ] and perm_S_ordstep [OF rule2, of τ]
        show ?thesis by (auto, metis GROUND_subset Un_absorb1 Un_iff join_mono)
      next
        case True
        then have "ooverlap {≻} (ℛ ∪ ℰ) (l1, r1) (l2, r2) p μ s t"
          using * and disj and p [THEN funposs_imp_poss] and irrefl
          by (auto simp: ooverlap_def mgu s t ctxt_of_pos_term_subst simp del: mgu.simps)
            (auto dest!: perm_R_less dest: subst trans)
        moreover have "∀u ⊲ l1 ⋅ μ. u ∈ NF (ordstep {≻} (ℛ ∪ ℰ))"
          using NF and NF_ordstep_S_subset[OF rules] unfolding Sconv by force
        ultimately have "(s, t) ∈ PCP_ext (ℛ ∪ ℰ)"
          by  (force simp: PCP_ext_def)
        then show ?thesis by blast
      qed }
    moreover
    { fix π1 and π2 assume rule1: 1 ∙ (l1, r1) ∈ ℛ" and 2 ∙ (l2, r2) ∈ ?Eo"
      then obtain u2 v2 τ2' where rule2: "(u2, v2) ∈ ℰ"
        and 2 ∙ l2 = u2 ⋅ τ2'" and 2 ∙ r2 = v2 ⋅ τ2'"
        and 2 ∙ l2 ≻ π2 ∙ r2" and "u2 ⋅ τ2' ≻ v2 ⋅ τ2'"
        by (auto simp: E_ord_def eqvt)
      moreover define τ2 where 2 = τ2' ∘s sop (-π2)"
      ultimately have [simp]: "l2 = u2 ⋅ τ2" "r2 = v2 ⋅ τ2"
        and gr2: "u2 ⋅ τ2 ≻ v2 ⋅ τ2" by (auto simp: term_pt.permute_flip perm_less)
      from eq have eq': "□⟨u2 ⋅ τ2s τ⟩ = C⟨l1 ⋅ τ⟩" by simp

      have rule2': "0 ∙ (u2, v2) ∈ ℛ ∪ ℰ" using ‹(u2, v2) ∈ ℰ by (auto)
      have "(s ⋅ σ, t ⋅ σ) ∈ ?R"
      proof (cases "hole_pos C ∈ funposs u2")
        case False

        from t' have t'': "t ⋅ σ = □⟨v2 ⋅ τ2s τ⟩" by simp
        from non_ooverlap_GROUND_joinable [OF RE eq' refl False _ _ _ s' t'' ground rule2']
          and perm_R_ordstep [OF rule1, of τ] and E_ordstep [OF rule2 gr2, of τ]
        have "(s ⋅ σ, t ⋅ σ) ∈ (GROUND (ordstep {≻} (ℛ ∪ ℰ)))" by auto
        then show ?thesis
          using join_mono [OF ordstep_subset_S, OF rules] and GROUND_subset unfolding Sconv
          by (metis (no_types, lifting) UnI2 join_mono set_mp)
      next
        case True
        obtain π where disj: "vars_rule (l1, r1) ∩ vars_rule (π ∙ u2, π ∙ v2) = {}"
          using vars_rule_disjoint unfolding eqvt [symmetric] by blast
        have rule2'': "-π ∙ (π ∙ u2, π ∙ v2) ∈ ℛ ∪ ℰ"
          using rule2' by (simp)
        then have rule2'': "∃p. p ∙ (π ∙ u2, π ∙ v2) ∈ ℛ ∪ ℰ" ..
        define ν where "ν x = (if x ∈ vars_rule (l1, r1) then Var x else (sop (-π) ∘s τ2) x)" for x
        have 1: "π ∙ u2 ⋅ ν = u2 ⋅ τ2"
        proof -
          have "π ∙ u2 ⋅ ν = π ∙ u2 ⋅ sop (-π) ∘s τ2"
            using disj by (unfold term_subst_eq_conv) (auto simp add: ν_def vars_defs)
          then show ?thesis by simp
        qed
        have 2: "π ∙ v2 ⋅ ν = v2 ⋅ τ2"
        proof -
          have "π ∙ v2 ⋅ ν = π ∙ v2 ⋅ sop (-π) ∘s τ2"
            using disj by (unfold term_subst_eq_conv) (auto simp add: ν_def vars_defs)
          then show ?thesis by simp
        qed
        have [simp]: "r1 ⋅ ν = r1"
        proof -
          have "r1 ⋅ ν = r1 ⋅ Var" by (unfold term_subst_eq_conv) (simp add: ν_def vars_defs)
          then show ?thesis by simp
        qed
        have r2: "r2 = π ∙ v2 ⋅ ν"
        proof -
          have "π ∙ v2 ⋅ ν = π ∙ v2 ⋅ sop (-π) ∘s τ2"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis using p and True
            by (auto simp: eqvt dest!: funposs_imp_poss)
        qed
        have l2: "l2 = π ∙ u2 ⋅ ν"
        proof -
          have "π ∙ u2 ⋅ ν = π ∙ u2 ⋅ sop (-π) ∘s τ2"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis using p and True
            by (auto simp: eqvt dest!: funposs_imp_poss)
        qed
        then have "l2 |_ p = (π ∙ u2 |_ p) ⋅ ν"
          using True by (auto dest!: funposs_imp_poss)
        moreover
        have "l1 ⋅ ν ∘s μ = l1 ⋅ μ"
          by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def)
        ultimately
        have *: "l1 ⋅ ν ∘s μ = (π ∙ u2 |_ p) ⋅ ν ∘s μ" using ‹l1 ⋅ μ = (l2 |_ p) ⋅ μ› by simp
        then obtain μ' where mgu: "mgu l1 (π ∙ u2 |_ p) = Some μ'"
          using mgu_complete by (auto simp: unifiers_def simp del: subst_subst_compose)
        from mgu_sound [OF this, THEN is_imgu_imp_is_mgu] and * obtain δ where μ': " ν ∘s μ = μ' ∘s δ"
          by (auto simp: is_mgu_def unifiers_def simp del: subst_subst_compose)
        have l1: "l1 ⋅ μ = l1 ⋅ μ' ∘s δ"
          by (unfold term_subst_eq_conv, fold μ') (simp add: subst_compose ν_def vars_defs)
        let ?s = "(ctxt_of_pos_term p (π ∙ u2 ⋅ μ'))⟨r1 ⋅ μ'⟩"
        let ?t = "π ∙ v2 ⋅ μ'"
        have 3: "t ⋅ μ' ⋅ δ = t ⋅ ν ⋅ μ" for t by (fold subst_subst_compose) (simp add: μ')
        have "p ∈ poss (π ∙ u2)" using True by (simp add: eqvt funposs_imp_poss)
        then have 4: "s ⋅ σ = ?s ⋅ δ ⋅ σ"
          by (simp add: s ctxt_of_pos_term_subst [symmetric] 1 3)
        have 5: "t ⋅ σ = ?t ⋅ δ ⋅ σ" by (simp add: t 2 3)

        have "¬ π ∙ v2 ⋅ μ' ≻ π ∙ u2 ⋅ μ'" (is "¬ ?v ⋅ _ ≻ ?u ⋅ _")
        proof
          presume "¬ ?thesis"
          then have "?v ⋅ ν ∘s μ ≻ ?u ⋅ ν ∘s μ"
            unfolding μ' by (simp add: subst)
          then have "r2 ⋅ μ ≻ l2 ⋅ μ" unfolding l2 r2 by simp
          moreover have "l2 ≻ r2" using ‹π2 ∙ l2 ≻ π2 ∙ r2 by (simp add: perm_less)
          ultimately show False using irrefl by (auto dest: subst trans)
        qed simp
        then have "ooverlap {≻} (ℛ ∪ ℰ) (l1, r1) (π ∙ u2, π ∙ v2) p μ' ?s ?t"
          using mgu and disj and rule1 and rule2'' and True and irrefl
          by (auto simp: ooverlap_def perm_less simp del: mgu.simps)
            (auto dest!: perm_R_less dest: subst trans simp del: mgu.simps)
        moreover have "∀u ⊲ l1 ⋅ μ'. u ∈ NF (ordstep {≻} (ℛ ∪ ℰ))"
          using NF and NF_ordstep_S_subset[OF rules]
          unfolding l1
          by auto (meson NF_instance contra_subsetD instance_no_supt_imp_no_supt)
        ultimately have "(?s, ?t) ∈ PCP_ext (ℛ ∪ ℰ)" by (force simp: PCP_ext_def)
        then show ?thesis
          unfolding 4 and 5 by blast
      qed }
    moreover
    { fix π1 and π2 assume *: 1 ∙ (l1, r1) ∈ ?Eo" and rule2: 2 ∙ (l2, r2) ∈ ℛ"
      then have rule1: 1 ∙ (l1, r1) ∈ 𝒮 ℛ ℰ" and rule2': 2 ∙ (l2, r2) ∈ 𝒮 ℛ ℰ" by auto
      have "(s ⋅ σ, t ⋅ σ) ∈ ?R"
      proof (cases "hole_pos C ∈ funposs l2")
        case False
        from non_ooverlap_GROUND_joinable [OF S eq refl False _ _ _ s' t' ground rule2']
          and perm_S_ordstep [OF rule1, of τ] and perm_S_ordstep [OF rule2', of τ]
        show ?thesis by (auto) (metis GROUND_subset Un_absorb1 Un_iff join_mono)
      next
        from * obtain u1 v1 τ1' where rule1: "(u1, v1) ∈ ℰ"
          and 1 ∙ l1 = u1 ⋅ τ1'" and 1 ∙ r1 = v1 ⋅ τ1'"
          and 1 ∙ l1 ≻ π1 ∙ r1" and "u1 ⋅ τ1' ≻ v1 ⋅ τ1'"
          by (auto simp: E_ord_def eqvt)
        moreover define τ1 where 1 = τ1' ∘s sop (-π1)"
        ultimately have [simp]: "l1 = u1 ⋅ τ1" "r1 = v1 ⋅ τ1"
          and gr2: "u1 ⋅ τ1 ≻ v1 ⋅ τ1" by (auto simp: term_pt.permute_flip perm_less)
        case True
        obtain π where disj: "vars_rule (π ∙ u1, π ∙ v1) ∩ vars_rule (l2, r2) = {}"
          using vars_rule_disjoint unfolding eqvt [symmetric] by blast
        have rule1'': "-π ∙ (π ∙ u1, π ∙ v1) ∈ ℛ ∪ ℰ"
          using rule1 by (simp)
        then have rule1'': "∃p. p ∙ (π ∙ u1, π ∙ v1) ∈ ℛ ∪ ℰ" ..
        define ν where "ν x = (if x ∈ vars_rule (l2, r2) then Var x else (sop (-π) ∘s τ1) x)" for x
        have [simp]: "l2 ⋅ ν = l2 ⋅ Var" by (unfold term_subst_eq_conv) (simp add: ν_def vars_defs)
        then have [simp]: "(l2 |_ p) ⋅ ν = l2 |_ p"
          using True by (auto dest!: funposs_imp_poss) (metis ‹l2 ⋅ ν = l2 ⋅ Var› subst.cop_nil subt_at_subst)

        have [simp]: "r2 ⋅ ν = r2 ⋅ Var" by (unfold term_subst_eq_conv) (simp add: ν_def vars_defs)
        have [simp]: "π ∙ u1 ⋅ ν = u1 ⋅ τ1"
        proof -
          have "π ∙ u1 ⋅ ν = π ∙ u1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis by simp
        qed
        have [simp]: "π ∙ v1 ⋅ ν = v1 ⋅ τ1"
        proof -
          have "π ∙ v1 ⋅ ν = π ∙ v1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis by simp
        qed
        have l1: "l1 = π ∙ u1 ⋅ ν"
        proof -
          have "π ∙ u1 ⋅ ν = π ∙ u1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis using p and True
            by (auto simp: eqvt dest!: funposs_imp_poss)
        qed
        have r1: "r1 = π ∙ v1 ⋅ ν"
        proof -
          have "π ∙ v1 ⋅ ν = π ∙ v1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis using p and True
            by (auto simp: eqvt dest!: funposs_imp_poss)
        qed
        have *: "π ∙ u1 ⋅ ν ∘s μ = (l2 |_ p) ⋅ ν ∘s μ" using ‹l1 ⋅ μ = (l2 |_ p) ⋅ μ› by (simp)
        then obtain μ' where mgu: "mgu (π ∙ u1) (l2 |_ p) = Some μ'"
          using mgu_complete by (auto simp: unifiers_def simp del: subst_subst_compose)
        from mgu_sound [OF this, THEN is_imgu_imp_is_mgu] and * obtain δ where μ': " ν ∘s μ = μ' ∘s δ"
          by (auto simp: is_mgu_def unifiers_def simp del: subst_subst_compose)

        have 1: "t ⋅ μ' ⋅ δ = t ⋅ ν ⋅ μ" for t
          by (fold subst_subst_compose) (simp add: μ')
        let ?s = "(ctxt_of_pos_term p (l2 ⋅ μ'))⟨π ∙ v1 ⋅ μ'⟩"
        let ?t = "r2 ⋅ μ'"
        have 2: "s ⋅ σ = ?s ⋅ δ ⋅ σ"
          using True [THEN funposs_imp_poss]
          by (simp add: s ctxt_of_pos_term_subst [symmetric] 1)
        have 3: "t ⋅ σ = ?t ⋅ δ ⋅ σ" by (simp add: t 1)

        have "¬ π ∙ v1 ⋅ μ' ≻ π ∙ u1 ⋅ μ'"
          (is "¬ ?v ⋅ _ ≻ ?u ⋅ _")
        proof
          presume "¬ ?thesis"
          then have "?v ⋅ ν ∘s μ ≻ ?u ⋅ ν ∘s μ"
            unfolding μ' by (simp add: subst)
          then have "r1 ⋅ μ ≻ l1 ⋅ μ" unfolding l1 r1 by simp
          moreover have "l1 ≻ r1" using ‹π1 ∙ l1 ≻ π1 ∙ r1 by (simp add: perm_less)
          ultimately show False using irrefl by (auto dest: subst trans)
        qed simp
        then have "ooverlap {≻} (ℛ ∪ ℰ) (π ∙ u1, π ∙ v1) (l2, r2) p μ' ?s ?t"
          using mgu and disj and rule1'' and rule2 and True and irrefl
          by (auto simp: ooverlap_def perm_less) (auto dest!: perm_R_less dest: subst trans)
        moreover have "∀u ⊲ π ∙ u1 ⋅ μ'. u ∈ NF (ordstep {≻} (ℛ ∪ ℰ))"
          using NF and NF_ordstep_S_subset[OF rules]
          unfolding l1 and 1 [symmetric]
          by auto (meson NF_instance contra_subsetD instance_no_supt_imp_no_supt)
        ultimately have "(?s, ?t) ∈ PCP_ext (ℛ ∪ ℰ)" by (force simp: PCP_ext_def)
        then show ?thesis
          unfolding 2 and 3 by blast
      qed }
    moreover
    { fix π1 and π2 assume 1 ∙ (l1, r1) ∈ ?Eo" and 2 ∙ (l2, r2) ∈ ?Eo"
      then obtain u1 v1 τ1' and u2 v2 τ2'
        where rule1: "(u1, v1) ∈ ℰ" and 1 ∙ l1 = u1 ⋅ τ1'" and 1 ∙ r1 = v1 ⋅ τ1'"
          and 1 ∙ l1 ≻ π1 ∙ r1"
          and rule2: "(u2, v2) ∈ ℰ" and 2 ∙ l2 = u2 ⋅ τ2'" and 2 ∙ r2 = v2 ⋅ τ2'"
          and 2 ∙ l2 ≻ π2 ∙ r2"
        by (force simp: E_ord_def eqvt)
      moreover
      define τ1 and τ2 where 1 = τ1' ∘s sop (-π1)" and 2 = τ2' ∘s sop (-π2)"
      ultimately
      have [simp]: "u1 ⋅ τ1 = l1" "v1 ⋅ τ1 = r1" "u2 ⋅ τ2 = l2" "v2 ⋅ τ2 = r2"
        and gr1: "u1 ⋅ τ1 ≻ v1 ⋅ τ1" and gr2: "u2 ⋅ τ2 ≻ v2 ⋅ τ2"
        by (auto simp: term_pt.permute_flip [symmetric] perm_less)

      have rule2': "0 ∙ (u2, v2) ∈ ℛ ∪ ℰ" using rule2 by simp

      have "(s ⋅ σ, t ⋅ σ) ∈ ?R"
      proof (cases "hole_pos C ∈ funposs u2")
        case False

        from eq have eq': "□⟨u2 ⋅ τ2s τ⟩ = C⟨u1 ⋅ τ1s τ⟩" by simp
        from s' have s'': "s ⋅ σ = C⟨v1 ⋅ τ1s τ⟩" by simp
        from t' have t'': "t ⋅ σ = □⟨v2 ⋅ τ2s τ⟩" by simp
        from non_ooverlap_GROUND_joinable [OF RE eq' refl False _ _ _ s'' t'' ground rule2']
          and E_ordstep [OF rule1 gr1, of τ] and E_ordstep [OF rule2 gr2, of τ]
        have "(s ⋅ σ, t ⋅ σ) ∈ (GROUND (ordstep {≻} (ℛ ∪ ℰ)))" by auto
        then show ?thesis
          using join_mono [OF ordstep_subset_S[OF rules]] and GROUND_subset
          by (metis (no_types, lifting) UnI2 join_mono Sconv set_mp)
      next
        case True
        obtain π where disj: "vars_rule (π ∙ u1, π ∙ v1) ∩ vars_rule (u2, v2) = {}"
          using vars_rule_disjoint unfolding eqvt [symmetric] by blast
        have "-π ∙ (π ∙ u1, π ∙ v1) ∈ ℛ ∪ ℰ" using rule1 by simp
        then have rule1': "∃p. p ∙ (π ∙ u1, π ∙ v1) ∈ ℛ ∪ ℰ" ..
        have rule2'': "∃p. p ∙ (u2, v2) ∈ ℛ ∪ ℰ" using rule2' ..

        define ν where "ν x = (if x ∈ vars_rule (u2, v2) then τ2 x else (sop (-π) ∘s τ1) x)" for x

        have l1: "π ∙ u1 ⋅ ν = l1"
        proof -
          have "π ∙ u1 ⋅ ν = π ∙ u1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis by simp
        qed
        have [simp]: "π ∙ v1 ⋅ ν = r1"
        proof -
          have "π ∙ v1 ⋅ ν = π ∙ v1 ⋅ sop (-π) ∘s τ1"
            using disj
            by (unfold term_subst_eq_conv) (auto simp: subst_compose vars_defs ν_def eqvt)
          then show ?thesis by simp
        qed
        have [simp]: "u2 ⋅ ν = l2"
        proof -
          have "u2 ⋅ ν = u2 ⋅ τ2" by (unfold term_subst_eq_conv) (auto simp: ν_def vars_defs)
          then show ?thesis by simp
        qed
        then have [simp]: "u2 |_ p ⋅ ν = l2 |_ p"
          using True [THEN funposs_imp_poss] by (auto) (metis ‹u2 ⋅ ν = l2 subt_at_subst)

        have [simp]: "v2 ⋅ ν = r2"
        proof -
          have "v2 ⋅ ν = v2 ⋅ τ2" by (unfold term_subst_eq_conv) (auto simp: ν_def vars_defs)
          then show ?thesis by simp
        qed

        have *: "π ∙ u1 ⋅ ν ∘s μ = (u2 |_ p) ⋅ ν ∘s μ" using ‹l1 ⋅ μ = (l2 |_ p) ⋅ μ› by (simp add: l1)
        then obtain μ' where mgu: "mgu (π ∙ u1) (u2 |_ p) = Some μ'"
          using mgu_complete by (auto simp: unifiers_def simp del: subst_subst_compose)
        from mgu_sound [OF this, THEN is_imgu_imp_is_mgu] and * obtain δ where μ': " ν ∘s μ = μ' ∘s δ"
          by (auto simp: is_mgu_def unifiers_def simp del: subst_subst_compose)

        have 1: "t ⋅ μ' ⋅ δ = t ⋅ ν ⋅ μ" for t
          by (fold subst_subst_compose) (simp add: μ')
        let ?s = "(ctxt_of_pos_term p (u2 ⋅ μ'))⟨π ∙ v1 ⋅ μ'⟩"
        let ?t = "v2 ⋅ μ'"
        have 2: "s ⋅ σ = ?s ⋅ δ ⋅ σ"
          using True [THEN funposs_imp_poss] and p [THEN funposs_imp_poss]
          by (simp add: s ctxt_of_pos_term_subst [symmetric] 1)
        have 3: "t ⋅ σ = ?t ⋅ δ ⋅ σ" by (simp add: t 1)

        have "¬ π ∙ v1 ⋅ μ' ≻ π ∙ u1 ⋅ μ'" (is "¬ ?v ⋅ _ ≻ ?u ⋅ _")
        proof
          presume "¬ ?thesis"
          then have "?v ⋅ ν ∘s μ ≻ ?u ⋅ ν ∘s μ" unfolding μ' by (simp add: subst)
          then have "r1 ⋅ μ ≻ l1 ⋅ μ" by (simp add: l1)
          moreover have "l1 ≻ r1" using ‹π1 ∙ l1 ≻ π1 ∙ r1 by (simp add: perm_less)
          ultimately show False using irrefl by (auto dest: subst trans)
        qed simp
        moreover have "¬ v2 ⋅ μ' ≻ u2 ⋅ μ'" (is "¬ ?v ⋅ _ ≻ ?u ⋅ _")
        proof
          presume "¬ ?thesis"
          then have "?v ⋅ ν ∘s μ ≻ ?u ⋅ ν ∘s μ" unfolding μ' by (simp add: subst)
          then have "r2 ⋅ μ ≻ l2 ⋅ μ" by simp
          moreover have "l2 ≻ r2" using ‹π2 ∙ l2 ≻ π2 ∙ r2 by (simp add: perm_less)
          ultimately show False using irrefl by (auto dest: subst trans)
        qed simp
        ultimately
        have "ooverlap {≻} (ℛ ∪ ℰ) (π ∙ u1, π ∙ v1) (u2, v2) p μ' ?s ?t"
          using mgu and disj and rule1' and rule2'' and True
          by (auto simp: ooverlap_def perm_less)
        moreover
        have "∀u ⊲ π ∙ u1 ⋅ μ'. u ∈ NF ?RE"
          using NF and NF_ordstep_S_subset
          unfolding l1 [symmetric] and 1 [symmetric]
          by auto (meson NF_instance rules contra_subsetD instance_no_supt_imp_no_supt)
        ultimately have "(?s, ?t) ∈ PCP_ext (ℛ ∪ ℰ)" by (force simp: PCP_ext_def)
        then show ?thesis
          unfolding 2 and 3 by blast
      qed }
    ultimately have "(s ⋅ σ, t ⋅ σ) ∈ ?R" using o by (auto simp: overlap_def) }
  then show ?thesis
    using assms by (auto elim!: rstepE) blast+
qed

abbreviation "ℰsucc" where "ℰsucc ≡ ℰ ∩ {≻}"
definition ℛ':: "('a, 'b) term rel" where "ℛ' ≡ {(l, r) | l r. (l, r) ∈ dot (ℛ ∪ ℰsucc) ∧ (∄u l0 r0. (l, u) ∈ rstep {(l0, r0)} ∧ (l0, r0) ∈ 𝒮 ℛ ℰ ∧ l0 ⊲⋅ l)}"
definition ℰ' :: "('a, 'b) term rel" where "ℰ' ≡ {(some_NF (rstep ℛ') s, some_NF (rstep ℛ') t) | s t. (s, t) ∈ ℰ ∧ some_NF (rstep ℛ') s ≠ some_NF (rstep ℛ') t }"
"𝒮'"">definition "𝒮'":: "('a, 'b) term rel" where "𝒮' = ℛ' ∪ E_ord (op ≻) ℰ'"


lemma Esucc_E_ord: "ℰsucc ⊆ E_ord (op ≻) ℰ"
proof
  fix s t
  assume "(s, t) ∈ ℰsucc"
  note ord = this [unfolded Int_iff split]
  show "(s, t) ∈ E_ord (op ≻) ℰ" unfolding E_ord_def mem_Collect_eq
    by(rule exI [of _ s], rule exI [of _ t], rule exI [of _ Var], insert ord, auto)
qed

lemma encomp_perm: "u ⊴⋅ s = (u ⊴⋅ (π ∙ s))"
proof
  from term_apply_subst_Var_Rep_perm obtain σ where sub: "π ∙ s = s ⋅ σ" by metis
  show "u ⊴⋅ s ⟹ u ⊴⋅ π ∙ s" unfolding encompeq.simps sub
    by (metis subst_apply_term_ctxt_apply_distrib subst_subst)
next
  have "s = ( -π) ∙ π ∙ s" by auto
  with term_apply_subst_Var_Rep_perm obtain σ where sub: "s = (π ∙ s) ⋅ σ" by metis
  show "u ⊴⋅ π ∙ s ⟹ u ⊴⋅ s" unfolding encompeq.simps using sub
    by (metis subst_apply_term_ctxt_apply_distrib subst_subst)
qed

abbreviation "lexencomp ≡ lex_two {⋅⊳} {⋅⊵}  {(a, b :: bool) . a ∧ ¬ b}"

lemma SN_lexencomp: "SN lexencomp"
proof-
  have t: "SN {(a, b). a ∧ ¬ b}" using SN_inv_image [OF SN_nat_gt, of "λb. if b then 1 else 0"]
    by (metis (no_types, lifting) SN_onI mem_Collect_eq split_conv)
  then show ?thesis using SN_encomp encomp_order.less_le_trans encomp_order.le_less_trans
    by(intro lex_two, auto)
qed

lemma gterms_ER_conv_implies_S_conv:
  assumes RR_less: "RR ⊆ {≻}"
  assumes ground: "ground s" "ground t"
  assumes conv: "(s, t) ∈ (rstep (RR ∪ EE))*" (is "_ ∈ ?R*")
  shows "(s, t) ∈ (GROUND (rstep (RR ∪ E_ord (op ≻) EE)))*"
proof -
  from gterm_conv_GROUND_conv assms have "(s, t) ∈ (GROUND (rstep (RR ∪ EE)))*" by auto
  with conversion_mono [OF GROUND_rstep_ordstep] show ?thesis unfolding conversion_conversion_idemp by auto
qed

lemma reduced_ground_complete:
  assumes "𝒮 ℛ ℰ ⊆ {≻}" and "GCR (rstep (𝒮 ℛ ℰ)) ∧ SN (GROUND (rstep (𝒮 ℛ ℰ)))"
  shows "GCR (rstep 𝒮') ∧
         SN (GROUND (rstep 𝒮')) ∧
         (GROUND (rstep (𝒮 ℛ ℰ)))! = (GROUND (rstep 𝒮'))!"
proof-
  let ?S = "ℛ' ∪ (E_ord (op ≻) ℰ')"
  have sub_dot: "ℛ' ⊆ dot (ℛ ∪ ℰsucc)" unfolding ℛ'_def by fast
  with rstep_subset_less ‹ℛ ⊆ {≻}› have **: "ℛ ∪ ℰsucc ⊆ {≻}" unfolding rstep_union by auto
  with rstep_subset_less ‹ℛ ⊆ {≻}› have *: "rstep (ℛ ∪ ℰsucc) ⊆ {≻}" unfolding rstep_union by auto
  with SN_less SN_subset have "SN (rstep (ℛ ∪ ℰsucc))" by auto
  interpret SN_trs "ℛ ∪ ℰsucc" by standard fact

  { fix l r
    assume a: "(l, r) ∈ dot' (ℛ ∪ ℰsucc)"
    then have "∃ r'. (l, r') ∈ (ℛ ∪ ℰsucc) ∧ r = some_NF (rstep (ℛ ∪ ℰsucc)) r'"
      unfolding dot'_def mem_Collect_eq by fast
    with some_NF_rtrancl rstep.rule [OF a] have "(l, r) ∈ (rstep (ℛ ∪ ℰsucc))+"
      by (meson rstep.rule rtrancl_into_trancl2)
  }
  with rsteps_subset_less [OF **] rm_variants_subset have "dot (ℛ ∪ ℰsucc) ⊆ {≻}"
    unfolding dot_def by blast
  with sub_dot rstep_subset_less have R'succ: "ℛ' ⊆ {≻}" by auto
  with rstep_subset_less have "rstep ℛ' ⊆ {≻}" unfolding E_ord_def by blast
  with SN_less SN_subset have "SN (rstep ℛ')" unfolding 𝒮'_def by metis
  from R'succ rstep_subset_less have "rstep ?S ⊆ {≻}" unfolding E_ord_def by blast
  with SN_less SN_subset GROUND_subset have x: "SN (GROUND (rstep 𝒮'))" unfolding 𝒮'_def by metis

  { fix s t
    assume "(s, t) ∈ GROUND (rstep (𝒮 ℛ ℰ))"
    then have g: "ground s" "ground t" "(s, t) ∈ rstep (𝒮 ℛ ℰ)" unfolding GROUND_def by auto
    { fix l r
      assume in_S: "(l, r) ∈ 𝒮 ℛ ℰ"
      let ?orig = "λ rl. ¬ rl ∈ ℛ ∪ ℰsucc"
      from in_S have "∃ r'. (l, r') ∈ rstep 𝒮'"
      proof (induct "(l, ?orig (l, r))" arbitrary: l r rule: SN_induct [OF SN_lexencomp])
        case (1 l)
        then show ?case
        proof (cases "(l, r) ∈ ℛ ∪ ℰsucc")
          case True
          note REsucc = this
          show ?thesis
          proof (cases "∃u l0 r0. (l, u) ∈ rstep {(l0, r0)} ∧ (l0, r0) ∈ 𝒮 ℛ ℰ ∧ l0 ⊲⋅ l")
            case True
            then obtain u l0 r0 where step: "(l, u) ∈ rstep {(l0, r0)}" "(l0, r0) ∈ 𝒮 ℛ ℰ" and ll0: "l0 ⊲⋅ l" by blast
            then obtain C σ where C: "l = C⟨l0 ⋅ σ⟩" "(l0, r0) ∈ 𝒮 ℛ ℰ" using rstep.simps by auto
            from ll0 have "((l, ?orig (l, r)), (l0, ?orig (l0, r0))) ∈ lexencomp" by force
            with 1(1) [OF this] step have "∃ r0'. (l0, r0') ∈ rstep 𝒮'" by auto
            with C(1) rstep.ctxt show ?thesis by force
          next
            case False
            from REsucc obtain u where in_dot: "(l, u) ∈ dot' (ℛ ∪ ℰsucc)" unfolding dot'_def by auto
            from rm_variants_representative [OF this]
            obtain π l' r' where var: "(l', r') ∈ dot(ℛ ∪ ℰsucc)" "π ∙ (l, u) = (l', r')"
              unfolding dot_def by blast
            then have "π ∙ l = l' ∧ π ∙ u = r'" by (simp add: rule_pt.permute_prod.simps)
            then have lu: "l = ( - π) ∙ l'" "u = ( - π) ∙ r'" by auto
            { assume "∃u l0 r0. (l', u) ∈ rstep {(l0, r0)} ∧ (l0, r0) ∈ 𝒮 ℛ ℰ ∧ l0 ⊲⋅ l'"
              then obtain v l0 r0 where step': "(l', v) ∈ rstep {(l0, r0)} "
                and in_S: "(l0, r0) ∈ 𝒮 ℛ ℰ" and enc: "l0 ⊲⋅ l'" by blast
              with rstep_r_p_s_imp_rstep have step: "(l', v) ∈ rstep {(l0, r0)}" by auto
              with permute_rstep lu have step: "(l, ( - π) ∙ v) ∈ rstep {(l0, r0)}" by auto
              from encomp_perm enc lu have enc: "l0 ⊲⋅ l" unfolding encomp_def
                by (metis encomp_def encomp_order.le_less_trans encompeq_refl)
              with False enc step in_S have False by blast
            }
            with var have "(l', r') ∈ ℛ'" unfolding ℛ'_def mem_Collect_eq by blast
            with var(2) perm_rstep_conv rstep.rule have "(l, u) ∈ rstep ℛ'" by metis
            then show ?thesis unfolding rstep_union 𝒮'_def by fast
          qed
        next
          case False
          note unorientable = this
          with 1(2) obtain u v σ where uv: "(u, v) ∈ ℰ" "l = u ⋅ σ" "r = v ⋅ σ" "u ⋅ σ ≻ v ⋅ σ"
            unfolding E_ord_def by blast
          let ?u = "some_NF (rstep ℛ') u"
          let ?v = "some_NF (rstep ℛ') v"
          interpret SN_R': SN_ars "rstep ℛ'" by standard fact
          from ‹SN (rstep ℛ')› have uu: "(u, ?u) ∈ (rstep ℛ')*"
            by (simp add: SN_ars.some_NF_rtrancl Normalization_Equivalence.SN_ars.intro)
          from ‹SN (rstep ℛ')› have vv: "(v, ?v) ∈ (rstep ℛ')*"
            by (simp add: SN_ars.some_NF_rtrancl Normalization_Equivalence.SN_ars.intro)
          show ?thesis
          proof (cases "?u = u")
            case True
            from vv rsteps_subset_less [OF ‹ℛ' ⊆ {≻}›] have "v ≽ ?v"
              unfolding rtrancl_eq_or_trancl by auto
            with subst [of v ?v σ] trans [OF uv(4)] uv(4) have gt: "u ⋅ σ ≻ ?v ⋅ σ" by auto
            with irrefl have "u ≠ ?v" by auto
            with uv(1) True have in_E': "(u, ?v) ∈ ℰ'"
              unfolding ℰ'_def Un_iff mem_Collect_eq converse_iff by metis
            from gt uv in_E' have "(l, ?v ⋅ σ) ∈ E_ord op ≻ ℰ'"
              unfolding E_ord_def mem_Collect_eq by blast
            with uv show ?thesis unfolding rstep_union 𝒮'_def by blast
          next
            case False
            with uu have "(u, ?u) ∈ (rstep ℛ')+" by (simp add: rtrancl_eq_or_trancl)
            with tranclD obtain w where "(u, w) ∈ (rstep ℛ')" "(w, ?u) ∈ (rstep ℛ')*"
              by (meson tranclD)
            then obtain l0 r0 C τ where l0r0: "u = C⟨l0 ⋅ τ⟩" "w = C⟨r0 ⋅ τ⟩" "(l0, r0) ∈ ℛ'" by auto
            then have "(l0, r0) ∈ dot (ℛ ∪ ℰsucc)" unfolding ℛ'_def by auto
            with lhss_dot obtain r0' where oriented: "(l0, r0') ∈ ℛ ∪ ℰsucc" by fastforce
            from Esucc_E_ord oriented have in_S: "(l0, r0') ∈ 𝒮 ℛ ℰ" unfolding Un_iff by auto
            from uv(2) have "l ⋅⊵ l0" unfolding encompeq.simps l0r0(1)
                subst_apply_term_ctxt_apply_distrib subst_subst by fast
            with unorientable oriented have "((l, ?orig (l, r)), (l0, ?orig (l0, r0'))) ∈ lexencomp" by simp
            from 1(1) [OF this in_S] show ?thesis unfolding uv l0r0 by fast
          qed
        qed
      qed
    }
    with g have "∃t'. (s, t') ∈ rstep 𝒮'" by fast
    then obtain t' where step: "(s, t') ∈ rstep 𝒮'" by auto
    with ‹rstep ?S ⊆ {≻}› ground_less [OF g(1)] have "ground t'" unfolding 𝒮'_def by auto
    with g step have "∃t'.(s, t') ∈ GROUND (rstep 𝒮')" unfolding GROUND_def by auto
  }
  then have NFs: "NF (GROUND (rstep 𝒮')) ⊆ NF (GROUND (rstep (𝒮 ℛ ℰ)))" unfolding NF_def by auto

  have step_conv: "GROUND(rstep 𝒮') ⊆ (GROUND (rstep (𝒮 ℛ ℰ)))*"
  proof
    fix s t
    assume "(s, t) ∈ GROUND(rstep 𝒮')"
    then have g: "ground s" "ground t" "(s, t) ∈ rstep 𝒮'" unfolding GROUND_def by auto
    then obtain l r σ C where C: "s = C⟨l ⋅ σ⟩" "t = C⟨r ⋅ σ⟩" "(l, r) ∈ 𝒮'" by auto
    { fix l r
      assume "(l, r) ∈ ℛ'"
      with rm_variants_subset have "(l, r) ∈ dot' (ℛ ∪ ℰsucc)" unfolding ℛ'_def mem_Collect_eq dot_def by fast
      then obtain r' where r: "r = some_NF (rstep (ℛ ∪ ℰsucc)) r'" "(l, r') ∈ ℛ ∪ ℰsucc" unfolding dot'_def  by auto
      interpret REsucc: SN_ars "rstep (ℛ ∪ ℰsucc)" by standard fact
      from ‹SN (rstep (ℛ ∪ ℰsucc))› have "(r', r) ∈ (rstep (ℛ ∪ ℰsucc))*" unfolding r
        by (simp add: SN_ars.some_NF_rtrancl Normalization_Equivalence.SN_ars.intro)
      with rstep_mono [of "ℛ ∪ ℰsucc" "ℛ ∪ ℰ", THEN rtrancl_mono] have rr: "(r', r) ∈ (rstep  (ℛ ∪ ℰ))*" by blast
      from r(2) have "(l, r') ∈ rstep (ℛ ∪ ℰ)" by fastforce
      with rr have "(l, r) ∈ (rstep (ℛ ∪ ℰ))*" unfolding conversion_def
        by (meson converse_rtrancl_into_rtrancl in_rtrancl_UnI)
    } note R'_RE = this
    {  fix s t
      assume "(s, t) ∈ rstep ℛ'"
      then obtain l r C σ where "(l, r) ∈ ℛ'" "s = C⟨l ⋅ σ⟩" "t = C⟨r ⋅ σ⟩" by auto
      with R'_RE [OF this(1)] ctxt.closed_conversion [OF ctxt_closed_rstep, THEN ctxt.closedD]
        subst.closed_conversion [OF subst_closed_rstep, THEN subst.closedD] have "(s, t) ∈ (rstep (ℛ ∪ ℰ))*" by metis
    } note R'_step_RE = this
    then have "rstep ℛ' ⊆ (rstep (ℛ ∪ ℰ))*" by auto
    from conversion_conversion_idemp conversion_mono [OF this]
    have R'_steps_RE: "⋀ s t.(s, t) ∈ (rstep ℛ')* ⟹ (s, t) ∈ (rstep (ℛ ∪ ℰ))*" by auto
    show "(s, t) ∈ (GROUND (rstep (𝒮 ℛ ℰ)))*"
    proof (cases "(l, r) ∈ ℛ'")
      case True
      from R'_RE [OF this] ctxt.closed_conversion [OF ctxt_closed_rstep, THEN ctxt.closedD]
        subst.closed_conversion [OF subst_closed_rstep, THEN subst.closedD]
      have "(s, t) ∈ (rstep (ℛ ∪ ℰ))*" unfolding C by metis
      with gterms_ER_conv_implies_S_conv [OF ‹ℛ ⊆ {≻}› g(1) g(2)] show ?thesis .
    next
      case False
      with C(3) [unfolded 𝒮'_def] have "(l, r) ∈ E_ord op ≻ ℰ'" by auto
      then obtain u v τ where uv: "l = u⋅ τ" "r = v ⋅ τ" "(u, v) ∈ ℰ'" "u⋅ τ ≻ v ⋅ τ"
        unfolding E_ord_def mem_Collect_eq by blast
      from uv(3) obtain u' v' where uv': "(u', v') ∈ ℰ" "u = some_NF (rstep ℛ') u'" "v = some_NF (rstep ℛ') v'"
        unfolding ℰ'_def Un_iff mem_Collect_eq by auto
      interpret SN_R': SN_ars "rstep ℛ'" by standard fact
      from ‹SN (rstep ℛ')› have "(u', u) ∈ (rstep ℛ')*" unfolding uv'
        by (simp add: SN_ars.some_NF_rtrancl Normalization_Equivalence.SN_ars.intro)
      with rsteps_closed_subst rsteps_closed_ctxt have "(C⟨u' ⋅ τ ⋅ σ⟩, C⟨u ⋅ τ ⋅ σ⟩) ∈ (rstep ℛ')*" by metis
      with R'_steps_RE conversion_inv have uconv: "(C⟨u ⋅ τ ⋅ σ⟩, C⟨u' ⋅ τ ⋅ σ⟩) ∈ (rstep (ℛ ∪ ℰ))*" by metis
      from ‹SN (rstep ℛ')› uv' have "(v', v) ∈ (rstep ℛ')*"
        by (simp add: SN_ars.some_NF_rtrancl Normalization_Equivalence.SN_ars.intro)
      with rsteps_closed_subst rsteps_closed_ctxt have "(C⟨v' ⋅ τ ⋅ σ⟩, C⟨v ⋅ τ ⋅ σ⟩) ∈ (rstep ℛ')*" by metis
      with R'_steps_RE have vconv: "(C⟨v' ⋅ τ ⋅ σ⟩, C⟨v ⋅ τ ⋅ σ⟩) ∈ (rstep (ℛ ∪ ℰ))*" by auto
      from uv' uv rstepI have "(C⟨u' ⋅ τ ⋅σ⟩, C⟨v' ⋅ τ ⋅σ⟩) ∈ rstep (ℰ)" by auto
      then have "(C⟨u' ⋅ τ ⋅σ⟩, C⟨v' ⋅ τ ⋅σ⟩) ∈ (rstep (ℛ ∪ ℰ))*" unfolding conversion_def using r_into_rtrancl
        by fastforce
      from conversion_trans' [OF uconv this, THEN conversion_trans'] vconv
      have "(s, t) ∈ (rstep (ℛ ∪ ℰ))*" unfolding C uv by auto
      with gterms_ER_conv_implies_S_conv [OF ‹ℛ ⊆ {≻}› g(1) g(2)] show ?thesis .
    qed
  qed

  from assms have CR_SN: "CR (GROUND (rstep (𝒮 ℛ ℰ))) ∧ SN (GROUND (rstep (𝒮 ℛ ℰ)))"
    by (meson GROUND_subset SN_subset)
  interpret complete_ars "GROUND (rstep (𝒮 ℛ ℰ))" by (standard, insert CR_SN, auto)
  from complete_NE_intro1 [OF step_conv x NFs] complete_ars.CR show ?thesis
    unfolding complete_ars_def Normalization_Equivalence.SN_ars_def by auto
qed
end


    
lemma Einf_sym_without_Ew_sym: assumes "(l,r) ∈ E" and "(l,r) ∉  Eω"
  shows "∃j. ((l, r) ∈ (E j) ∧ (l, r) ∉ (E (Suc j))) ∨ ((r, l) ∈ (E j) ∧ (r, l) ∉ (E (Suc j)))"
by (cases "(l,r) ∈ E", insert assms Einf_without_Ew, auto)

lemma Ew_reducible_aux:
  assumes "(∃r'. (l,r') ∈ Eω ∧ l⋅σ ≻ r'⋅σ) ∨ l ∉ NF (ordstep {≻} (R ∪ Eω))"
  shows "l⋅σ ∉ NF (ordstep {≻} (R ∪ Eω))"
proof-
  from assms consider "∃r'. (l,r') ∈ Eω ∧ l⋅σ ≻ r'⋅σ" | "l ∉ NF (ordstep {≻} (R ∪ Eω))" by auto
  thus ?thesis proof(cases)
    case 1
    then obtain r' where r':"(l, r') ∈ Eω" "l ⋅ σ ≻ r' ⋅ σ" by auto
    have "(l⋅σ, r'⋅σ) ∈ ordstep {≻} (R ∪ Eω)"
      by (rule ordstep.intros[of _ _ _ _ ], insert r', auto)
    then show ?thesis by auto
  next
    case 2
    then obtain t where "(l, t) ∈ ordstep {≻} (R ∪ Eω)" by auto
    from ordstep_subst_ctxt[OF subst_closed_less this, of ]
      have "(l ⋅ σ, t ⋅ σ) ∈ ordstep {≻} (R ∪ Eω)" by auto
    thus ?thesis by auto
  qed  
qed
  
lemma Ei_oriented_reducible_in_RiEw: 
  assumes "(l,r) ∈ E" and "l⋅σ ≻ r⋅σ"
  shows "(∃r'. (l,r') ∈ Eω ∧ l⋅σ ≻ r'⋅σ) ∨ l ∉ NF (ordstep {≻} (R ∪ Eω))"
  using assms
proof (induct "{#l, r#}" arbitrary: l r σ rule: SN_induct [OF lessenc_op.SN])
  case 1
  note IH = this
  have comm: "⋀t u. {#t, u#} =  {#u, t#}" using add_mset_commute by auto
  show ?case
  proof (cases "(l, r) ∈ Eω")
    case True
    with 1 show ?thesis by auto
  next
    case False
    with 1(2) Einf_sym_without_Ew_sym obtain j where lr_cases:
      "((l, r) ∈ (E j) ∧ (l, r) ∉ (E (Suc j))) ∨ ((r, l) ∈ (E j) ∧ (r, l) ∉ (E (Suc j)))" by auto
    define u' where "u' ≡ if (l, r) ∈ (E j) ∧ (l, r) ∉ (E (Suc j)) then l else r"
    define v' where "v' ≡ if (l, r) ∈ (E j) ∧ (l, r) ∉ (E (Suc j)) then r else l"
    from lr_cases have uv:"(u', v') ∈ (E j)" "(u', v') ∉ E (Suc j)" by (auto simp: u'_def v'_def)
    have uv_cases:"(u' = l ∧ v' = r) ∨ (u' = r ∧ v' = l)" by (auto simp: u'_def v'_def)
    
    let ?ordstep = "ordstep {≻} (R ∪ Eω)"
    from ordstep_rstep_conv[OF _ Rinf_less] subst.closedI[of "{≻}"] subst have
      ordstep_rstep:"ordstep {≻} R = rstep R" by auto
    note ordstep_mono = ordstep_mono[of _ "R ∪ Eω", OF _ subset_refl, of _ "{≻}"]
      
    from subst[of r l σ] trans[OF IH(3)] irrefl have not_r_succ_l:"¬ (r ≻ l)" by auto
    let ?enc = "encstep1 (E (Suc j)) (R (Suc j)) O (E (Suc j))"
    from uv oKBi_E_supset' [OF irun [of j]]
      consider "(u', v') ∈ R (Suc j)" | "(v', u') ∈ R (Suc j)" | "u' = v'" | "(u', v') ∈ ?enc"
      by auto blast
    then show ?thesis
    proof (cases)
      case 1
      from not_r_succ_l Ri_less[of "Suc j"] 1 have uv:"u' = l" "v' = r" using uv_cases by auto
      from 1 ordstep_rstep have "(l,r) ∈ ordstep {≻} R" unfolding uv by auto
      with ordstep_mono show ?thesis by auto
    next
      case 2
      from not_r_succ_l Ri_less[of "Suc j"] 2 have uv:"u' = r" "v' = l" using uv_cases by auto
      from 2 ordstep_rstep have "(l,r) ∈ ordstep {≻} R" unfolding uv by auto
      with ordstep_mono show ?thesis by auto
    next
      case 3
      from irrefl IH(3) uv_cases show ?thesis unfolding 3 by auto
    next
      case 4
      then obtain u v w where v: "(u, w) ∈ encstep1 (E (Suc j)) (R (Suc j))" "(w, v) ∈ (E (Suc j))"
        and uv_cases:"(u = l ∧ v = r) ∨ (u = r ∧ v = l)" using uv_cases by auto
          
      show ?thesis proof(cases "u = l")
        case True
        from v(1) have "u ∉ NF (ordstep {≻} (R ∪ Eω))"
        proof (cases)
          case (estep l0 r0 D τ)
          from encompeq.intros [of u D "l0 ⋅ τ" Var] encompeq.intros [of "r0 ⋅ τ" ] estep(2)
            have u_l0:"l0 ⋅ τ ⊴⋅ u" "r0 ⊴⋅ r0 ⋅ τ" by auto
          with estep have ur0: "u ⋅≻ r0" "u ⋅≻ l0" unfolding lessencp_def by blast+
          note mul = ns_s_mul_ext_union_multiset_l[OF ns_mul_ext_bottom, of "{#u#}" "{#l0, r0#}"]    
          with ur0 have "({#u, v#}, {#l0, r0#}) ∈ mulless" using multi_member_last by auto
          with uv_cases comm[of l r] have "({#l, r#}, {#l0, r0#}) ∈ mulless" by presburger
          from IH(1)[OF this _ estep(4)] estep(1) Ew_reducible_aux  
            have NF1:"l0 ⋅ τ ∉ NF (ordstep {≻} (R ∪ Eω))" by auto
          from u_l0 obtain Cu σu where u:"u = Cu⟨l0 ⋅ τ ⋅ σu⟩" unfolding encompeq.simps by auto
          from NF1 ordstep_subst_ctxt[OF subst_closed_less] show ?thesis unfolding u by blast
        next
          case rstep
          with ordstep_rstep have "(u, w) ∈ ordstep {≻} R" by fast
          with ordstep_mono[of "R"] show ?thesis by auto
        qed
        with True show ?thesis by auto
      next
        case False
        with uv_cases have uv:"u = r" "v = l" by auto
        from v(1) encstep1_less Ri_less have rw:"r ≻ w" unfolding uv by auto
        from rw mset_two2[of r w l] have "({#l, r#}, {#l, w#}) ∈ mulless" using lessencp_def by fast
        from IH(1)[OF this _ trans[OF IH(3) subst, OF rw]] v(2) show ?thesis unfolding uv by fast
      qed
    qed
  qed
qed

lemma Sw_SS:"Sω = 𝒮 Rω Eω" unfolding Sw_def E_ord_def by auto
  
lemma Ri_reducible_in_RwEw: 
  assumes "(l,r) ∈ R"
  shows "l ∉ NF (rstep Sω)"
  using assms
proof (induct "(l, r)" arbitrary: l r rule: SN_induct [OF SN_lexless])
  case 1
  note IH = this
  let ?ordstep = "ordstep {≻} (Rω ∪ Eω)"
  show ?case 
  proof (cases "(l, r) ∈ Rω")
    case True
    thus ?thesis unfolding Sw_def by auto
  next
    case False
    with Rinf_without_Rw[OF IH(2)] obtain i where
      not_i: "(l, r) ∉ R (Suc i)" and pred_i: "(l, r) ∈ R i" by auto
    let ?Ri = "R (Suc i)" and ?Ei = "E (Suc i)"
    from pred_i oKBi_R_supset [OF irun [of i]] and not_i
      consider "(l, r) ∈ encstep2 ?Ei ?Ri O ?Ei" | "(l, r) ∈ ?Ri O (ostep ?Ei ?Ri)¯" by blast
    then show ?thesis
    proof (cases)
      case 1
      then obtain u where u: "(l, u) ∈ encstep2 ?Ei ?Ri" by auto
      thus ?thesis
      proof(cases)
        case (estep l0 r0 C σ)
        with Ei_oriented_reducible_in_RiEw[OF _ estep(4)] consider
          "∃r'. (l0, r') ∈ Eω ∧ l0 ⋅ σ ≻ r' ⋅ σ" | "l0 ∉ NF (ordstep {≻} (R ∪ Eω))" by auto
        thus ?thesis proof(cases)
          case 1
          then obtain r' where r':"(l0, r') ∈ Eω" "l0 ⋅ σ ≻ r' ⋅ σ" by auto
          hence "(l0 ⋅ σ, r' ⋅ σ) ∈ rstep Sω" unfolding Sw_def by auto
          thus ?thesis unfolding estep by fast
        next
          case 2
          then obtain t where "(l0, t) ∈ ordstep {≻} (R ∪ Eω)" by auto
          with ordstep_Un consider "(l0, t) ∈ ordstep {≻} (Eω)" | "(l0, t) ∈ ordstep {≻} R" by auto
          thus ?thesis proof(cases)
            case 1
            from ordstep_subst_ctxt[OF subst_closed_less 1] ordstep_mono
              have "(l, C⟨t ⋅ σ⟩) ∈ ordstep {≻} (Eω)" unfolding estep by auto
            with ordstep_mono[of "Eω" "Rω ∪ Eω"]  have "(l, C⟨t ⋅ σ⟩) ∈ ?ordstep" by auto
            with Sw_SS ordstep_subset_S[OF Rw_less]
              have "(l, C⟨t ⋅ σ⟩) ∈ rstep Sω" unfolding ordstep_S_conv[OF Rw_less] by auto
            thus ?thesis by fast
          next
            case 2
            from ordstep_imp_rstep[OF this] obtain l1 r1 D τ
              where lr1:"l0 = D⟨l1⋅τ⟩" "t = D⟨r1⋅τ⟩" "(l1,r1) ∈ R" by fast
            from encompeq.intros[OF lr1(1)] estep(5) have "l ⋅≻ l1" using lessencp_def by auto
            hence "((l, r), l1, r1) ∈ lexless" by auto
            from IH(1)[OF this lr1(3)] obtain t' where lt:"(l1, t') ∈ rstep Sω" by auto
            thus ?thesis unfolding lr1 estep by fast
          qed
        qed
      next
        case (rstep l0 r0 C σ)
        from rstep(4) have "l ⋅≻ l0" using lessencp_def by auto
        hence "((l, r), l0, r0) ∈ lexless" using lessencp_def by auto
        from IH(1)[OF this] rstep(1) have "l0 ∉ NF_trs Sω" by auto
        then show ?thesis unfolding rstep(2) by blast
      qed
    next
      case 2
      then obtain r' where r':"(l,r') ∈ ?Ri" "(r,r') ∈ ostep ?Ei ?Ri" by fast
      from ostep_imp_less[OF Ri_less this(2)] have "r ⋅≻ r'" using lessencp_def by fast
      hence "((l, r), l, r') ∈ lexless" by auto
      from IH(1)[OF this] r'(1) show ?thesis by auto
    qed
  qed
qed

lemma Ei_oriented_reducible_in_RwEw: 
  assumes "(l,r) ∈ E" and "l⋅σ ≻ r⋅σ"
  shows "l⋅σ ∉ NF_trs Sω"
proof-
  from Ew_reducible_aux[OF Ei_oriented_reducible_in_RiEw[OF assms]]
    have "∃t. (l⋅σ, t) ∈ ordstep {≻} (R ∪ Eω)" by auto
  then obtain t where t:"(l⋅σ, t) ∈ ordstep {≻} (R ∪ Eω)" by auto
  { assume "(l⋅σ, t) ∈ ordstep {≻} R"
    from ordstep_imp_rstep[OF this] obtain C τ l0 r0 where
      lr0:"(l0, r0) ∈ R" and lt:"l ⋅ σ = C⟨l0 ⋅ τ⟩" by fast
    from Ri_reducible_in_RwEw[OF lr0] have "l ⋅ σ ∉ NF_trs Sω" unfolding lt by fastforce
  } note Rinf = this
  { assume a:"(l⋅σ, t) ∈ ordstep {≻} (Eω)"
    with ordstep_mono[of "Eω" "Rω ∪ Eω"]  have "(l⋅σ, t) ∈ ordstep {≻} (Rω ∪ Eω)" by auto
            with Sw_SS ordstep_subset_S[OF Rw_less]
              have "(l⋅σ, t) ∈ rstep Sω" unfolding ordstep_S_conv[OF Rw_less] by auto
  }
  with t Rinf show ?thesis unfolding ordstep_Un by auto
qed
  
lemma NF_Sw_subset_NF_Sinf: "NF_trs Sω ⊆ NF_trs (𝒮 R E)"
proof-
  { fix s t
    assume "(s,t) ∈ rstep (𝒮 R E)"
    with rstep_union consider "(s,t) ∈ rstep R" | "(s,t) ∈ rstep (E_ord (op ≻) E)" by fast
    hence "s ∉ NF_trs Sω" proof(cases)
      case 1
      from rstepE[OF this] NF_subterm[of s "R"] show ?thesis using Ri_reducible_in_RwEw
        by (metis NF_instance NF_subterm ctxt_imp_supteq)
    next
      case 2
      then obtain l r C σ where "(l,r) ∈ E_ord (op ≻) E" and st:"s = C⟨l⋅σ⟩" "t = C⟨r⋅σ⟩" by auto
      then obtain l' r' τ where lr:"l = l' ⋅ τ" "r = r' ⋅ τ" and ord:"(l', r') ∈ E" "l' ⋅ τ ≻ r' ⋅ τ"
        unfolding E_ord_def by fast
      from subst[OF ord(2)] Ei_oriented_reducible_in_RwEw[OF ord(1), of "τ ∘s σ"] show ?thesis
        unfolding st lr by fastforce
    qed
  }
  thus ?thesis unfolding NF_def by auto
qed

lemma Sw_subset_Sinf:"Sω ⊆ 𝒮 R E"
  unfolding Sw_SS using Rw_subset_Rinf E_ord_mono[OF Ew_subset_Einf] by auto

lemma NF_Sw_NF_Sinf: "NF_trs Sω = NF_trs (𝒮 R E)"
  using NF_Sw_subset_NF_Sinf NF_trs_mono[OF Sw_subset_Sinf] by auto

lemma SN_Sinf:"SN (rstep (𝒮 R E))"
  using ordstep_S_conv[OF Rinf_less] by (metis SN_ordstep)

context
  assumes fair: "∀(s, t)∈ PCP_ext (Rω ∪ Eω). (s, t) ∈ (rstep E)"
begin
  
lemma msteps_imp_source_steps:
  assumes "∀ t ∈# M. s ≻ t"
  shows "(mstep M ℛ)* ⊆ (⋃z ∈ {z. s ≻ z}. slab ℛ z)*"
proof-
  have "mstep M ℛ ⊆ (⋃z ∈ {z. s ≻ z}. slab ℛ z)"
  proof
    fix t u
    assume a: "(t, u) ∈ mstep M ℛ"
    then have tu: "(t, u) ∈ rstep ℛ" unfolding mstep_def mem_Collect_eq by auto
    from a have "∃ v ∈# M. v ≽ t" unfolding mstep_def mem_Collect_eq by auto
    with assms transD [OF trans_less] have st: "s ≻ t" by blast
    from tu have "(t, u) ∈ slab ℛ t" unfolding source_step_def by auto
    with st show "(t, u) ∈ (⋃z ∈ {z. s ≻ z}. slab ℛ z)" by auto
  qed
  from conversion_mono [OF this] show ?thesis by auto
qed

text ‹Establish peak decreasingness of ‹Sω› steps on ground terms for correctness result below.›
lemma ground_Sw_peak_decreasing:
  assumes "ground s" and "(s, t) ∈ slab Sω s" and "(s, u) ∈ slab Sω s"
  shows "(t, u) ∈ (⋃ v ∈ {v. ground v ∧ s ≻ v}. slab Sω v)*"
proof-
  from assms have rsteps: "(s, t) ∈ rstep Sω" "(s, u) ∈ rstep Sω" unfolding source_step_def by auto
  note vc = SN_imp_variable_condition [OF SN_Sw_step]
  from vc peak_imp_nabla2 [OF _ rsteps] have tu: "(t, u) ∈ nabla Sω s ^^ 2" by auto
  have Sw_less: "Sω ⊆ {≻}" using Ri_less unfolding Sw_def Rω_def by fast
  from compatible_rstep_imp_less [OF Sw_less] have Sw_step_less: "rstep Sω ⊆ {≻}" by blast
  let ?less_s_conv = "(⋃z∈{z. s ≻ z}. slab Sω z)*"
  let ?leq_conv = "λ v. (⋃z∈{z. v ≽ z}. slab Sω z)*"
  { fix v w
    assume vw: "(v, w) ∈ nabla Sω s" and "ground v" and "ground w" and "s ≻ v" and "s ≻ w"
    then have vw: "(v, w) ∈ (rstep Sω) ∨ (v, w) ∈ (rstep (PCP Sω))" unfolding nabla_def by auto
    { assume "(v, w) ∈ (rstep Sω)"
      then obtain x where vx: "(v, x) ∈ (rstep Sω)*" and wx: "(w, x) ∈ (rstep Sω)*"
        unfolding join_def rtrancl_converse by blast
      from rsteps_slabI [OF vx _ Sw_step_less, of v] have "(v, x) ∈ ?leq_conv v" by auto
      from slab_conv_less_label [OF this] ‹s ≻ v› have vx: "(v, x) ∈ ?less_s_conv" by auto
      from rsteps_slabI [OF wx _ Sw_step_less, of w] have "(w, x) ∈ ?leq_conv w" by auto
      from slab_conv_less_label [OF this] ‹s ≻ w› have "(x, w) ∈ ?less_s_conv"
        unfolding conversion_inv by auto
      from conversion_trans' [OF vx this] have "(v, w) ∈ ?less_s_conv" by auto
    } note if_joinable = this
    { assume "(v, w) ∈ (rstep (PCP Sω))"
      from PCP_xPCP [OF Rw_less ‹ground v› ‹ground w› this [unfolded Sw_def, folded E_ord_def]]
      have vw: "(v, w) ∈ (rstep Sω) ∨ (v, w) ∈ (rstep (PCP_ext (Rω ∪ Eω)))"
        unfolding Sw_def E_ord_def by auto
      { assume "(v, w) ∈ (rstep (PCP_ext (Rω ∪ Eω)))"
        with fair rstep_rstep have "(v, w) ∈ (rstep E)" by fast
        from ground_Einf_Sw [OF this ‹ground v› ‹ground w›] have vw: "(v, w) ∈ (mstep {#v, w#} Sω)*" .
        from ‹s ≻ v› ‹s ≻ w› have "∀ t ∈# {#v, w#}. s ≻ t" by auto
        from vw msteps_imp_source_steps [OF this] have "(v, w) ∈ ?less_s_conv" by auto
      }
      with vw if_joinable have "(v, w) ∈ ?less_s_conv" by auto
    }
    with vw if_joinable have "(v, w) ∈ ?less_s_conv" by auto
  } note nabla_slab = this
  from tu obtain v where v: "(t, v) ∈ nabla Sω s" "(v, u) ∈ nabla Sω s" by auto
  note gimp = rstep_ground [OF _ ‹ground s›]
  from this [OF _ rsteps(1)] this [OF _ rsteps(2)] vc have "ground t" "ground u" by auto
  from v [unfolded nabla_def] rstep_ground [OF _ ‹ground s›] have sv: "(s, v) ∈ (rstep Sω)+" by auto
  from rsteps_ground [OF _ ‹ground s› trancl_into_rtrancl [OF this]] vc have "ground v" by fast
  from sv rsteps_subset_less [OF Sw_less] have sv: "s ≻ v" by auto
  from Sw_step_less rsteps have st: "s ≻ t" and su: "s ≻ u" by auto
  note tv = nabla_slab [OF v(1) ‹ground t› ‹ground v› st sv]
  note vu = nabla_slab [OF v(2) ‹ground v› ‹ground u› sv su]
  from ground_less [OF ‹ground s›] have "{v. ground v ∧ s ≻ v} = {v. s ≻ v}" by auto
  with conversion_trans' [OF tv vu] show ?thesis by auto
qed

lemma correctness_okb: "GCR (rstep Sω) ∧ SN (rstep Sω) ∧ (rstep (E 0))* = (rstep (Rω ∪ Eω))*"
proof -
  interpret ars_peak_decreasing "slab Sω" "{t. ground t}" "op ≻"
    by (unfold_locales, insert ground_Sw_peak_decreasing SN_less, auto)
  from CR have CR: "CR (⋃t ∈ {t. ground t}. slab Sω t)" by auto
  from rstep_ground SN_imp_variable_condition [OF SN_Sw_step]
  have "⋀ s t. ground s ⟹ (s, t) ∈ rstep Sω ⟹ ground t" by fastforce
  then have "(⋃t ∈ {t. ground t}. slab Sω t) = GROUND (rstep Sω)" unfolding source_step_def GROUND_def
    by blast
  with CR have "GCR (rstep Sω)" unfolding GROUND_def source_step_def by auto
  with SN_Sw_step oKBi_conversion_ERw show ?thesis by auto
qed

lemma PCP_oriented_xPCP:
  assumes R_less:"ℛ ⊆ {≻}"
  shows "PCP ℛ ⊆ PCP_ext ℛ"
proof
  fix s t
  assume "(s, t) ∈ PCP ℛ"
  then obtain μ l1 r1 p l2 r2 where o: "overlap ℛ ℛ (l1, r1) p (l2, r2)"
    and "the_mgu l1 (l2 |_ p) = μ"
    and s: "s = replace_at l2 p r1 ⋅ μ"
    and t: "t = r2 ⋅ μ"
    and NF: "∀u ⊲ l1 ⋅ μ. u ∈ NF (rstep ℛ)"
    by (auto simp: PCP_def the_mgu_def split: option.splits)
  then have mgu: "mgu l1 (l2 |_ p) = Some μ"
    by (auto simp: the_mgu_def unifiers_def overlap_def simp del: mgu.simps split: option.splits dest!: mgu_complete)
  have p: "p ∈ funposs l2" using o by (auto simp: overlap_def)
  from o have rules:"∃p. p ∙ (l1, r1) ∈ ℛ" "∃p. p ∙ (l2, r2) ∈ ℛ" "vars_rule (l1, r1) ∩ vars_rule (l2, r2) = {}"
    unfolding overlap_def by auto
  { fix l r p
    assume "p ∙ (l,r) ∈ ℛ"
    with R_less have "l ≻ r" using perm_R_less perm_less by blast
    with subst have "l ⋅ μ ≻ r ⋅ μ" by auto
    from trans[OF this] have "r ⋅ μ ≻ l ⋅ μ ⟹ l ⋅ μ ≻ l ⋅ μ" by auto
    with SN_imp_acyclic[OF SN_less] trans have "(r ⋅ μ, l ⋅ μ) ∉ {≻}" unfolding acyclic_def by blast
  }
  hence orientation:"(r1 ⋅ μ, l1 ⋅ μ) ∉ {≻} ∧ (r2 ⋅ μ, l2 ⋅ μ) ∉ {≻}"
    using o[unfolded overlap_def] by auto
  from s[unfolded subst_apply_term_ctxt_apply_distrib]
  have s:"s = (ctxt_of_pos_term p (l2 ⋅ μ))⟨r1 ⋅ μ⟩"
    using ctxt_of_pos_term_subst[OF funposs_imp_poss[OF p]] by metis
  from rules p mgu orientation
  have ooverlap:"ooverlap {≻} ℛ (l1, r1) (l2, r2) p μ s t"
    unfolding ooverlap_def fst_conv snd_conv s t by simp
  from NF ordstep_rstep_conv[OF _ R_less] subst have prime: "∀u⊲l1 ⋅ μ. u ∈ NF (ordstep {≻} ℛ)"
    by fastforce
  with ooverlap show "(s, t) ∈ PCP_ext ℛ" unfolding PCP_ext_def by force
qed

lemma Ew_empty_CR_Rw_gtotal:
  assumes "Eω = {}"
  shows "CR (rstep Rω)"
proof-
  from fair PCP_oriented_xPCP[OF Rw_less] have fair:"PCP Rω ⊆ (rstep E)" unfolding assms by auto
  interpret okb_irun_nonfailing by (unfold_locales, insert fair assms, auto)
  from Ew_empty_implies_CR_Rw show ?thesis by auto
qed

end
end

locale permuted_ordered_completion_inf = ordered_completion_inf + gtotal_reduction_order less +
 fixes R E n
 assumes R0: "R 0 = {}"
    and ERn: "(E n ∪ R n) - Id ≠ {}"
    and permuted_run: "∀i < n. (E i, R i) ⊢oKBπ (E (Suc i), R (Suc i))"
    and variant_free: "∀i ≤ n. variant_free_trs (E i) ∧ variant_free_trs (R i)"
    and fair:"∀(s, t) ∈ gtotal_okb_irun_inf.PCP_ext (op ≻) (R n ∪ (E n)). (s, t) ∈ (rstep (⋃i≤n. E i))"
begin

abbreviation Sw_permuted ("Sωπ") where "Sωπ ≡ (R n) ∪ E_ord (op ≻) (E n)"

lemma (in gtotal_okb_irun_inf) pcp_litsim :
  assumes "(s, t) ∈ PCP_ext ℛ1" and "ℛ1 ≐ ℛ2"
  shows "(s,t) ∈ PCP_ext ℛ2"
proof-
  note litsim_defs = subsumable_trs.litsim_def subsumeseq_trs_def
  from assms[unfolded PCP_ext_def] obtain ρ ρ' p μ where
    o:"ooverlap {≻} ℛ1 ρ ρ' p μ s t" and prime:"∀u⊲fst ρ ⋅ μ. u ∈ NF (ordstep {≻} ℛ1)" by auto
  from this(1)[unfolded ooverlap_def] have "(∃p. p ∙ ρ ∈ ℛ1) ∧ (∃p. p ∙ ρ' ∈ ℛ1)" by auto
  with assms(2) have "∃p. p ∙ ρ ∈ ℛ2 ∧ (∃p. p ∙ ρ' ∈ ℛ2)" unfolding litsim_defs
    by (metis rule_pt.permute_plus)
  with o have ooverlap:"ooverlap {≻} ℛ2 ρ ρ' p μ s t" unfolding ooverlap_def by meson
  have "∀u⊲fst ρ ⋅ μ. u ∈ NF (ordstep {≻} ℛ2)"
  proof(rule, rule)
    fix u
    assume "fst ρ ⋅ μ ⊳ u"
    with prime have nf:"u ∈ NF (ordstep {≻} ℛ1)" by auto
    from ordstep_permute_litsim[OF _ subsumable_trs.litsim_sym[OF assms(2)], of _ _ 0]
    have "ordstep {≻} ℛ2 ⊆ ordstep {≻} ℛ1" by auto 
    from NF_anti_mono[OF this] nf show "u ∈ NF (ordstep {≻} ℛ2)" by auto
  qed
  with ooverlap show ?thesis unfolding PCP_ext_def by blast
qed

lemma correctness_oKB_permute: "GCR (rstep Sωπ) ∧ SN (rstep Sωπ) ∧ (rstep (E 0))* = (rstep (R n ∪ E n))*"
proof-
  note litsim_defs = subsumable_trs.litsim_def subsumeseq_trs_def
  from oKB_permute_run[OF permuted_run variant_free] R0 obtain E' R' where
    run:"∀i<n. (E' i, R' i) ⊢oKB (E' (Suc i), R' (Suc i))" "∀i≤n. E i ≐ E' i ∧ R i ≐ R' i" by force
  from run(2)[rule_format, of 0] R0 have "R' 0 = {}" unfolding litsim_defs by fast
  with run(1)[rule_format] oKBi_conversion R0 have conversion:"(rstep (E' 0))* = (rstep (E' n ∪ R' n))*" by (induct n) auto
  from run(2)[rule_format, of n] have ERn_litsim:"E n ≐ E' n" "R n ≐ R' n" by auto
  from ERn obtain s' t' where "(s',t') ∈ E n ∪ R n" "s' ≠ t'" by auto
  with ERn_litsim obtain s t where st:"(s,t) ∈ E' n ∪ R' n" "s ≠ t" unfolding litsim_defs
    by (metis UnE UnI1 UnI2 rule_pt.permute_prod.simps term_pt.permute_eq_iff)
  define Ei ("Ei") where "Ei ≡ λi. if i ≤ n then E' i else if (t,t) ∈ E' n then E' n else if even (i - n) then E' n - {(t,t)} else E' n ∪ {(t,t)}"
  define Ri ("Ri") where "Ri ≡ λi. if i ≤ n then R' i else R' n"
  { fix i
    have "(Ei i, Ri i) ⊢oKB (Ei (Suc i), Ri (Suc i))" proof(cases "i < n")
      case True
      with run(1)[rule_format, of i] show ?thesis unfolding Ri_def Ei_def by auto
    next
      case False
      hence "i ≥ n" by auto
      then show ?thesis proof(cases "(t,t) ∈ E' n")
        case True
        with ‹i ≥ n› have eqs:"Ei i = E' n" "Ri i = R' n" "Ei (Suc i) = E' n" "Ri (Suc i) = R' n"
          unfolding Ri_def Ei_def by auto
        from st have "(s,t) ∈ rstep (R' n ∪ (E' n))" by auto
        from oKBi.deduce[OF this this] insert_absorb[OF True] show ?thesis unfolding eqs by auto
      next
        case False
        note tt_nomem = this
      then show ?thesis proof(cases "i=n")
        case True
        with ‹i ≥ n› have eqs:"Ei i = E' n" "Ri i = R' n" "Ei (Suc i) = E' n ∪ {(t,t)}" "Ri (Suc i) = R' n"
          unfolding Ri_def Ei_def by auto
        from st have "(s,t) ∈ rstep (R' n ∪ (E' n))" by auto
        with oKBi.deduce show ?thesis unfolding eqs by auto
      next
        case False
        show ?thesis proof(cases "even (i - n)")
        case True
        hence eqs:"Ei i = E' n - {(t,t)}" "Ri i = R' n" "Ei (Suc i) = E' n ∪ {(t,t)}" "Ri (Suc i) = R' n"
          using ‹i ≥ n› False tt_nomem unfolding Ri_def Ei_def by auto
        from st have "(s,t) ∈ rstep (R' n ∪ (E' n - {(t,t)}))" by auto
        from st(2) oKBi.deduce[OF this this] show ?thesis unfolding eqs by force
      next
        case False
        hence eqs:"Ei i = E' n ∪ {(t,t)}" "Ri i = R' n" "Ei (Suc i) = E' n - {(t,t)}" "Ri (Suc i) = R' n" 
          using ‹i ≠ n› ‹i ≥ n› tt_nomem unfolding Ri_def Ei_def by auto
        with oKBi.delete[of t "Ei i"] show ?thesis unfolding eqs by auto
      qed
    qed
  qed
qed
  }
  note irun = this
  from ‹R' 0 = {}› have "Ri 0 = {}" unfolding Ri_def by auto
  interpret gtotal_okb_irun_inf Ri Ei "op ≻" by (standard, insert gtotal irun ‹Ri 0 = {}›, auto)
  have Rw:"Rω = R' n" proof
    have "⋀i. (⋂j∈{j. j≥i}. Ri j) ⊆ R' n" by (rule Inter_lower, insert Ri_def, auto)
    thus "Rω ⊆ R' n" unfolding Rω_def by blast
  next
    have "R' n ⊆ (⋂j∈{j. j≥n}. Ri j)" by (rule Inter_greatest, insert Ri_def, auto)
    thus "R' n ⊆ Rω" unfolding Rω_def by fast
  qed
  have Ew:"Eω = E' n" proof(cases "(t,t) ∈ E' n")
    case True
    hence En:"⋀i. i ≥ n ⟹ Ei i = E' n" unfolding Ei_def by auto
    show ?thesis proof
      have "⋀i. (⋂j∈{j. j≥i}. Ei j) ⊆ E' n" by (rule Inter_lower, insert En Ei_def True, auto) 
      thus "Eω ⊆ E' n" unfolding Eω_def by blast
    next
      have "E' n ⊆ (⋂j∈{j. j≥n}. Ei j)" by (rule Inter_greatest, insert En, force)
      thus "E' n ⊆ Eω" unfolding Eω_def by fast
    qed
  next
    case False
    hence En:"⋀i. i>n ⟹ Ei i = E' n - {(t,t)} ∨ Ei (Suc i) = E' n - {(t,t)}" unfolding Ei_def by auto
    with False have En':"Ei (n + 2) = E' n - {(t,t)}" unfolding Ei_def by auto
    { fix i
      from En' have 1:"i ≤ n ⟹ ∃j ≥ i. Ei j = E' n - {(t,t)}" by (meson trans_le_add1)
      from En[of i] have "i > n ⟹ ∃j ≥ i. Ei j = E' n - {(t,t)}" by (meson le_SucI order_refl)
      with 1 have "∃j ≥ i. Ei j = E' n - {(t,t)}" by (cases "i > n", auto)
    } note ex = this
    show ?thesis proof
      have "⋀i. (⋂j∈{j. j≥i}. Ei j) ⊆ E' n - {(t,t)}" by (rule Inter_lower, insert ex, force)
      thus "Eω ⊆ E' n" unfolding Eω_def by blast
    next
      from En have above_i:"⋀i. i ≥ n ⟹ E' n - {(t,t)} ⊆ Ei i" unfolding Ei_def by auto
      have "E' n - {(t,t)} ⊆ (⋂j∈{j. j≥n}. Ei j)" by (rule Inter_greatest, insert above_i, force)
      with False show "E' n ⊆ Eω" unfolding Eω_def by fast
    qed
  qed
  have fair:"∀(s, t) ∈ PCP_ext (Rω ∪ Eω). (s, t) ∈ (rstep Einf)"
  proof
    fix u v
    assume pcp:"(u, v) ∈ PCP_ext (Rω ∪ Eω)"
    from ERn_litsim have litsim:"(Rω ∪ (Eω)) ≐ (R n ∪ (E n))" using litsim_union litsim_symcl
      unfolding Ew Rw by (metis subsumable_trs.litsim_sym)
    with pcp pcp_litsim have "(u,v) ∈ PCP_ext (R n ∪ (E n))" by auto
    with fair have in_Einf:"(u,v) ∈  (rstep (⋃i≤n. E i))" by auto
    {fix k
      have "k ≤ n ⟹ (⋃i≤k. E i) ≐ (⋃i≤k. E' i)" proof(induct k, insert run(2),simp)
        case (Suc k)
        hence k:"k ≤ n" by auto
        have "(⋃i≤Suc k. E i) = ((⋃i≤k. E i) ∪ E (Suc k))" by (simp add: atMost_Suc sup.commute)
        with Suc(1)[OF k, THEN litsim_union, of "E (Suc k)" "E' (Suc k)"] run(2)[rule_format, OF Suc(2)]
        show ?case by (simp add: atMost_Suc sup.commute)
      qed
    }
    hence "(⋃i≤n. E i) ≐ (⋃i≤n. E' i)" by auto
    with in_Einf litsim_rstep_eq have step:"(u,v) ∈ (rstep (⋃i≤n. E' i))" by metis
    have "(⋃i≤n. E' i) = (⋃i≤n. Ei i)" using Ei_def by auto
    hence "(⋃i≤n. E' i) ⊆ Einf" by auto
    from rstep_mono[OF this] have "rstep (⋃i≤n. E' i) ⊆ rstep Einf" by auto
    with step show "(u,v) ∈ (rstep Einf)" by blast
qed
  from fair correctness_okb have ground_complete:
    "GCR (rstep Sw) ∧ SN (rstep Sw) ∧ (rstep (Ei 0))* = (rstep (Rω ∪ Eω))*" by auto
  from ERn_litsim ‹Rω = R' n› subsumable_trs.litsim_sym have "Rω ≐ R n" by auto
  have Eord:"E_ord (op ≻) (E n) = E_ord (op ≻) Eω" 
  proof(rule, rule)
    fix l r
    assume "(l,r) ∈ E_ord op ≻ (E n)"
    then obtain l' r' σ where lr:"l = l'⋅σ" "r = r'⋅σ" "(l', r') ∈ (E n)" "l ≻ r"
      unfolding E_ord_def mem_Collect_eq by auto
    hence "l ≠ r"  using irrefl by auto
    from litsim_symcl[OF ERn_litsim(1)] lr(3) obtain π
      where pi:"(π ∙ l',π ∙ r') ∈ (E' n)" unfolding litsim_defs rule_pt.permute_prod.simps[symmetric] by meson
    with ‹l ≠ r› have "(π ∙ l',π ∙ r') ∈ (E' n - {(t,t)})" unfolding lr by fastforce
    note facts = less_set_permute pi irrefl lr
    show "(l,r) ∈ E_ord op ≻ Eω" unfolding E_ord_def mem_Collect_eq Ew
      by (rule exI[of _ "π ∙ l'"], rule exI[of _ "π ∙ r'"], rule exI[of _ "sop (-π) ∘s σ"], insert facts, auto)
  next
    show "E_ord op ≻ Eω ⊆ E_ord op ≻ (E n)" proof
      fix l r
      assume "(l,r) ∈ E_ord op ≻ Eω"
      then obtain l' r' σ where lr:"l = l'⋅σ" "r = r'⋅σ" "(l', r') ∈ (E' n)" "l ≻ r" 
        unfolding E_ord_def mem_Collect_eq using Ew by auto
    from litsim_symcl[OF ERn_litsim(1)] lr(3) obtain π
      where pi:"(π ∙ l',π ∙ r') ∈ (E n)" unfolding litsim_defs rule_pt.permute_prod.simps[symmetric] by metis
    note facts = less_set_permute pi irrefl lr
    show "(l,r) ∈ E_ord op ≻ (E n)" unfolding E_ord_def mem_Collect_eq
      by (rule exI[of _ "π ∙ l'"], rule exI[of _ "π ∙ r'"], rule exI[of _ "sop (-π) ∘s σ"], insert facts, auto)
    qed
  qed
  note facts = this ‹Rω ≐ R n› litsim_union subsumable_trs.litsim_refl
  have Sw_eq:"Sw ≐ Sωπ" unfolding Sw_def Eord
    by (rule litsim_union, insert facts, unfold E_ord_def[of "op ≻"], fast+)
  with litsim_rstep_eq have rstep:"rstep Sw = rstep Sωπ" by auto
  from run(2)[rule_format, of 0] litsim_rstep_eq[of "E 0" "Ei 0"] have conversion0:"(rstep (E 0))* = (rstep (E' 0))*"
    unfolding Ei_def by force
  have conversion':"(rstep (R' n ∪ E' n))* = (rstep (Rω ∪ Eω))*"
    by (metis Ei_def Un_commute conversion le0 oKBi_conversion_ERw)
  from run(2)[rule_format, of n] litsim_union subsumable_trs.litsim_sym have "R' n ∪ E' n ≐ R n ∪ E n" by blast
  with conversion' litsim_rstep_eq have "(rstep (R n ∪ E n))* = (rstep (Rω ∪ Eω))*" by metis
  with conversion0 ground_complete have conversion:"(rstep (E 0))* = (rstep (R n ∪ E n))*"
    unfolding Ei_def by auto
  with ground_complete show ?thesis unfolding litsim_rstep_eq[OF Sw_eq] by auto
qed
end

locale homomorphism =
  fixes hF :: "'a ⇒ 'c"
begin
definition h where "h ≡ map_funs_term hF"
definition hR where "hR ≡ map_funs_trs hF"
definition hs where "hs ≡ map_funs_subst hF"
definition hC where "hC ≡ map_funs_ctxt hF"

lemma [simp]: "h = map_funs_term hF" unfolding h_def by auto
lemma [simp]: "hC = map_funs_ctxt hF" unfolding hC_def by auto
lemma [simp]: "hs = map_funs_subst hF" unfolding hs_def by auto
lemma [simp]: "hR = map_funs_trs hF" unfolding hR_def by auto
  
lemma ctxt_apply_imp:"t = C⟨u⟩ ⟹ (h t = (hC C)⟨h u⟩)"
using map_funs_term_ctxt_distrib by auto

lemma hR_union [simp]: "hR (R1 ∪ R2) = hR R1 ∪ hR R2"
 by (simp add: map_funs_trs_union)

lemma hR_sym [simp]: "hR (R1) = (hR R1)"
  unfolding hR_def hR_union map_funs_trs.simps map_funs_rule.simps image_Un by force
    
lemma subst_apply_h:"h (t ⋅ σ) = (h t) ⋅ (hs σ)"
  using map_funs_subst_distrib by fastforce
end
  
locale inj_homomorphism = homomorphism "hF"
  for hF :: "'a ⇒ 'c" +
  assumes injF:"inj hF"
begin
  
lemma inj: "inj h"
proof-
  { fix s t::"('a, 'b) term"
    assume st:"s ≠ t"
    hence "h s ≠ h t" proof(induct s arbitrary:t)
      case (Var x)
      with term.simps show ?case by (cases t, auto)
    next
      case (Fun f ss)
      note Funs = this
      thus ?case proof(cases t)
        case (Var x)
        then show ?thesis by auto
      next
        case (Fun g ts)
        then show ?thesis proof(cases "f=g")
          case True
          from Funs[unfolded Fun True] have "ss ≠ ts" by auto
          hence diff:"(∃i < length ss. ss ! i ≠ ts ! i) ∨ (length ss ≠ length ts)"
            using nth_equalityI by blast
          then show ?thesis proof(cases "length ss ≠ length ts")
            case True
            then show ?thesis unfolding Fun using map_eq_imp_length_eq by auto
          next
            case False
            with diff obtain i where i:"i < length ss ∧ ss ! i ≠ ts ! i" by auto
            hence "ss ! i ∈ set ss" by auto
            with Funs(1)[OF this] i have "h (ss ! i) ≠ h (ts ! i)" by auto
            with i show ?thesis unfolding Fun using map_nth_conv by fastforce
          qed
        next
          case False
          with Funs injF[unfolded inj_on_def] show ?thesis unfolding Fun by auto
        qed
      qed
        
    qed
  }
  thus ?thesis unfolding inj_on_def by auto
qed
  
definition h_inv ("h-1") where "h-1 ≡ map_funs_term (inv hF)"
abbreviation hs_inv ("hs-1") where "hs_inv σ ≡ map_funs_subst (inv hF) σ"
abbreviation hC_inv ("hC-1") where "hC_inv σ ≡ map_funs_ctxt (inv hF) σ"
  
lemma ctxt_apply_h:"t = C⟨u⟩ = (h t = (hC C)⟨h u⟩)"
proof
  assume "t = C⟨u⟩"
  thus "h t = (hC C)⟨h u⟩" using ctxt_apply_imp by auto
next
  assume "h t = (hC C)⟨h u⟩"
  thus "t = C⟨u⟩" proof(induct C arbitrary:t u)
    case Hole
    with inj[unfolded inj_on_def] show ?case by auto
  next
    case (More f bef C aft)
    note h_t = More(2)[unfolded ctxt_apply_term.simps]
    then obtain g ts where t:"t = Fun g ts" by (cases t, auto)
    from h_t injF[unfolded inj_on_def] have fg:"f = g" unfolding t by fastforce
    from h_t have lists:"map h ts = map h bef @ (hC C)⟨h u⟩ # map h aft" (is "?L = ?R")
      unfolding t by fastforce
    hence ith:"?L ! (length bef) = ?R ! (length bef)" by auto
    from lists have len:"length bef < length ts" using length_map[of h ts] by auto
    with ith have h_ti:"h (ts ! (length bef)) = (hC C)⟨h u⟩" (is "h ?ti = _")
      unfolding nth_map[OF len, of h]
      by (metis length_map nth_append_length)
    with More(1) have ti:"?ti = C⟨u⟩" by auto
    from lists[unfolded h_ti[symmetric]] have "map h ts = map h (bef @ (ts ! length bef) # aft)" by auto
    with map_injective[OF _ inj] have "ts = bef @ (?ti # aft)" by blast
    thus ?case using t ti fg by simp
  qed  
qed
  
lemma ctxt_apply_h':"h C⟨u⟩ = (hC C)⟨h u⟩" using ctxt_apply_h by auto
 
lemma inv_h_h: "h-1 (h t) = t"
  unfolding h_def h_inv_def map_funs_term_comp inv_o_cancel[OF injF] by simp

lemma map_funs_subst_compose:
fixes σ::"('d,'e) subst" and f::"'d ⇒ 'f"
shows "map_funs_subst f (σ ∘s τ) = (map_funs_subst f σ ∘s (map_funs_subst f τ))"
proof-
  have "⋀x f σ τ. map_funs_subst f (σ ∘s τ) x = ((map_funs_subst f σ) ∘s (map_funs_subst f τ)) x"
    unfolding subst_compose_def using subst_apply_h by simp
  thus ?thesis unfolding hs_def by fast
qed

lemma ctxt_image:
  assumes "h t = Ch⟨uh⟩"
  shows "Ch = hC (hC_inv Ch) ∧ uh = h (h-1 uh)"
proof-
  from assms have "h-1 (h t) = (hC-1 Ch)⟨h-1 uh⟩" unfolding h_inv_def by simp
  from this[unfolded inv_h_h] have "t = (hC-1 Ch)⟨h-1 uh⟩" by auto
  hence "h t = (hC (hC-1 Ch))⟨h (h-1 uh)⟩" by auto
  with assms have eq:"Ch⟨uh⟩ = (hC (hC-1 Ch))⟨h (h-1 uh)⟩" by auto
  hence "hC (hC-1 Ch) = Ch" by (induct Ch, auto)
  with eq[unfolded this ctxt_eq] show ?thesis by argo
qed
  
lemma subst_image:
  fixes t::"('a, 'b) term"
  assumes "h t = (h u) ⋅ σh"
  shows "∀x ∈ vars_term (h u). σh x = h (hs-1 σh x)"
proof -
  from map_funs_subst_distrib[of "inv hF" "h u" σh] have t:"t = u ⋅ (hs-1 σh)"
    using map_funs_term_ident unfolding assms[symmetric] h_inv_def[symmetric] inv_h_h by auto
  from assms have "(h u) ⋅ σh = (h u) ⋅ hs (hs-1 σh)" unfolding t subst_apply_h by auto
  thus ?thesis unfolding term_subst_eq_conv by auto
qed
  
lemma rule_h:"(s,t) ∈ RR = ((h s, h t) ∈ (hR RR))"
  unfolding hR_def map_funs_trs.simps map_funs_rule.simps
  by (smt h_def image_iff inv_h_h prod.collapse snd_conv swap_simp)

lemma rstep_h:"(s,t) ∈ rstep RR = ((h s, h t) ∈ rstep (hR RR))"
proof
  assume rstep:"(s,t) ∈ rstep RR"
  then obtain C σ l r where lr:"(l, r) ∈ RR" and sC:"s = C⟨l ⋅ σ⟩" and tC:"t = C⟨r ⋅ σ⟩" by auto
  from sC tC ctxt_apply_h have C:"h s = (hC C)⟨h (l ⋅ σ)⟩" "h t = (hC C)⟨h (r ⋅ σ)⟩" by auto
  from subst_apply_h have sub:"h (l ⋅ σ) = h l ⋅ (hs σ)" "h (r ⋅ σ) = h r ⋅ (hs σ)" by auto
  with lr[unfolded rule_h] C show "(h s, h t) ∈ rstep (hR RR)" by auto
next
  assume "(h s, h t) ∈ rstep (hR RR)"
  then obtain Ch σh lh rh where lr:"(lh, rh) ∈ hR RR" and sC:"h s = Ch⟨lh ⋅ σh⟩" and tC:"h t = Ch⟨rh ⋅ σh⟩" by auto
  from lr obtain l r where lr:"(l, r) ∈ RR" "lh = h l" "rh = h r"
    unfolding hR_def map_funs_trs.simps map_funs_rule.simps by force
  with sC[THEN ctxt_image] obtain C u where C:"Ch = hC C" and lu:"h u = h l ⋅ σh" by metis
  from lr tC[THEN ctxt_image] obtain v where rv:"h v = h r ⋅ σh" by metis
  from lu[unfolded lr, THEN subst_image] rv[THEN subst_image] obtain σ where
    "∀x ∈ vars_term (h l) ∪ vars_term (h r). (σh x = hs σ x)" by force
  hence σ:"(h l) ⋅ σh = (h l) ⋅ (hs σ)" "(h r) ⋅ σh = (h r) ⋅ (hs σ)" using term_subst_eq_conv by (fast,fast)
  from sC[unfolded C lr] have s:"s = C⟨l ⋅ σ⟩"
    unfolding σ subst_apply_h[symmetric] ctxt_apply_h[symmetric] by fast
  from tC[unfolded C lr] have t:"t = C⟨r ⋅ σ⟩"
    unfolding σ subst_apply_h[symmetric] ctxt_apply_h[symmetric] by fast
  with s lr(1) show "(s,t) ∈ rstep RR" by auto
qed

lemma subterm_h: "s ⊵ t = (h s ⊵ h t)"
proof
  assume "s ⊵ t"
  with ctxt_apply_h show "h s ⊵ h t" by fast
next
  assume "h s ⊵ h t"
  then obtain Ch where ht:"h s = Ch⟨h t⟩" by auto
  from ctxt_image[OF this] obtain C u where C:"Ch = hC C" and u:"h u = h s" by force
  from ht[unfolded C ctxt_apply_h[symmetric]] show "s ⊵ t" by auto
qed
  
lemma subterm_strict_h: "s ⊳ t = (h s ⊳ h t)"
  using subterm_h supt_supteq_set_conv by auto

lemma enc_h: "s ⊴⋅ t = (h s ⊴⋅ h t)"
proof
  assume "s ⊴⋅ t"
  with encompeq_termE obtain C σ where "t = C⟨s ⋅ σ⟩" by auto
  with ctxt_apply_h have C:"h t = (hC C)⟨h (s ⋅ σ)⟩" by auto
  from subst_apply_h have "h (s ⋅ σ) = h s ⋅ (hs σ)" by auto
  with C show "h s ⊴⋅ h t" unfolding encompeq.simps by force
next
  assume "h s ⊴⋅ h t"
  with encompeq_termE obtain Ch σh where ht:"h t = Ch⟨h s ⋅ σh⟩" by auto
  from ctxt_image[OF this] obtain C u where C:"Ch = hC C" and u:"h u = h s ⋅ σh" by metis
  with subst_image[OF u] have σ:"∀x ∈ vars_term (h s). (σh x = hs (hs-1 σh) x)" by auto
  hence σ:"(h s) ⋅ σh = (h s) ⋅ (hs (hs-1 σh))" using term_subst_eq_conv by blast
  from ht[unfolded C] show "s ⊴⋅ t"
    unfolding σ subst_apply_h[symmetric] ctxt_apply_h[symmetric] by fast
qed
  
lemma inv_hR_rule[simp]:
  assumes "(hl,hr) ∈ hR R"
  shows "(h-1 hl, h-1 hr) ∈ R ∧ h (h-1 hl) = hl ∧ h (h-1 hr) = hr"
proof-
  note hlr = assms[unfolded hR_def map_funs_trs.simps map_funs_rule.simps]
  from imageE[OF hlr, unfolded prod.inject] obtain l r where "(l, r) ∈ R" and hlr:"hl = h l" "hr = h r"
    unfolding h_def by (metis prod.collapse)
  thus ?thesis unfolding hlr inv_h_h by auto
qed
end
  
locale gtotal_okb_irun_h = gtotal_okb_irun_inf R E less + inj_homomorphism "hF"
  for R :: "nat ⇒ (('a, 'b::infinite) term × ('a, 'b) term) set" 
  and E :: "nat ⇒ (('a, 'b) term × ('a, 'b) term) set" 
  and less :: "('a, 'b) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50)
  and hF :: "'a ⇒ 'c" +
  fixes less_h :: "('c, 'b) term ⇒ ('c, 'b) term ⇒ bool" (infix "≻h" 50)
  assumes SN_less_h: "SN {(x, y). x ≻h y}"
    and ctxt_h: "s ≻h t ⟹ C⟨s⟩ ≻h C⟨t⟩"
    and subst_h: "s ≻h t ⟹ s ⋅ σ ≻h t ⋅ σ"
    and trans_h: "s ≻h t ⟹ t ≻h u ⟹ s ≻h u"
    and gtotal_h: "ground s ∧ ground t ⟶ s = t ∨ s ≻h t ∨ t ≻h s"
    and compat_h:"s' ≻ t' ⟹ (h s') ≻h (h t')"
begin
  
abbreviation less_sk_set ("{≻h}") where "{≻h} ≡ {(x, y). x ≻h y}"

abbreviation Eh where "Eh i ≡ hR (E i)"
abbreviation Rh where "Rh i ≡ hR (R i)"
abbreviation oKBi_h (infix "⊢h" 55)
  where "oKBi_h ER ER' ≡ ordered_completion.oKBi less_h ER ER'"
    
sublocale okb_h: ordered_completion less_h
  by (unfold_locales, insert SN_less_h ctxt_h subst_h trans_h gtotal_h, blast+)
  
lemma ordstep_h:
  assumes "(s,t) ∈ ordstep {≻} EE"
  shows "(h s, h t) ∈ ordstep {≻h} (hR EE)"
proof-
  from assms obtain C l r σ where lr:"(l, r) ∈ EE" and sC:"s = C⟨l ⋅ σ⟩" and tC:"t = C⟨r ⋅ σ⟩" and
    less:"l ⋅ σ ≻ r ⋅ σ" unfolding ordstep.simps by blast
  hence rule:"(h l, h r) ∈ hR EE" unfolding hR_def map_funs_trs.simps by force
  from sC tC ctxt_apply_h have C:"h s = (hC C)⟨h (l ⋅ σ)⟩" "h t = (hC C)⟨h (r ⋅ σ)⟩" by auto
  from subst_apply_h have sub:"h (l ⋅ σ) = h l ⋅ (hs σ)" "h (r ⋅ σ) = h r ⋅ (hs σ)" by auto
  with rule C compat_h[OF less] show "(h s, h t) ∈ ordstep {≻h} (hR EE)"
    unfolding ordstep.simps by auto
qed
    
lemma ostep_h:
  assumes "(s,t) ∈ ostep EE RR"
  shows "(h s, h t) ∈ okb_h.ostep (hR (EE)) (hR RR)"
  using assms rstep_h[of s t RR] ordstep_h[of s t "EE"]
  unfolding ostep_def okb_h.ostep_def hR_sym by fast

lemma encstep2_h:
  assumes "(s,t) ∈ encstep2 EE RR"
  shows "(h s, h t) ∈ okb_h.encstep2 (hR EE) (hR RR)"
using assms proof(cases)
  case (estep l r C σ)
  hence rule:"(h l, h r) ∈ (hR EE)" unfolding hR_def map_funs_trs.simps by force
  from estep ctxt_apply_h have C:"h s = (hC C)⟨h (l ⋅ σ)⟩" "h t = (hC C)⟨h (r ⋅ σ)⟩" by auto
  from subst_apply_h have sub:"h (l ⋅ σ) = h l ⋅ (hs σ)" "h (r ⋅ σ) = h r ⋅ (hs σ)" by auto
  from enc_h estep(5) have "h l ⊲⋅ h s" unfolding encomp_def by auto
  with okb_h.encstep2.estep rule sub rule C compat_h[OF estep(4)] show ?thesis by auto
next
  case (rstep l r C σ)
  hence rule:"(h l, h r) ∈ hR RR" unfolding hR_def map_funs_trs.simps by force
  from rstep ctxt_apply_h have C:"h s = (hC C)⟨h (l ⋅ σ)⟩" "h t = (hC C)⟨h (r ⋅ σ)⟩" by auto
  from subst_apply_h have sub:"h (l ⋅ σ) = h l ⋅ (hs σ)" "h (r ⋅ σ) = h r ⋅ (hs σ)" by auto
  from enc_h rstep(4) have "h l ⊲⋅ h s" unfolding encomp_def by auto
  with okb_h.encstep2.rstep rule sub rule C show ?thesis by auto
qed

lemma enc_hstep1:
  assumes "(s,t) ∈ encstep1 EE RR"
  shows "(h s, h t) ∈ okb_h.encstep1 (hR EE) (hR RR)"
using assms proof(cases)
  case (estep l r C σ)
  hence rule:"(h l, h r) ∈ (hR EE)" unfolding hR_def map_funs_trs.simps by force
  from estep ctxt_apply_h have C:"h s = (hC C)⟨h (l ⋅ σ)⟩" "h t = (hC C)⟨h (r ⋅ σ)⟩" by auto
  from subst_apply_h have sub:"h (l ⋅ σ) = h l ⋅ (hs σ)" "h (r ⋅ σ) = h r ⋅ (hs σ)" by auto
  from enc_h estep(5) have "h l ⊲⋅ h s" unfolding encomp_def by auto
  with okb_h.encstep1.estep rule sub rule C compat_h[OF estep(4)] show ?thesis by auto
next
  case rstep
  from this[unfolded rstep_h] okb_h.encstep1.rstep show ?thesis by simp
qed

lemma oKBi_h: "(Eh i, Rh i) ⊢h (Eh (Suc i), Rh (Suc i))"
proof-
  from hR_union hR_sym have s1 [simp]:"rstep (Rh i ∪ (Eh i)) = rstep (hR(R i ∪ (E i)))" by metis
  let ?h_rule = "λ(s,t).(h s, h t)"
  have h_rule:"⋀A. hR A = ?h_rule ` A" unfolding hR_def map_funs_trs.simps by force
  from inj have inj':"inj ?h_rule" unfolding inj_on_def by (simp add: inj_on_def, auto)
  have hR_diff [simp]:"⋀A B. hR (A - B) = hR A - hR B"
    using image_set_diff[OF inj'] unfolding h_rule by metis
  have h_rule [simp]:"⋀s t. hR {(s,t)} = {(h s, h t)}" unfolding hR_def map_funs_trs.simps by force
  have h_lift [simp]:"⋀s t S. (s, t) ∈ S ⟹ (h s, h t) ∈ hR S" unfolding hR_def map_funs_trs.simps by force
  from irun[of i] show ?thesis proof(cases)
    case (deduce s t u)
    with rstep_h s1 have st:"(h s, h t) ∈ rstep (Rh i ∪ (Eh i))" by blast
    from deduce rstep_h s1 have su:"(h s, h u) ∈ rstep (Rh i ∪ (Eh i))" by blast
    from deduce(1) have "Eh (Suc i) = Eh i ∪ {(h t, h u)}" unfolding hR_def map_funs_trs.simps by force
    with okb_h.oKBi.deduce[OF st su] deduce(2) show ?thesis by simp
  next
    case (orientl s t)
    with compat_h have st:"h s ≻h h t" by auto
    from orientl(4) have st':"(h s, h t) ∈ Eh i" using h_lift by force
    have E:"Eh (Suc i) = Eh i - {(h s, h t)}" unfolding orientl(1) hR_diff h_rule by auto
    have "Rh (Suc i) = Rh i ∪ {(h s, h t)}" unfolding orientl(2) hR_union h_rule by auto
    with okb_h.oKBi.orientl[OF st st'] E show ?thesis by auto
  next
    case (orientr t s)
    with compat_h have st:"h t ≻h h s" by auto
    from orientr(4) have st':"(h s, h t) ∈ Eh i" using h_lift by force
    have E:"Eh (Suc i) = Eh i - {(h s, h t)}" unfolding orientr(1) hR_diff h_rule by auto
    have "Rh (Suc i) = Rh i ∪ {(h t, h s)}" unfolding orientr(2) hR_union h_rule by auto
    with okb_h.oKBi.orientr[OF st st'] E show ?thesis by auto
  next
    case (delete s)
    have E:"Eh (Suc i) = Eh i - {(h s, h s)}" unfolding delete(1) hR_diff h_rule by force
    from delete(3) h_lift have "(h s, h s) ∈ Eh i" by force
    from okb_h.oKBi.delete[OF this] E delete(2) show ?thesis by auto
  next
    case (compose t u s)
    have R:"Rh (Suc i) = Rh i - {(h s, h t)} ∪ {(h s, h u)}"
      unfolding compose(2) hR_union hR_diff h_rule by simp
    from compose(3)[THEN ostep_h] h_rule
    have ostep:"(h t, h u) ∈ okb_h.ostep (Eh i) (Rh i - {(h s, h t)})"
      unfolding hR_diff h_rule by simp
    from okb_h.oKBi.compose[OF ostep h_lift[OF compose(4)]] R compose(1) show ?thesis by auto
  next
    case (simplifyl s u t)
    have E:"Eh (Suc i) = Eh i - {(h s, h t)} ∪ {(h u, h t)}"
      unfolding simplifyl(1) hR_union hR_diff h_rule by force
    from simplifyl(3)[THEN enc_hstep1]
      have "(h s, h u) ∈ okb_h.encstep1 (Eh i - {(h s, h t)}) (Rh i)" unfolding hR_diff h_rule by auto
    from okb_h.oKBi.simplifyl[OF this h_lift[OF simplifyl(4)]] E simplifyl(2) show ?thesis by argo
  next
    case (simplifyr t u s)
    have E:"Eh (Suc i) = Eh i - {(h s, h t)} ∪ {(h s, h u)}"
      unfolding simplifyr(1) hR_union hR_diff h_rule by force
    from simplifyr(3)[THEN enc_hstep1]
      have "(h t, h u) ∈ okb_h.encstep1 (Eh i - {(h s, h t)}) (Rh i)" unfolding hR_diff h_rule by auto
    from okb_h.oKBi.simplifyr[OF this h_lift[OF simplifyr(4)]] E simplifyr(2) show ?thesis by argo
  next
    case (collapse t u s)
    have E:"Eh (Suc i) = Eh i ∪ {(h u, h s)}" unfolding collapse(1) hR_union h_rule by force
    from hR_def have R:"Rh (Suc i) = Rh i - {(h t, h s)}" unfolding collapse(2) hR_diff h_rule by force
    from collapse(3)[THEN encstep2_h]
      have "(h t, h u) ∈ okb_h.encstep2 (Eh i) (Rh i - {(h t, h s)})" unfolding hR_diff h_rule by auto
    from okb_h.collapse[OF this h_lift[OF collapse(4)]] R E show ?thesis by auto
  qed
qed
  
sublocale irun_h: gtotal_okb_irun_inf Rh Eh less_h
proof(unfold_locales)
  show "⋀s t. ground s ⟹ ground t ⟹ s = t ∨ s ≻h t ∨ t ≻h s" using gtotal_h by auto
next
  show "⋀i. okb_h.oKBi (Eh i, Rh i) (Eh (Suc i), Rh (Suc i))" using oKBi_h by auto
qed(auto simp:R0 map_funs_trs.simps)
    
lemma hEinf [simp]:"hR E = irun_h.Einf" unfolding hR_def
    by (smt SUP_cong image_UN map_funs_trs.simps)
    
lemma hRinf [simp]:"hR R = irun_h.Rinf" unfolding hR_def
    by (smt SUP_cong image_UN map_funs_trs.simps)

lemma hEw [simp]:"hR Eω = irun_h.Eω"
proof-
  have eq:"hR Eω = (map_funs_rule hF) ` Eω" unfolding hR_def map_funs_trs.simps by force
  from inj have inj_rule:"inj (map_funs_rule hF)" unfolding h_def map_funs_rule.simps inj_on_def by auto
  { fix j ::nat
    let ?C = "Collect (op ≤ j)"
    have univ:"∀x∈Collect (op ≤ j). E x ⊆ UNIV" using subset_UNIV by blast
    from le0 have "0 ∈ {i. i ≤ j}" by simp
    with image_INT[OF inj_rule univ] have "map_funs_rule hF ` INTER ?C E = (⋂x∈?C. map_funs_rule hF ` E x)" by auto
  }
  with eq have "hR Eω = (⋃i. ⋂j∈Collect (op ≤ i). map_funs_rule hF ` E j)"
    unfolding Eω_def image_UN by presburger
  thus ?thesis unfolding irun_h.Eω_def unfolding hR_def map_funs_trs.simps by simp
qed

lemma hRw [simp]:"hR Rω = irun_h.Rω"
proof-
  have eq:"hR Rω = (map_funs_rule hF) ` Rω" unfolding hR_def map_funs_trs.simps by force
  from inj have inj_rule:"inj (map_funs_rule hF)" unfolding h_def map_funs_rule.simps inj_on_def by auto
  { fix j ::nat
    let ?C = "Collect (op ≤ j)"
    have univ:"∀x∈Collect (op ≤ j). R x ⊆ UNIV" using subset_UNIV by blast
    from le0 have "0 ∈ {i. i ≤ j}" by simp
    with image_INT[OF inj_rule univ] have "map_funs_rule hF ` INTER ?C R = (⋂x∈?C. map_funs_rule hF ` R x)" by auto
  }
  with eq have "hR Rω = (⋃i. ⋂j∈Collect (op ≤ i). map_funs_rule hF ` R j)"
    unfolding Rω_def image_UN by presburger
  thus ?thesis unfolding irun_h.Rω_def unfolding hR_def map_funs_trs.simps by simp
qed

lemma perm_h:"p ∙ (h t) = h (p ∙ t)"
  by (induct t, unfold h_def, auto)

lemma perm_inv_h:"p ∙ (h-1 t) = h-1 (p ∙ t)"
  by (induct t, unfold h_inv_def, auto)
  
lemma poss_inv_h:"poss (h-1 t) = poss t" unfolding h_inv_def
  by (induct t, unfold term.map, auto)
      
lemma funposs_inv_h:"funposs (h-1 t) = funposs t" unfolding h_inv_def
  by (induct t, unfold term.map, auto)
    
lemma pos_term_h:"p ∈ poss t ⟹ (h t |_ p) = h (t |_ p)" unfolding h_def
proof (induct t arbitrary:p)
  case (Var x)
  hence p:"p = ε" by auto
  then show ?case unfolding p term.map subt_at.simps by auto
next
  case (Fun f ts)
  show ?case proof(cases p)
    case Empty
    show ?thesis unfolding Empty term.map subt_at.simps by auto
  next
    case (PCons i ps)
    from Fun(2) have "i < length ts" "ps ∈ poss (ts ! i)" unfolding PCons poss.simps by auto
    with Fun(1)[OF _ this(2)]  show ?thesis unfolding PCons using subt_at.simps by auto
  qed
qed

lemma unify_h:
  fixes es::"(('d, 'e) term × ('d, 'e) term) list" and f::"'d ⇒ 'f"
  assumes "unify es bs = Some cs"
  defines "fh ≡ map_funs_term f"
  shows "unify (map (λ(u,v).(fh u, fh v)) es) (map (λ(x,t).(x, fh t)) bs) = Some (map (λ(x,t).(x, fh t)) cs)"
proof-
  let ?he = "(λ(u,v).(fh u, fh v))"
  let ?mhe = "map ?he"
  let ?hb = "(λ(x,t).(x, fh t))"
  let ?mhb = "map ?hb"
  { fix s t::"('d, 'e) term"
    fix us::"(('d, 'e) term × ('d, 'e) term) list"
    assume "decompose s t = Some us"
    hence "decompose (fh s) (fh t) = Some (map ?he us)" proof(induct s arbitrary:t)
      case (Var x)
      then show ?case unfolding decompose_def h_def term.map by simp
    next
      case (Fun f ss)
      from Fun(2)[unfolded decompose_def] obtain g ts where t:"t = Fun g ts" by (cases t, auto)
      from Fun(2)[unfolded decompose_def t, simplified] option.simps(3) have fg:"f = g" by metis
      from Fun(2)[unfolded decompose_def t] have d:"decompose (Fun f ss) t = zip_option ss ts"
        unfolding t fg decompose_def h_def term.map by simp
      from Fun(2)[unfolded this] zip_option_zip_conv have
        l:"length ts = length ss" "length us = length ss" "us = zip ss ts" by auto
      from this(3) zip_map_map have z:"zip (map fh ss) (map fh ts) = map ?he us" by auto
      have zo:"zip_option (map fh ss) (map fh ts) = Some (map ?he us)"
        by (rule zip_option_intros(1), unfold length_map, insert l z, auto)
      have "decompose (fh (Fun f ss)) (fh t) = zip_option (map fh ss) (map fh ts)"
        unfolding t fg decompose_def fh_def term.map by simp
      with zo show ?case by auto
    qed
  } note dec = this
  {fix x::"'e" and t E
    { fix u v
      have "map_funs_subst f (Var(x := t)) = Var(x := fh t)" unfolding fh_def by auto
      hence "⋀u. fh (u ⋅ Var(x := t)) = (fh u) ⋅ Var(x := fh t)"
        unfolding fh_def map_funs_subst_distrib by auto
      hence "(fh (u ⋅ Var(x := t)), fh (v ⋅ Var(x := t))) = ((fh u) ⋅ Var(x := fh t), fh v ⋅ Var(x := fh t))"
        by auto
    }
    hence "?mhe (subst_list (subst x t) E) = subst_list (subst x (fh t)) (?mhe E)"
      unfolding subst_list_def map_map subst_def o_def split by auto
  } note subst_list = this
show ?thesis using assms(1)
proof (induction es bs arbitrary: cs rule: unify.induct)
  case (1 bs)
  then show ?case by force
next
  case (2 f ss g ts E bs)
  then obtain us where d: "decompose (Fun f ss) (Fun g ts) = Some us"
    and [simp]: "f = g" "length ss = length ts" "us = zip ss ts"
    and u:"unify (us @ E) bs = Some cs" by (auto split: option.splits)
  from dec[OF this(1)] have
   "unify (?mhe ((Fun f ss, Fun g ts) # E)) (?mhb bs) = unify ((?mhe us) @ (?mhe E)) (?mhb bs)"
    unfolding list.map split fh_def term.map unify.simps unfolding fh_def[symmetric] by auto
  with 2(1)[OF d u] show ?case by auto
next
  case (3 x t E bs)
  show ?case
  proof (cases "t = Var x")
    assume t:"t = Var x"
    with 3(3) have u1:"unify E bs = Some cs" by simp
    have u2:"unify (?mhe ((Var x, t) # E)) (?mhb bs) = unify (?mhe E) (?mhb bs)"
      unfolding list.map split fh_def term.map t unify.simps by simp
    with 3(1)[OF t u1] show ?case by argo
  next
    assume t:"t ≠ Var x"
    let ?sub = "subst_list (subst x t) E"
    let ?sub' = "subst_list (subst x (fh t)) (?mhe E)"
    from t 3(3) have xt:"x ∉ vars_term t" unfolding unify.simps by auto
    with 3(3) t have "unify (subst_list (subst x t) E) ((x, t) # bs) = Some cs" by auto
    with 3(2)[OF t xt] have u1:"unify (?mhe ?sub) ((x, fh t) # (?mhb bs)) = Some (?mhb cs)" by auto
    from t have ht:"fh t ≠ Var x" unfolding fh_def by (cases t, auto)
    with xt have u2:"unify (?mhe ((Var x, t) # E)) (?mhb bs) = unify ?sub' ((x, fh t) # (?mhb bs))"
      unfolding list.map split fh_def term.map unify.simps vars_term_map_funs_term2 by auto
    with u1 show ?case unfolding list.map subst_list[symmetric] split by argo
  qed
next
  case (4 f ss x E bs)
    let ?t = "Fun f ss"
    let ?sub = "subst_list (subst x ?t) E"
    let ?sub' = "subst_list (subst x (fh ?t)) (?mhe E)"
    from 4(2) have xt:"x ∉ vars_term ?t" unfolding unify.simps by force
    with 4(2) have "unify (subst_list (subst x ?t) E) ((x, ?t) # bs) = Some cs" by auto
    with 4(1)[OF xt] have u1:"unify (?mhe ?sub) ((x, fh ?t) # (?mhb bs)) = Some (?mhb cs)" by auto
    have ht:"fh ?t ≠ Var x" unfolding fh_def by auto
    with xt have u2:"unify (?mhe ((?t, Var x) # E)) (?mhb bs) = unify ?sub' ((x, fh ?t) # (?mhb bs))"
      unfolding list.map split fh_def term.map unify.simps vars_term_map_funs_term2 by auto
    with u1 show ?case unfolding list.map subst_list[symmetric] split by argo
  qed
qed
  
lemma mgu_h:
  fixes s::"('a, 'b) term"
  assumes "mgu (h s) (h t) = Some μ"
  shows "mgu s t = Some ((hs_inv μ)) ∧ hs (hs_inv μ) = μ"
proof-
  let ?bmap = "map (λ(x,t). (x, h-1 t))"
  let ?bmap' = "map (λ(x,t). (x, h t))"
  from assms[unfolded mgu.simps] obtain bs where
    uh:"unify [(h s, h t)] [] = Some bs" and mu:"μ = subst_of bs" by (cases "unify [(h s, h t)] []", auto)
  from unify_h[OF this(1), of "inv hF"] have u:"unify [(s, t)] [] = Some (?bmap bs)"
    unfolding list.map split h_def map_funs_term_comp o_def inv_f_f[OF injF] term.map_ident h_inv_def
    by auto
  from unify_h[OF this, of hF] uh have bs:"bs = ?bmap' (?bmap bs)"
    unfolding list.map split h_def[symmetric] by auto
  have x:"⋀ x f t. subst x (map_funs_term f t) = map_funs_subst f (subst x t)" unfolding subst_def by auto
  { fix bs::"('e × ('d, 'e) term) list"
    fix f::"'d ⇒ 'f"
    have "subst_of (map (λ(x, t). (x, map_funs_term f t)) bs) = map_funs_subst f (subst_of bs)"
    proof-
      let ?m = "λ(x, t). (x, map_funs_term f t)"
      show ?thesis proof(induct bs, simp)
        case (Cons xt bs)
        then obtain x t where xt:"xt = (x,t)" by fastforce
        have "subst_of (map ?m (xt # bs)) = subst_of (map ?m bs) ∘s (subst x (map_funs_term f t))"
          unfolding xt list.map split by auto
        also from Cons x[of x f] have "… = (map_funs_subst f (subst_of bs)) ∘s (map_funs_subst f (subst x t))" by simp
        also from map_funs_subst_compose[of f] have "… = map_funs_subst f (subst_of bs ∘s (subst x t))" by metis
        finally show ?case unfolding xt by auto
      qed
    qed
  }
  hence "subst_of (map (λ(x, t). (x, h-1 t)) bs) = hs-1 (subst_of bs)"
    unfolding h_inv_def by auto
  hence mgu:"mgu s t = Some ((hs_inv μ))" unfolding mgu.simps u mu by simp
  { fix y
    from bs have "map_funs_subst hF (map_funs_subst (inv hF) (subst_of bs)) y = (subst_of bs) y" (is "?l y = ?r y")
    proof(induct bs, simp)
      case (Cons xt bs)
      then obtain x t where xt:"xt = (x,t)" by fastforce
      let ?h = "map_funs_subst hF"
      let ?hi = "map_funs_subst (inv hF)"
      let ?xt = "subst x t"
      from Cons(2)[unfolded xt list.map split] have t:"h (h-1 t) = t" by auto
      have "?h (?hi (subst_of (xt # bs))) y = ?h (?hi (subst_of bs ∘s ?xt)) y" (is "?t = _ ")
        unfolding xt by auto
      with map_funs_subst_compose[of "inv hF"] have "?t = ?h (?hi (subst_of bs) ∘s (?hi ?xt)) y" by metis
      with map_funs_subst_compose[of hF] have "?t = (?h (?hi (subst_of bs)) ∘s (?h (?hi ?xt))) y" by metis
      hence "?t = (?h (?hi (subst_of bs)) y) ⋅ (?h (?hi ?xt))" using subst_compose_def by metis
      with Cons have "?t = (subst_of bs y) ⋅ (?h (?hi ?xt))" by auto
      with t have "?t = (subst_of bs y) ⋅ (?h (?hi ?xt))" unfolding subst_of_def by auto
      with x[of x "inv hF"] have "?t = (subst_of bs y) ⋅ (?h (subst x (h-1 t)))"
        unfolding subst_of_def h_inv_def by auto
      with x[of x hF] have "?t = (subst_of bs y) ⋅ (subst x (h (h-1 t)))" unfolding subst_of_def by auto
      with t have "?t = (subst_of bs y) ⋅ ?xt" by auto
      hence "?t = ((subst_of bs) ∘s ?xt) y" unfolding subst_compose_def by metis
      then show ?case unfolding xt by auto
    qed
  }
  hence "hs (hs_inv μ) = μ" unfolding hs_def mu by auto
  with mgu show ?thesis by auto
qed
   
lemma PCP_h:
  assumes "(hs,ht) ∈ irun_h.PCP_ext (hR Rω ∪ ((hR Eω)))"
  shows "(h-1 hs, h-1 ht) ∈ PCP_ext (Rω ∪ Eω) ∧ hs = h (h-1 hs) ∧ ht = h (h-1 ht)"
proof-
  let ?s = "h-1 hs"
  let ?t = "h-1 ht"
  from assms[unfolded irun_h.PCP_ext_def] obtain hrl hrl' pos μ where
    ooverlap:"ooverlap {≻h} (hR Rω ∪ (hR Eω)) hrl hrl' pos μ (hs) (ht)" and
    nf:"∀hs⊲fst hrl ⋅ μ. hs ∈ NF (ordstep {≻h} (hR Rω ∪ (hR Eω)))" by blast
  (* ooverlap *)
  obtain hl hr hl' hr' where hrl:"hrl = (hl, hr)" "hrl' = (hl', hr')" by force
  { fix p hl hr
    assume p:"∃p. p ∙ (hl, hr) ∈ hR Rω ∪ (hR Eω)"
    then obtain p where "(p ∙ hl, p ∙ hr) ∈ hR Rω ∪ (hR Eω)"
      unfolding rule_pt.permute_prod_eqvt by auto
    hence p:"(p ∙ hl, p ∙ hr) ∈ hR (Rω ∪ Eω)" using hR_union hR_sym by blast
    let ?l = "p ∙ hl" and ?r = "p ∙ hr"
    from inv_hR_rule[OF p] have x:"(h-1 ?l, h-1 ?r) ∈ (Rω ∪ Eω)" "h (h-1 ?l) = ?l ∧ h (h-1 ?r) = ?r"
      unfolding perm_inv_h by auto
    from this(2) have "h (h-1 hl) = hl ∧ h (h-1 hr) = hr"
      unfolding perm_inv_h[symmetric] perm_h[symmetric] term_pt.permute_eq_iff by auto
    with x have "(∃p. p ∙ (h-1 hl, h-1 hr) ∈ Rω ∪ (Eω)) ∧ h (h-1 hl) = hl ∧ h (h-1 hr) = hr"
      unfolding rule_pt.permute_prod_eqvt perm_inv_h by meson
  } note perm_rule = this
  note ooverlap = ooverlap[unfolded ooverlap_def hrl snd_conv fst_conv]
  hence
    p:"∃p. p ∙ (hl, hr) ∈ hR Rω ∪ (hR Eω)" and
    p':"∃p'. p' ∙ (hl', hr') ∈ hR Rω ∪ (hR Eω)" and
    vars:"vars_rule (hl, hr) ∩ vars_rule (hl', hr') = {}" and
    pos:"pos ∈ funposs hl'" and
    mgu:"mgu hl (hl' |_ pos) = Some μ" and
    less:"(hr ⋅ μ, hl ⋅ μ) ∉ {≻h}" and less':"(hr' ⋅ μ, hl' ⋅ μ) ∉ {≻h}" and
    cs:"hs = (ctxt_of_pos_term pos (hl' ⋅ μ))⟨hr ⋅ μ⟩" and ht_mu:"ht = hr' ⋅ μ" by auto
  note p = perm_rule[OF p] perm_rule[OF p']
  from p have 1:"∃p. p ∙ (h-1 hl, h-1 hr) ∈ Rω ∪ Eω" "∃p. p ∙ (h-1 hl', h-1 hr') ∈ Rω ∪ Eω" by auto
  from p have inv:"h (h-1 hl) = hl" "h (h-1 hr) = hr" "h (h-1 hl') = hl'" "h (h-1 hr') = hr'" by auto
  have vt:"⋀t. vars_term (h-1 t) = vars_term t" by (unfold h_inv_def, auto)
  from vars have 2:"vars_rule (h-1 hl, h-1 hr) ∩ vars_rule (h-1 hl', h-1  hr') = {}"
    unfolding vars_rule_def fst_conv snd_conv vt by auto
  from pos have 3:"pos ∈ funposs (h-1 hl')" using funposs_inv_h by auto
  from mgu inv(3) have "mgu hl (h (h-1 hl') |_ pos ) = Some μ" by argo
  with pos_term_h[OF funposs_imp_poss[OF 3]] have "mgu (h (h-1 hl)) (h (h-1 hl' |_ pos )) = Some μ"
    using inv(1) by auto
  from mgu_h[OF this] have 4:"mgu (h-1 hl) (h-1 hl' |_ pos) = Some (hs_inv μ)" and μ:"μ = hs (hs_inv μ)"
    unfolding hs_def by auto
  let  = "hs_inv μ"
  from μ less inv have "(h (h-1 hr) ⋅ (hs ?μ), h (h-1 hl) ⋅ (hs ?μ)) ∉ {≻h}" by auto
  with subst_apply_h have "(h (h-1 hr ⋅ ?μ), h (h-1 hl ⋅ ?μ)) ∉ {≻h}" by auto
  with compat_h have 5:"((h-1 hr) ⋅ ?μ, (h-1 hl) ⋅ ?μ) ∉ {≻}" by fast
  from μ less' inv have "(h (h-1 hr') ⋅ (hs ?μ), h (h-1 hl') ⋅ (hs ?μ)) ∉ {≻h}" by auto
  with subst_apply_h have "(h (h-1 hr' ⋅ ?μ), h (h-1 hl' ⋅ ?μ)) ∉ {≻h}" by auto
  with compat_h have 6:"((h-1 hr') ⋅ ?μ, (h-1 hl') ⋅ ?μ) ∉ {≻}" by fast
  let ?C = "ctxt_of_pos_term pos"
  from cs inv μ have "hs = (ctxt_of_pos_term pos (h (h-1 hl') ⋅ hs ?μ))⟨h (h-1 hr) ⋅ hs ?μ⟩" by auto
  with subst_apply_h have hs:"hs = (?C (h (h-1 hl' ⋅ ?μ)))⟨h (h-1 hr ⋅ ?μ)⟩" by auto
      
  from ctxt_of_pos_term_map_funs_term_conv funposs_imp_poss[OF 3]
    have "?C (h (h-1 hl' ⋅ ?μ)) = hC (?C (h-1 hl' ⋅ ?μ))"
    unfolding h_def hC_def by auto
  with ctxt_apply_h have hs:"hs = h ((?C (h-1 hl' ⋅ ?μ))⟨h-1 hr ⋅ ?μ⟩)" unfolding hs by auto
  note inv_subst_apply_h = map_funs_subst_distrib[of "inv hF", unfolded h_inv_def[symmetric]]
  note inv_ctxt_apply_h = map_funs_term_ctxt_distrib[of "inv hF", unfolded h_inv_def[symmetric]]
  from funposs_imp_poss[OF 3] have pos':"pos ∈ poss (hl' ⋅ μ)" unfolding poss_inv_h by auto
  from hs have "hs = h (h-1 (?C (hl' ⋅ μ))⟨hr ⋅ μ⟩)"
    unfolding inv_subst_apply_h[symmetric] unfolding h_inv_def
    ctxt_of_pos_term_map_funs_term_conv[OF pos', symmetric]
    inv_ctxt_apply_h[symmetric, unfolded h_inv_def] by simp
  with cs have hs':"h (h-1 hs) = hs" by auto
  have 7:"h-1 hs = (ctxt_of_pos_term pos (h-1 hl' ⋅ hs-1 μ))⟨h-1 hr ⋅ hs-1 μ⟩" unfolding hs inv_h_h by auto
  
  from ht_mu μ inv(4) have "h (h-1 ht) = h (h-1 (h (h-1 hr') ⋅ hs (hs_inv μ)))" by auto
  with subst_apply_h have "h (h-1 ht) = h (h-1 (h (h-1 hr' ⋅ hs_inv μ)))" by auto
  hence "h (h-1 ht) = h (h-1 hr') ⋅ (hs (hs_inv μ))" unfolding inv_h_h unfolding subst_apply_h by auto
  with inv(4) μ ht_mu have ht':"h (h-1 ht) = ht" by auto
  hence "?t = h-1 (hr' ⋅ μ)" unfolding ht_mu using inv_h_h by auto
  hence 8:"?t = h-1 hr' ⋅ ?μ" unfolding inv_subst_apply_h by blast   
  from 1 2 3 4 5 6 7 8 have o:"ooverlap {≻} (Rω ∪ (Eω)) (h-1 hl, h-1 hr) (h-1 hl', h-1 hr') pos ?μ ?s ?t"
    unfolding ooverlap_def fst_conv snd_conv by argo
  have nf:"∀u ⊲ (h-1 hl) ⋅ ?μ. u ∈ NF (ordstep {≻} (Rω ∪ Eω))"
  proof(rule, rule)
    fix u
    assume "(h-1 hl) ⋅ ?μ ⊳ u"
    with subterm_strict_h have "h ((h-1 hl) ⋅ ?μ) ⊳ h u" by metis
    with inv(1) μ subst_apply_h have a:"hl ⋅ μ ⊳ h u" by auto
    { fix v
      assume "(u,v) ∈ ordstep {≻} (Rω ∪ Eω)"
      hence False proof(cases)
        case (1 l r C σ)
        from 1(1) have "(h l, h r) ∈ hR (Rω ∪ Eω)"
          unfolding hR_def h_def map_funs_trs.simps map_funs_rule.simps by force
        with hR_union hR_sym have hlr:"(h l, h r) ∈ hR Rω ∪ (hR Eω)" by blast
        from 1(2) subst_apply_h ctxt_apply_h have 2:"h u = (hC C)⟨h l ⋅ hs σ⟩" by auto
        from 1(3) subst_apply_h ctxt_apply_h have 3:"h v = (hC C)⟨h r ⋅ hs σ⟩" by auto
        from 1(4) compat_h subst_apply_h have 4:"(h l ⋅ hs σ, h r ⋅ hs σ) ∈ {≻h}" by force
        from hlr 2 3 4 have "(h u,h v) ∈ ordstep {≻h} (hR Rω ∪ (hR Eω))"
          unfolding ordstep.simps by blast
        with nf a show False unfolding hrl fst_conv by auto
      qed
    }
    thus "u ∈ NF (ordstep {≻} (Rω ∪ Eω))" by fast
  qed
  from o nf have "(?s,?t) ∈ PCP_ext (Rω ∪ Eω)" unfolding PCP_ext_def by force
  with hs' ht' show ?thesis by auto
qed
  
end


datatype ('a,'b) f_ext = FOrig 'a | FFresh 'b
  
text ‹Towards completeness of ordered completion.›
locale gtotal_okb_irun2 = gtotal_okb_irun_inf R E less
  for R :: "nat ⇒ (('a, 'b::infinite) term × ('a, 'b) term) set" 
  and E :: "nat ⇒ (('a, 'b) term × ('a, 'b) term) set" 
  and less :: "('a, 'b) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50)
  +
  fixes "c"::'a
  assumes c_min:"∀(t::('a, 'b) term). t ≽ (Fun c [])"
begin

text ‹Replace variables in a term by fresh function symbols.›
  
sublocale sk: inj_homomorphism FOrig
  by (unfold_locales, insert inj_on_def, auto)

definition sk" where sk ≡ (λx. Fun (FFresh x) [])"
    
definition sk :: "('a, 'b) term ⇒ (('a, 'b)f_ext, 'b) term"
  where "sk t = (sk.h t) ⋅ σsk"
  
lemma ground_sk [simp]: "ground (sk t)"
  by (induct t, insert sk_def sk.h_def, unfold σsk_def, auto)
  
definition orig_sig :: "(('a,'b) f_ext, 'b) term ⇒ bool"
  where "orig_sig t = (∀x. FFresh x ∉ funs_term t)"

definition orig_sigC :: "(('a,'b) f_ext, 'b) ctxt ⇒ bool"
   where "orig_sigC t = (∀x. FFresh x ∉ funs_ctxt t)"
  
lemma inj_on_sigma_sk: "inj_on (λ t. t ⋅ σsk) {t|t. orig_sig t}"
proof-
  { fix s::"(('a,'b) f_ext, 'b) term"
    fix t
    assume st:"s ≠ t" and s_dom:"∀x. FFresh x ∉ funs_term s" and t_dom:"∀x. FFresh x ∉ funs_term t"
    hence "s ⋅ σsk ≠ t ⋅ σsk" proof(induct s arbitrary:t)
      case (Var x)
      thus ?case using subst_apply_term.simps unfolding σsk_def by (cases t, auto)
    next
      case (Fun f ss)
      note Funs = this
      thus ?case proof(cases t)
        case (Var x)
        from Funs show ?thesis unfolding Var subst_apply_term.simps σsk_def by auto
      next
        case (Fun g ts)
        then show ?thesis proof(cases "f=g")
          case True
          from Funs[unfolded Fun True] have "ss ≠ ts" by auto
          hence diff:"(∃i < length ss. ss ! i ≠ ts ! i) ∨ (length ss ≠ length ts)"
            using nth_equalityI by blast
          then show ?thesis proof(cases "length ss ≠ length ts")
            case True
             then show ?thesis unfolding Fun subst_apply_term.simps using map_eq_imp_length_eq by blast
          next
            case False
            with diff obtain i where i:"i < length ss" "ss ! i ≠ ts ! i" by auto
            from i have in_set:"ss ! i ∈ set ss" by auto
            from i False have in_set2:"ts ! i ∈ set ts" by auto
            with in_set Funs(3) Funs(4)
              have all:"∀x. FFresh x ∉ funs_term (ss ! i)" " ∀x. FFresh x ∉ funs_term (ts ! i)"
              unfolding Fun by auto
            from Funs(1)[OF in_set i(2) all]
             i show ?thesis unfolding Fun using map_nth_conv by fastforce
          qed
        next
          case False
          with Funs show ?thesis unfolding Fun by auto
        qed
      qed
    qed
  }
  thus ?thesis unfolding inj_on_def orig_sig_def by auto
qed
    
lemma orig_sig_h:"orig_sig (sk.h t)"
  unfolding orig_sig_def sk.h_def using funs_term_map_funs_term[of FOrig t]
  by blast

lemma sk_Fun: "sk (Fun f ts) = Fun (FOrig f) (map sk ts)" unfolding sk_def by auto

lemma sk_ctxt_exists:
  fixes t::"('a, 'b) term"
  fixes us::"(('a, 'b) f_ext, 'b) term"
  assumes "sk t = Cs⟨us⟩"
  shows "∃ C u. Cs = (sk.hC C) ⋅c σsk ∧ us = sk u"
  using assms
proof (induct Cs arbitrary: t)
  case Hole
  show ?case unfolding orig_sigC_def by (rule exI[of _ Hole], rule exI[of _ t], insert Hole, auto)
next
  case (More g bef C aft)
  from More(2) obtain f ts where t: "t = Fun f ts" "g = FOrig f" unfolding sk_def σsk_def by (cases t, auto)
  from More(2) sk_Fun have args:"bef @ C⟨us⟩ # aft = map sk ts" unfolding t ctxt_apply_term.simps by auto
  define i where "i ≡ length bef" let ?ti = "ts ! i"
  from arg_cong[OF args, of length] have l:"i < length ts" unfolding i_def by force
  from arg_cong[OF args, of "λl. l ! i"] i_def have "sk ?ti = C⟨us⟩" unfolding nth_map[OF l] by simp
  from More(1)[OF this] obtain D v where D:"C = sk.hC D ⋅c σsk" "us = sk v" by auto
  let ?C = "More f (take i ts) D (drop (Suc i) ts)"
  from args l have b:"bef = map sk (take i ts)" by (metis append_eq_conv_conj take_map i_def)
  from l i_def arg_cong[OF args, of "λl. drop (Suc i) l"]
    have a:"aft = map sk (drop (Suc i) ts)" unfolding drop_map by auto
  have *:"More g bef C aft = (sk.hC ?C) ⋅c σsk"
    unfolding b a sk.hC_def ctxt.map t(2) subst_apply_ctxt.simps map_map sk_def D by auto
  show ?case by (rule exI[of _ ?C], rule exI[of _ v], insert * D, auto)
qed

fun deskolemize :: "(('a, 'b) f_ext, 'b) term ⇒ ('a, 'b) term"  where
    "deskolemize (Var x) = Var x"
  | "deskolemize (Fun (FOrig f) ts) = Fun f (map deskolemize ts)"
  | "deskolemize (Fun (FFresh f) _) = Var f"

lemma deskolemize_sk[simp]:"deskolemize (sk t) = t"
proof(induct t)
case (Var x)
  then show ?case unfolding sk_def σsk_def by auto
next
  case (Fun f ts)
  hence "map (deskolemize ∘ sk) ts = ts" by (simp add: map_idI)
  thus ?case unfolding sk_def sk.h_def term.map subst_apply_term.simps map_map deskolemize.simps(2) o_def by simp
qed

lemma deskolemize_orig_sig: "orig_sig t ⟹ deskolemize t = sk.h_inv t"
proof(induct t)
  case (Var x)
  then show ?case unfolding deskolemize.simps sk.h_inv_def term.map by simp
next
  case (Fun f' ts)
  from Fun(2) obtain f where f:"f' = FOrig f" unfolding orig_sig_def by (cases f', auto)
  from Fun have "map deskolemize ts = map sk.h_inv ts" unfolding orig_sig_def by (simp add: map_idI)
  thus ?case unfolding f deskolemize.simps sk.h_inv_def term.map by (simp add: sk.injF)
qed

lemma deskolemize_subst: "orig_sig t ⟹ deskolemize (t ⋅ σ) = (deskolemize t) ⋅ (deskolemize ∘ σ)"
proof(induct t)
  case (Var x)
  then show ?case by simp
next
  case (Fun f' ts)
  then obtain f where f:"f' = FOrig f" unfolding orig_sig_def by (cases f', auto)
  from Fun have "⋀ti. ti ∈ set ts ⟹ orig_sig ti" using orig_sig_def by auto
  with Fun have "map (λti. deskolemize (ti ⋅ σ)) ts = map (λti. (deskolemize ti) ⋅ (deskolemize ∘ σ)) ts"
    by (meson map_eq_conv)
  then show ?case unfolding f subst_apply_term.simps deskolemize.simps by auto
qed

lemma sk_deskolemize_subst_vars:
  assumes "sk s = l ⋅ τ"
  shows "∀x ∈ vars_term l. τ x = sk (deskolemize (τ x))"
  using assms
proof(induct l arbitrary:s)
  case (Var x)
  hence tx:"τ x = sk s" by auto
  have "sk (deskolemize (τ x)) = τ x" using deskolemize_sk unfolding tx by metis
  thus ?case by auto
next
  case (Fun f ls)
  note Funl = this
  hence eq:"sk s = Fun f (map (λti. ti ⋅ τ) ls)" by auto
  show ?case proof(cases s)
    case (Var x)
    from eq[unfolded Var] have "f = FFresh x" unfolding sk_def σsk_def by auto
    from eq[unfolded this] have "ls = []" unfolding sk_def Var σsk_def by auto
    with Fun show ?thesis by auto
  next
    case (Fun g ss)
    from eq[unfolded Fun subst_apply_term.simps] have
      fg:"f = FOrig g" "map sk ss = map (λti. ti ⋅ τ) ls" unfolding sk_Fun by auto
    hence len:"length ss = length (map (λti. ti ⋅ τ) ls)" using length_map by metis
    { fix i
      assume i:"i < length ls"
      from fg(2) map_nth_eq_conv[OF len] i have "sk (ss ! i) = (ls ! i) ⋅ τ" by force
      from Funl(1)[OF _ this] i have "∀x∈vars_term (ls ! i). τ x = sk (deskolemize (τ x))" by auto
    }
    thus ?thesis using var_imp_var_of_arg[of _ f ls] by force
  qed
qed

lemma deskolemize_s:
  assumes "sk t  = sk.h (u :: ('a, 'b) term) ⋅ σ"
  shows "t = u ⋅ (deskolemize ∘ σ)"
proof-
  have xs:"⋀x. (deskolemize ∘ σsk) x = Var x" using deskolemize_sk[of "Var _"] unfolding sk_def by simp
  with term_subst_eq_conv[of t Var "deskolemize ∘ σsk"] have t:"t ⋅ (deskolemize ∘ σsk) = t" by simp
  from assms[unfolded sk_def] have "deskolemize (sk.h t ⋅ σsk) = deskolemize (sk.h u ⋅ σ)" by auto
  thus ?thesis
    using deskolemize_subst[OF orig_sig_h] deskolemize_orig_sig[OF orig_sig_h] sk.inv_h_h t by metis
qed

lemma rstep_imp_deskolemize_rstep:
  assumes "(s, t) ∈ rstep (sk.hR RR)"
  shows "(deskolemize s, deskolemize t) ∈ (rstep RR)="
proof-
  from assms obtain l r σ C where lr:"(l,r) ∈ sk.hR RR" and st:"s = C⟨l⋅σ⟩" "t = C⟨r⋅σ⟩" by auto
  show ?thesis unfolding st proof(induct C)
    case Hole
    from orig_sig_h have sig:"orig_sig l" "orig_sig r" unfolding orig_sig_def by (metis lr sk.inv_hR_rule)+
    let ?l = "sk.h_inv l" and ?r = "sk.h_inv r"
    from sig deskolemize_orig_sig deskolemize_subst have d:
      "deskolemize (l ⋅ σ) = ?l ⋅ (deskolemize ∘ σ)" "deskolemize (r ⋅ σ) = ?r ⋅ (deskolemize ∘ σ)" by auto
    from sig lr deskolemize_orig_sig sk.inv_hR_rule have "(?l,?r) ∈ RR" by auto
    thus ?case unfolding ctxt_apply_term.simps d by auto
  next
    case (More f' bef C aft)
    then show ?case proof(cases f')
      case (FOrig f)
      let ?C = "More f (map deskolemize bef) Hole (map deskolemize aft)"
      from More show ?thesis unfolding FOrig ctxt_apply_term.simps deskolemize.simps map_append list.map
        using rstep.ctxt[of _ _ _ ?C] by auto
    next
      case (FFresh f)
      from More show ?thesis unfolding lr FFresh ctxt_apply_term.simps deskolemize.simps by auto
    qed
  qed
qed

lemma rstep_imp_deskolemize_rsteps:
  assumes "(s, t) ∈ (rstep (sk.hR RR))*"
  shows "(deskolemize s, deskolemize t) ∈ (rstep RR)*"
  using assms
proof(induct, simp)
  case (step t u)
  with rstep_imp_deskolemize_rstep have "(deskolemize t, deskolemize u) ∈ (rstep RR)=" by auto
  with step show ?case by auto
qed

lemma rstep_imp_rstep_sk:"(s,t) ∈ rstep RR ⟹ (sk s, sk t) ∈ rstep (sk.hR RR)"
  using sk.rstep_h[of s t RR] rstep.subst unfolding sk_def by auto

lemma rstep_sk_imp_rstep:"(sk s, sk t) ∈ rstep (sk.hR RR) ⟹ (s,t) ∈ (rstep RR)="
  using rstep_imp_deskolemize_rstep deskolemize_sk by metis

lemma rstep_step_sk:"s ≠ t ⟹ (s,t) ∈ rstep RR = ((sk s, sk t) ∈ rstep (sk.hR RR))"
  using rstep_imp_rstep_sk[of s t] rstep_sk_imp_rstep[of s t] by blast

lemma sk_C:
  assumes "sk s = C⟨t⟩"
  shows "hole_pos C ∈ poss s ∧ sk (s |_ (hole_pos C)) = t"
  using assms
proof(induct C arbitrary: s t)
  case Hole
  then show ?case by auto
next
  case (More f bef ctx aft)
  then obtain g ts where f:"s = Fun g ts" "f = FOrig g" unfolding sk_def σsk_def by (cases s, auto)
  let ?ti = "ts ! (length bef)"
  from More[unfolded f] sk_Fun have ts:"bef @ ctx⟨t⟩ # aft = map sk ts" by auto
  from ts have "length (bef @ ctx⟨t⟩ # aft) = length ts" by auto
  hence l:"length bef < length ts" unfolding length_append length_Cons by linarith
  with ts have "sk ?ti = ctx⟨t⟩" by (metis nth_append_length nth_map)
  with More(1)[OF this] have ti:"hole_pos ctx ∈ poss ?ti" "sk (?ti |_ hole_pos ctx) = t" by auto
  with l have p:"hole_pos (More f bef ctx aft) ∈ poss s" unfolding f poss.simps(2) by simp
  with ti(2) show ?case unfolding f subt_at.simps by simp
qed

lemma sk_inj:"sk s = sk t ⟹ s = t"
  using sk.inj inj_on_sigma_sk orig_sig_h unfolding sk_def inj_on_def by blast

definition sk_redord :: "((('a,'b) f_ext, 'b) term ⇒ (('a,'b) f_ext, 'b) term ⇒ bool) ⇒ bool"
  where "sk_redord Rel ≡
  ((∀ s t. s ≻ t ⟶ Rel (sk.h s) (sk.h t)) ∧
   (∀ s t. ground s ∧ ground t ⟶ s = t ∨ Rel s t ∨ Rel t s) ∧
   SN {(x, y). Rel x y} ∧
   (∀ s t C. Rel s t ⟶ Rel C⟨s⟩ C⟨t⟩) ∧
   (∀ s t σ. Rel s t ⟶ Rel (s ⋅ σ) (t ⋅ σ)) ∧
   (∀ s t u. Rel s t ⟶ Rel t u ⟶ Rel s u) ∧
   (∀ t. ground t ⟶ Rel== t (sk.h (Fun c []))))"
  
text ‹Assume a ground total extension of ≻ on image of homomorphism terms .›
  
definition preco :: "(('a × nat) × ('a × nat)) set" where
  "preco ≡ (SOME Rel. well_order_on (UNIV - {(c,0)}) Rel) ∪ {((c,0),fn) |fn. True}"
  
definition precf :: "(('b × nat) × ('b × nat)) set" where
  "precf ≡ (SOME Rel. well_order Rel)"
  
definition prec_set where
 "prec_set ≡ {((FOrig f, n), (FFresh g, m)) |f g n m. True} ∪
             {((FFresh f, n), (FFresh g, m)) |f g n m. ((f,n),(g,m)) ∈ precf - Id} ∪
             {((FOrig f, n), (FOrig g, m)) |f g n m. ((f,n),(g,m)) ∈ preco - Id}"

definition precsk where "precsk gm fn ≡ (fn, gm) ∈ prec_set"

lemma well_order_precf:"well_order precf"
proof-
  from Zorn.well_order_on obtain r ::"(('b × nat) × ('b × nat)) set" where r:"well_order r" by blast
  have "∃Rel :: (('b × nat) × ('b × nat)) set. well_order Rel" unfolding precf_def
    by (rule exI[of _r], insert r, auto)
  from someI_ex[OF this] show ?thesis using precf_def by auto
qed

lemma well_order_preco:"well_order preco"
proof-  
  define R ::"(('a × nat) × ('a × nat)) set" where "R ≡ SOME Rel. well_order_on (UNIV - {(c,0)}) Rel"
  have p:"preco = R ∪ {((c,0),fn) |fn. True}" unfolding preco_def R_def by auto
      
  from Zorn.well_order_on obtain r ::"(('a × nat) × ('a × nat)) set" where r:"well_order_on (UNIV - {(c,0)}) r" by blast
  have ex:"∃Rel ::(('a × nat) × ('a × nat)) set. well_order_on (UNIV - {(c,0)}) Rel"
    by (rule exI[of _ r], insert r, auto)
  from someI_ex[OF ex] have x:"well_order_on (UNIV - {(c,0)}) R" unfolding R_def by auto
  hence wf1:"wf (R - Id)" (is "wf ?R1") unfolding well_order_on_def by auto 
  have wf0:"wf ({((c,0),fn) |fn. True} - Id)" (is "wf ?R0") unfolding wf_def
    by (smt CollectD DiffD1 DiffD2 Pair_inject pair_in_Id_conv)
  from well_order_on_Field[OF x, unfolded Field_def]
    have "Domain ({((c,0),fn) |fn. True} - Id) ∩ Range (R - Id) = {}" by auto
  from wf_Un[OF wf0 wf1 this] have "wf (({((c,0),fn) |fn. True} - Id) ∪ (R - Id))" by metis
  hence wf:"wf (preco - Id)" unfolding p Un_Diff using Un_commute by metis
  
  let ?U = "UNIV - {(c,0)}"
  from well_order_on_Field[OF x, unfolded Field_def] have c:"(c,0) ∉ Range R" unfolding R_def by auto
  from x have "linear_order_on?U R" unfolding well_order_on_def by auto
  hence po:"partial_order_on ?U R" and total:"Relation.total_on ?U R"
    unfolding linear_order_on_def by auto
  from po[unfolded partial_order_on_def] have pre:"preorder_on ?U R" and antisym:"antisym R" by auto
  hence refl:"refl_on ?U R" and trans:"trans R" unfolding preorder_on_def by auto
  have total:"total preco"
  proof
    fix f g :: "('a × nat)"
    assume "f ≠ g"
    with total[unfolded Relation.total_on_def] show "(f, g) ∈ preco ∨ (g, f) ∈ preco"
      unfolding p by (cases "f = (c,0)", auto)
  qed
  from refl have refl:"refl preco" unfolding p refl_on_def by auto
  from trans c have trans:"trans preco" unfolding p trans_def by blast
  from refl trans have pre:"preorder_on UNIV preco" unfolding preorder_on_def by auto
  have antisym:"antisym preco" proof
    fix f g :: "('a × nat)"
    assume "(f, g) ∈ preco" and "(g, f) ∈ preco"
    with antisym[unfolded antisym_def] c show "f = g" unfolding p by blast
  qed
  from pre antisym have partial:"partial_order_on UNIV preco" unfolding partial_order_on_def by auto
  from partial total have "linear_order preco" unfolding linear_order_on_def by auto
  with wf show ?thesis unfolding well_order_on_def by auto
qed
  
lemma SN_prec:"SN {(gm,fn) |fn gm. precsk gm fn }"
proof-
  let ?Rf = "{((FFresh f, n), (FFresh g, m)) |f g n m. ((f,n),(g,m)) ∈ precf - Id}"
  let ?Ro = "{((FOrig f, n), (FOrig g, m)) |f g n m. ((f,n),(g,m)) ∈ preco - Id}"
  let ?Rof = "{((FOrig f, n), (FFresh g, m)) |f g n m. True}"
  let ?F = "λC. (λ(f,n).(C f, n))"
  have inj:"inj (?F FFresh)" unfolding inj_on_def by auto
  have aux:"⋀C P.(map_prod (?F C) (?F C) ` (P - Id)) = {(?F C (f,n), ?F C (g,m)) |f g n m. ((f,n), (g,m)) ∈ P - Id}"
    unfolding split by force
  from well_order_precf have "wf (precf - Id)" using well_order_on_def by auto
  from wf_map_prod_image[OF this inj] have wff:"wf ?Rf" unfolding aux[of FFresh] split by auto
      
  have inj:"inj (?F FOrig)" unfolding inj_on_def by auto
  from well_order_preco have "wf (preco - Id)" unfolding well_order_on_def by auto
  from wf_map_prod_image[OF this inj] have wfo:"wf ?Ro" unfolding aux by simp
      
  have wf:"wf ?Rof" by (smt Pair_inject f_ext.distinct(1) mem_Collect_eq wf_def)
  have "?Rof O ?Rf ⊆ ?Rof" by blast
  from wf_union_compatible[OF wf wff this] have wf:"wf (?Rof ∪ ?Rf)" by auto
  from wf_union_compatible[OF this wfo] have wf:"wf ((?Rof ∪ ?Rf) ∪ ?Ro)" by blast
  hence "wf prec_set" unfolding prec_set_def by auto
  thus ?thesis unfolding precsk_def SN_iff_wf converse_unfold mem_Collect_eq split by simp
qed
  
text ‹Replace Skolem constants by minimal constant c›
definition Dc :: "(('a, 'b) f_ext, 'b) term ⇒ ('a, 'b) term"
  where "Dc t ≡ map_funs_term (λf. case f of FOrig f ⇒ f | FFresh z ⇒ c) t"
    
definition Dcc :: "(('a, 'b) f_ext, 'b) ctxt ⇒ ('a, 'b) ctxt"
  where "Dcc C ≡ map_funs_ctxt (λf. case f of FOrig f ⇒ f | FFresh z ⇒ c) C"
    
definition Dsc :: "(('a, 'b) f_ext, 'b) subst ⇒ ('a, 'b) subst"
  where "Dsc σ ≡ map_funs_subst (λf. case f of FOrig f ⇒ f | FFresh z ⇒ c) σ"
     
lemma Dc_h: "Dc (sk.h t) = t" unfolding Dc_def sk.h_def map_funs_term_comp
  by (simp add: funs_term_map_funs_term_id)
    
abbreviation kbo_sk :: "(('a, 'b) f_ext, 'b) term ⇒ (('a, 'b) f_ext, 'b) term ⇒ bool × bool"
  where "kbo_sk ≡ kbo.kbo (λf.1) 1 (λf n. 1) (λf. f = FOrig c) precsk precsk=="

lemma irrefl_precsk:"∀fn. ¬ (fn,fn) ∈ prec_set"
  using SN_prec unfolding refl_on_def  precsk_def by fastforce
    
lemma trans_preceq_sk:
  assumes "fn = gm ∨ precsk fn gm" and "gm = hk ∨ precsk gm hk"
  shows "(fn = gm ∧ gm = hk) ∨ precsk fn hk"
proof-
  obtain f n g m h k where fn:"fn = (f,n)" and gm:"gm = (g,m)" and hk:"hk = (h,k)" by fastforce
  from assms have a:"(f,n) = (g,m) ∨ precsk (f,n) (g,m)" "(g,m) = (h,k) ∨ precsk (g,m) (h,k)"
    unfolding fn gm hk by auto
  have "⋀R. well_order R ⟹ trans R"
    unfolding well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def by auto
  note trans = this[OF well_order_precf] this[OF well_order_preco]
  { assume p:"((h,k), (g,m)) ∈ prec_set" "((g,m),(f,n)) ∈ prec_set" and neq:"(h,k) ≠ (f,n)"
    have "((h,k),(f,n)) ∈ prec_set" proof(cases h)
      case (FFresh a)
      from p[unfolded this] prec_set_def obtain b where g:"g = FFresh b" by auto
      from p[unfolded this] prec_set_def obtain c where f:"f = FFresh c" by auto
      from p[unfolded FFresh g f] prec_set_def have "((a,k),(b,m)) ∈ precf" "((b,m),(c,n)) ∈ precf" by auto
      note t = trans(1)[unfolded trans_def, rule_format, OF this]
      with neq show ?thesis unfolding FFresh f prec_set_def by auto
    next
      case (FOrig a)
      note h = this
      then show ?thesis proof(cases g)
        case (FFresh b)
        then show ?thesis using neq p unfolding prec_set_def FOrig FFresh by (cases h, auto)
      next
        case (FOrig b)
        show ?thesis using neq trans(2)[unfolded trans_def, rule_format, of "(a,k)" "(b,m)"] p
          unfolding prec_set_def h FOrig by (cases f, auto)
      qed
    qed
  }
  note trans = this
  show ?thesis
  proof (cases "(g, m) = (h, k)")
    case True
    from a show ?thesis unfolding fn gm hk precsk_def split True by auto
  next
    case False
    with a have p:"((h, k), g, m) ∈ prec_set" unfolding precsk_def by auto
    show ?thesis proof (cases "(f, n) = (g, m)")
      case True
      with a p show ?thesis unfolding fn gm hk precsk_def by auto
    next
      case False
      from False a have p':"((g, m), f, n) ∈ prec_set" unfolding precsk_def by auto
      { assume eq:"(h, k) = (f, n)"
        have aux:"{(gm, fn) |fn gm. (fn, gm) ∈ prec_set}¯ = prec_set" by fast
        from p p' have "(fn,fn) ∈ ({(gm, fn) |fn gm. (fn, gm) ∈ prec_set}¯)+" unfolding fn eq precsk_def aux by auto
        hence False using wf_acyclic[OF SN_imp_wf[OF SN_prec]] unfolding precsk_def
          unfolding acyclic_def by blast
      }
      with p p' trans show ?thesis unfolding fn gm hk precsk_def by fast
    qed
  qed
qed
 
lemma trans_prec_sk:
  assumes "precsk fn gm" and "precsk gm hk"
  shows "precsk fn hk"
proof-
  from trans_preceq_sk assms have "(fn = gm ∧ gm = hk) ∨ precsk fn hk" by presburger
  with assms irrefl_precsk show ?thesis unfolding refl_on_def by auto
qed
 
lemma antisym_prec_sk:
  assumes "precsk fn gm"
  shows "¬ precsk gm fn"
using trans_prec_sk[OF assms, of fn] irrefl_precsk unfolding precsk_def by auto
    
lemma adm:"admissible_kbo (λf.1) 1 precsk precsk== (λf. f = FOrig c) (λf n. 1)"
proof-
  have min:"∀g. g ≠ FOrig c ⟶ precsk (g, 0) (FOrig c, 0)"
  proof(rule, rule)
    fix g :: "('a, 'b) f_ext"
    assume gc:"g ≠ FOrig c"
    have "⋀a. a ≠ c ⟹ ((FOrig c,0), (FOrig a,0)) ∈ prec_set" unfolding preco_def prec_set_def by auto
    with gc show "precsk (g, 0) (FOrig c, 0)" unfolding precsk_def prec_set_def by (cases g, auto)
  qed
  { fix f
    assume a:"∀g. g = f ∨ precsk (g, 0) (f, 0)" "f ≠ FOrig c"
    hence p:"precsk (FOrig c, 0) (f, 0)" by auto
    from min a(2) have p':"precsk (f, 0) (FOrig c, 0)" by auto
    from trans_prec_sk[OF p p'] irrefl_precsk[unfolded refl_on_def] have False
      using SN_on_irrefl SN_prec unfolding precsk_def by blast
  }
  with min irrefl_precsk[unfolded refl_on_def] have
    min:"⋀f. (f = (FOrig c)) = (∀g. g = f ∨ precsk (g, 0) (f, 0))" by auto
  show ?thesis  
  proof(unfold_locales)
    fix fn gm hk
    from trans_preceq_sk show "precsk== fn gm ⟹ precsk== gm hk ⟹ precsk== fn hk" by fast
  next
    fix fn gm
    from antisym_prec_sk show "precsk fn gm = strict precsk== fn gm" by blast
  next
    show "SN {(x, y). precsk x y}" using SN_prec by fast
  next
  qed (insert min, auto)
qed
 
lemma precsk_gtotal:"fn = gm ∨ precsk fn gm ∨ precsk gm fn"
proof-
  { assume neq:"fn ≠ gm"
  obtain f g n m where p:"fn = (f,n)" "gm = (g,m)" by fastforce
  from well_order_preco well_order_precf have t:"total preco" "total precf"
    unfolding well_order_on_def linear_order_on_def by auto
  have "precsk fn gm ∨ precsk gm fn" proof(cases f)
    case (FOrig a)
    from t[unfolded Relation.total_on_def] neq show "precsk fn gm ∨ precsk gm fn"
      unfolding precsk_def prec_set_def p FOrig by (cases g, auto)
  next
    case (FFresh a)
    note a = this
    show "precsk fn gm ∨ precsk gm fn" proof (cases g)
      case (FOrig b)
      show ?thesis unfolding precsk_def prec_set_def p FFresh FOrig by force
    next
      case (FFresh b)
      from t(2)[unfolded Relation.total_on_def] neq show ?thesis  
        unfolding precsk_def prec_set_def p FFresh a by auto
    qed
  qed
  }
  thus ?thesis by auto
qed

abbreviation lex_sk
  where "lex_sk ≡ lex_two {≻} Id  {(s, t). fst (kbo_sk s t)}"
  
definition less_sk :: "(('a,'b) f_ext, 'b) term ⇒ (('a,'b) f_ext, 'b) term ⇒ bool" (infix "≻sk" 50)
  where "less_sk s t = (((Dc s, s), (Dc t, t)) ∈ lex_sk)"
  
abbreviation less_sk_set ("{≻sk}") where "less_sk_set ≡ {(s, t). s ≻sk t}"

abbreviation lesseq_sk (infix "≽sk" 50) where "s ≽sk t ≡ (op ≻sk)== s t"

  
text ‹There exists a ground total extension of ≻ on skolemized terms.›
lemma less_sk: "sk_redord less_sk"
proof-
  define σc :: "'b ⇒ ('a, 'b) term" where c ≡ (λx. Fun c [])"
  have 1:"∀s t. s ≻ t ⟶ sk.h s ≻sk sk.h t"
  proof(rule, rule, rule)
    fix s t
    assume "s ≻ t"
    with subst have "Dc (sk.h s) ≻ Dc (sk.h t)" unfolding Dc_h by auto
    thus "sk.h s ≻sk sk.h t" unfolding less_sk_def by auto
  qed
    
  have 2:"∀s t. ground s ∧ ground t ⟶ s = t ∨ s ≻sk t ∨ t ≻sk s"
  proof(rule, rule, rule)
    fix s t :: "(('a,'b) f_ext, 'b) term"
    assume "ground s ∧ ground t"
    hence g:"ground s" "ground t" by auto
    from admissible_kbo.S_ground_total[OF adm refl precsk_gtotal _ g(1) _ g(2)]
    have kbo_gtotal:"s = t ∨ fst (kbo_sk s t) ∨ fst (kbo_sk t s)" by force
    from g have g':"ground (Dc s)" "ground (Dc t)" unfolding Dc_def by auto
    from gtotal[OF this] kbo_gtotal show "s = t ∨ s ≻sk t ∨ t ≻sk s" unfolding less_sk_def by force
  qed 
    
  have 3:"SN {≻sk}"
  proof-
    from admissible_kbo.S_SN[OF adm] have "SN {(s, t). fst (kbo_sk s t)}" by auto
    from lex_two[OF _ SN_less this, of Id] have "SN {(s,t).(s,t) ∈ lex_sk}" by auto
    then show ?thesis unfolding less_sk_def by fast
  qed
    
  have 4:"∀s t C. s ≻sk t ⟶ C⟨s⟩ ≻sk C⟨t⟩"
  proof(rule+)
    fix s t C
    assume "s ≻sk t"
    then consider "Dc s ≻ Dc t" | "(Dc s = Dc t ∧ fst (kbo_sk s t))" unfolding less_sk_def by auto
    thus "C⟨s⟩ ≻sk C⟨t⟩" proof(cases)
      case 1
      from ctxt[OF 1, of "Dcc C"] show ?thesis
        unfolding Dc_def Dcc_def map_funs_term_ctxt_distrib[symmetric] 
        unfolding Dc_def[symmetric] less_sk_def by auto
    next
      case 2
      hence eq:"Dc C⟨s⟩ = Dc C⟨t⟩" unfolding Dc_def by auto
      from 2 admissible_kbo.S_ctxt[OF adm, of s t C] have "fst (kbo_sk C⟨s⟩ C⟨t⟩)" by blast
      with eq show "C⟨s⟩ ≻sk C⟨t⟩" unfolding less_sk_def by force
    qed
  qed
    
  have 5:"∀s t σ. s ≻sk t ⟶ s ⋅ σ ≻sk t ⋅ σ"
  proof(rule+)
    fix s t σ
    assume "s ≻sk t"
    then consider "Dc s ≻ Dc t" | "(Dc s = Dc t ∧ fst (kbo_sk s t))" unfolding less_sk_def by auto
    thus "s ⋅ σ ≻sk t ⋅ σ" proof(cases)
      case 1
      from subst[OF 1, of "Dsc σ"] show ?thesis 
        unfolding Dc_def Dsc_def map_funs_subst_distrib[symmetric] 
        unfolding Dc_def[symmetric] less_sk_def by auto
    next
      case 2
      hence eq:"Dc (s ⋅ σ) = Dc (t ⋅ σ)" unfolding Dc_def by auto
      from 2 admissible_kbo.S_subst[OF adm, of s t σ] have "fst (kbo_sk (s ⋅ σ) (t ⋅ σ))" by blast
      with eq show ?thesis  unfolding less_sk_def by auto
    qed
  qed
    
  have 6:"∀s t u. s ≻sk t ⟶ t ≻sk u ⟶ s ≻sk u"
  proof(rule+)
    fix s t u
    assume "s ≻sk t" and tu:"t ≻sk u"
    then consider "Dc s ≻ Dc t" | "(Dc s = Dc t ∧ fst (kbo_sk s t))" unfolding less_sk_def by auto
    thus "s ≻sk u" proof(cases)
      case 1
      from tu have "Dc t ≽ Dc u" unfolding less_sk_def by auto
      with trans[unfolded trans_def, OF 1, of "Dc u"] 1 show ?thesis unfolding less_sk_def by auto
    next
      case 2
      have S_NS:"fst (kbo_sk t u) ⟶ snd (kbo_sk t u)"
        using adm admissible_kbo.S_imp_NS by blast
      note trans = admissible_kbo.kbo_trans[OF adm, THEN conjunct1, rule_format, of s t u]
      with 2 S_NS have "fst (kbo_sk t u) ⟹ fst (kbo_sk s u)" by blast
      with 2 tu show ?thesis unfolding less_sk_def by (cases "Dc t ≻ Dc u", auto)
    qed
  qed
    
  have "∀t. ground t ⟶ t ≽sk Fun (FOrig c) []"
  proof(rule+)
    fix t :: "(('a,'b) f_ext, 'b) term"
    assume g:"ground t" and t_neq_c:"t ≠ Fun (FOrig c) []"
    have Dc_c:"Dc (Fun (FOrig c) []) = Fun c []" unfolding Dc_def by auto
    show "t ≻sk Fun (FOrig c) []" proof(cases "∃a. t = Fun (FFresh a) []")
      case True
      then obtain a where t:"t = Fun (FFresh a) []" by auto
      have Dc_f:"Dc t = Fun c []" unfolding t Dc_def Nil by auto
      let ?w = "weight_fun.weight (λf. 1) 1 (λf n. 1)"
      have w:"?w (Fun (FOrig c) []) = ?w t" unfolding t Nil weight_fun.weight.simps by auto
      have p:"precsk (FFresh a, 0) (FOrig c, 0)" unfolding precsk_def prec_set_def by auto
      note kbo = kbo.kbo.simps[of _ _ _ _ _ _ t "Fun (FOrig c) []", simplified]
      from p w have "fst (kbo_sk t (Fun (FOrig c) []))"
        unfolding kbo unfolding t Nil by force
      with Dc_f show ?thesis unfolding less_sk_def Dc_c by auto
    next
      case False
      from g obtain f ts where t:"t = Fun f ts" by (cases t, auto)
      from False t_neq_c have "Dc t ≠ Fun c []" unfolding Dc_def t by (cases f, auto)
      with c_min[rule_format, of "Dc t"] have "Dc t ≻ Fun c []" unfolding Dc_c by auto
      thus ?thesis unfolding less_sk_def Dc_c by auto
    qed
  qed
  hence 7:"∀t. ground t ⟶ op ≻sk== t (sk.h (Fun c []))" by auto
  from 1 2 3 4 5 6 7 show ?thesis unfolding sk_redord_def by force
qed

lemma c_min_less_sk:"ground t ⟶ less_sk== t (Fun (FOrig c) [])"
  using less_sk[unfolded sk_redord_def] sk.h_def by auto
    
sublocale sk_run:gtotal_okb_irun_h R E "op ≻" FOrig less_sk
proof (unfold_locales)
  show "⋀s t u. less_sk s t ⟹ less_sk t u ⟹ less_sk s u" using less_sk[unfolded sk_redord_def] by fast
qed (insert less_sk[unfolded sk_redord_def], auto)

lemma rsteps_sk:"(s,t) ∈ (rstep RR)* ⟹ ((sk s, sk t) ∈ (rstep (sk.hR RR))* )"
  using rstep_imp_rstep_sk by (meson rtrancl_map)


lemma sk_clean_step:
  fixes l'::"('a,'b) term"
  assumes "((sk t) :: (('a, 'b) f_ext, 'b) term) = l⋅σ" "l⋅σ ≻sk r⋅σ" "l = sk.h l'" "r = sk.h r'" "ground (r⋅σ)"
  shows "∃τ u. l⋅σ = l⋅τ ∧ r⋅τ = sk u ∧ l⋅τ ≻sk r⋅τ"
proof-
  define τ where "τ ≡ ext_subst σ l (FOrig c)"
  from assms deskolemize_s have t:"t = l' ⋅ (deskolemize ∘ σ)" by auto
  from sk_deskolemize_subst_vars[OF assms(1)] have vs:"∀x∈vars_term l. σ x = sk (deskolemize (σ x))" by auto
  from assms subst_ext_subst'[of l] have eq:"l⋅σ = l ⋅ τ" unfolding τ_def by metis
  from local.sk_run.okb_h.ext_subst_less[of "FOrig c" r σ] c_min_less_sk assms(5)
    have  "op ≻sk== (r ⋅ σ) (r ⋅ τ)" unfolding τ_def by auto
  with assms(2) sk_run.trans_h[OF assms(2)] have less:"l⋅τ ≻sk r⋅τ" unfolding eq by auto
  define u where "u ≡ deskolemize (r⋅τ)"
  { fix x
    assume x:"x ∈ vars_term r'"
    have "τ x = (sk.hs (deskolemize ∘ τ) ∘s σsk) x" proof(cases "x ∈ vars_term l")
      case True
      hence "sk (deskolemize (τ x)) = τ x" using vs eq[unfolded term_subst_eq_conv] by auto
      then show ?thesis unfolding sk_def subst_compose_def by auto
    next
      case False
      with τ_def x have τx:"τ x = Fun (FOrig c) []" unfolding ext_subst_def by auto
      hence sig:"orig_sig (τ x)" unfolding orig_sig_def by auto
      have "(sk.hs (deskolemize ∘ τ) ∘s σsk) x = sk (deskolemize(τ x))" by (simp add: sk_def subst_compose_def)
      also have "… = τ x" unfolding τx deskolemize.simps sk_def by simp
      finally show ?thesis by auto
    qed
  }
  hence "sk.h r' ⋅ sk.hs (deskolemize ∘ τ) ⋅ σsk = sk.h r'  ⋅ τ"
    using term_subst_eq_conv[of "sk.h r'" τ] by force
  hence eq':"sk u = r⋅τ" unfolding u_def assms(4) deskolemize_subst[OF orig_sig_h] sk_def
    deskolemize_orig_sig[OF orig_sig_h] sk.inv_h_h sk.subst_apply_h by auto
  show ?thesis by (rule exI[of _ τ],rule exI[of _ u], insert eq eq' less, auto)
qed
    

context
  assumes fair:"∀(s, t)∈ PCP_ext (Rω ∪ Eω). (s, t) ∈ (rstep E)"
begin
  
abbreviation "Shω ≡ sk_run.irun_h.Sw"
abbreviation "Ehω ≡ sk_run.irun_h.Eω"
abbreviation "Rhω ≡ sk_run.irun_h.Rω"
abbreviation Ehinf ("Eh") where "Ehinf ≡ (⋃i. sk_run.Eh i)"
abbreviation Rhinf ("Rh") where "Rhinf ≡ (⋃i. sk_run.Rh i)"

  
lemma sk_run_fair: "∀(s, t)∈ sk_run.irun_h.PCP_ext (sk_run.irun_h.Rω ∪ sk_run.irun_h.Eω). (s, t) ∈ (rstep Eh)"
proof
  fix s t
  assume pcp:"(s, t) ∈ sk_run.irun_h.PCP_ext (Rhω ∪ Ehω)"
  from sk_run.PCP_h[unfolded sk_run.hRw sk_run.hEw, OF this]
    have "(sk.h_inv s, sk.h_inv t)∈ PCP_ext (Rω ∪ Eω)" and id:"s = sk.h (sk.h_inv s) ∧ t = sk.h (sk.h_inv t)" by auto
  with fair have "(sk.h_inv s, sk.h_inv t) ∈ (rstep E)" by auto
  with sk.rstep_h[of _ _ "E"] have "(sk.h (sk.h_inv s), sk.h (sk.h_inv t)) ∈ (rstep (sk.hR E))"
    by blast
  with sk_run.hEinf have  "(sk.h (sk.h_inv s), sk.h (sk.h_inv t)) ∈ (rstep Eh)" by metis
  with id show "(s,t)  ∈ (rstep Eh)" by auto    
qed

lemma ground_RE_conv_S_conv:
  assumes "(s,t) ∈ (GROUND (rstep (Rhω ∪ Ehω)))*"
  shows "(s,t) ∈ (GROUND (rstep Shω))*"
proof(cases "s=t")
  case True
  then show ?thesis by auto
next
  case False
  have "⋀Rel. GROUND (Rel) = (GROUND Rel)" unfolding GROUND_def by blast
  with assms[unfolded conversion_def] have c:"(s, t) ∈ (GROUND ((rstep (Rhω ∪ Ehω))))*" by metis
  from False GROUND_rtrancl[OF c]  GROUND_not_ground[OF c] have g:"ground s" "ground t"
    and conv:"(s,t) ∈ (rstep (Rhω ∪ Ehω))*" unfolding conversion_def by auto
  have E:"⋀E. (rstep (Rhω ∪ E))* = (rstep (Rhω ∪ E))*" unfolding conversion_def
    by (metis (no_types, lifting) rstep_simps(5) sup.right_idem sup_left_idem symcl_Un symcl_converse)
  note Sconv = sk_run.irun_h.gterms_ER_conv_implies_S_conv[OF sk_run.irun_h.Rw_less sk_run.irun_h.Rw_less g]
  with E conv show ?thesis unfolding sk_run.irun_h.Sw_def E_ord_def by blast
qed
  
context
  fixes ::"('a,'b) trs"
  assumes R_less:"ℛ ⊆ {≻}" and CR_R:"CR (rstep ℛ)"
  and conv_eq:"(rstep ℛ)* = (rstep (E 0))*"
begin

lemma SN_R:"SN (rstep ℛ)" 
  by (rule SN_subset [OF SN_less], insert compatible_rstep_imp_less[OF R_less], auto)
 
lemma R_NF_eq:
  assumes "(s,t) ∈ (rstep E)"
  shows "∃u. (s,u) ∈ (rstep ℛ)! ∧ (t,u) ∈ (rstep ℛ)!"
proof-
  from assms obtain i where "(s,t) ∈ (rstep (E i ∪ R i))*" by fast
  with oKBi_conversion_ERi conv_eq have "(s, t) ∈ (rstep ℛ)*" by fastforce
  with CR_imp_conversionIff_join[OF CR_R] obtain w where
    w:"(s, w) ∈ (rstep ℛ)*" "(t, w) ∈ (rstep ℛ)*" by auto
  from SN_reaches_NF[of "rstep ℛ"] SN_R have "∃w'. (w, w') ∈ (rstep ℛ)* ∧ w' ∈ NF_trs ℛ"
    unfolding normalizability_def by (simp add: SN_defs)
  then obtain w' where "(w, w') ∈ (rstep ℛ)*" and nf:"w' ∈ NF_trs ℛ" by auto
  with w have seqs:"(s, w') ∈ (rstep ℛ)*" "(t, w') ∈ (rstep ℛ)*" by auto
  show ?thesis by (rule exI[of _ w'], unfold normalizability_def, insert seqs nf, auto)
qed

lemma sk_less_compat:"s ≻ t ⟹ sk s ≻sk sk t"
  using sk_run.compat_h sk_run.subst_h unfolding sk_def by auto

lemma ground_ctxt_less:
  assumes "ground s" and "ground t" and "C⟨s⟩ ≻sk C⟨t⟩"
  shows "s ≻sk t"
proof-
  from assms sk_run.gtotal_h consider "s = t" | "s ≻sk t"  | "t ≻sk s" by auto
  thus ?thesis proof(cases)
    case 3
    with sk_run.okb_h.ctxt have "C⟨t⟩ ≻sk C⟨s⟩" by auto
    from sk_run.okb_h.irrefl sk_run.okb_h.trans[OF assms(3) this] show ?thesis by auto
  qed (insert assms(3) sk_run.okb_h.irrefl, auto)
qed 

lemma ground_subt_less:
  assumes "ground s" and "s ⊳ t"
  shows "s ≻sk t"
proof-
  from assms have gt:"ground t" "s ≠ t" by auto
  from assms obtain C where s:"s = C⟨t⟩" by auto
  { assume "t ≻sk s"
    hence less:"t ≻sk C⟨t⟩" unfolding s by auto
    { fix i
      have "(C ^ i)⟨t⟩ ≻sk (C ^ (Suc i))⟨t⟩" using less sk_run.ctxt_h by (induct i, auto)
    }
    with less sk_run.SN_less_h[unfolded SN_on_def] have False by fast
  }
  with sk_run.gtotal_h assms(1) gt ‹s ≠ t› show ?thesis by auto
qed 

lemma Sw_steps_lesseq:"(s,t) ∈ (rstep Sω)* ⟹ s ≽ t"
  unfolding rtrancl_eq_or_trancl using compatible_rstep_trancl_imp_less[OF Sw_less] by auto

lemma ooverlap_exists:
  assumes "C⟨l1⋅σ1⟩ = l2'⋅σ2'" and "hole_pos C ∈ funposs l2'" and "(l1,r1) ∈ RR" and "(l2',r2') ∈ RR"
  and "¬(r1⋅σ1 ≻ l1⋅σ1)" and "¬(r2'⋅σ2' ≻ l2'⋅σ2')"
  shows "∃r r' μ τ s t. ooverlap {≻} RR r r' (hole_pos C) μ s t ∧ C⟨r1 ⋅ σ1⟩ = s ⋅ τ ∧
                        r2' ⋅ σ2' = t ⋅ τ ∧ l1⋅σ1 = (fst r) ⋅ μ ⋅ τ"
proof-
  have rl1:"∃p. p ∙ (l1, r1) ∈ RR" by (rule exI[of _ 0], insert assms(3), auto)
  from vars_rule_disjoint obtain π
    where π: "vars_rule (π ∙ (l2', r2')) ∩ vars_rule (l1, r1) = {}" ..
  define l2 and r2 and σ2
    where "l2 = π ∙ l2'" and "r2 = π ∙ r2'" and 2 = (Var ∘ Rep_perm (-π)) ∘s σ2'"
  note rename = l2_def r2_def σ2_def
  from assms(4) have "-π ∙ (l2,r2) ∈ RR" unfolding rename by auto
  hence rl2:"∃p. p ∙ (l2,r2) ∈ RR" by (rule exI[of _ "-π"])
  from π have vars:"vars_rule (l1,r1) ∩ vars_rule (l2, r2) = {}"
    unfolding rename rule_pt.permute_prod_eqvt by auto
  define p where "p ≡ hole_pos C" 
  from assms(2) have p:"p ∈ funposs l2" by (auto simp:rename p_def)
  note poss_p = funposs_imp_poss [OF p]
    
  have r_l1:"l2' = -π ∙ l2" by (auto simp:rename)
  from assms(1) have eq:"l1 ⋅ σ1 = (l2' ⋅ σ2' |_ p)" unfolding p_def by (metis subt_at_hole_pos)
  note subt_l1 = subt_at_subst[OF poss_p]
  from eq[unfolded r_l1 permute_term_subst_apply_term] have
    eq':"l1 ⋅ σ1 = l2 |_ p ⋅ σ2" 
    unfolding subt_at_subst[OF poss_p] unfolding rename
    by (metis ‹l2 ≡ π ∙ l2'› eq r_l1 subst.cop_add subt_l1 term_apply_subst_Var_Rep_perm)
      
  note coinc = coincidence_lemma' [of _ "vars_rule (l1, r1)"]
  define σ where "σ x = (if x ∈ vars_rule (l1, r1) then σ1 x else σ2 x)" for x
  have "l1 ⋅ σ = l1 ⋅ (σ |s vars_rule (l1, r1))"
    using coinc[of l1] by (simp add: vars_rule_def)
  also have "… = l1 ⋅ (σ1 |s vars_rule (l1, r1))" by (simp add: σ_def [abs_def])
  finally have l1_coinc:"l1 ⋅ σ = l1 ⋅ σ1" using coinc[of l1] by (simp add: vars_rule_def)
  have disj: "vars_rule (l1, r1) ∩ vars_term l2 = {}"
    using π[unfolded rename] by (auto simp: vars_rule_def rule_pt.permute_prod_eqvt l2_def)
  have unif: "l1 ⋅ σ = (l2 |_ p) ⋅ σ"
  proof -
    from disj have disj: "vars_rule (l1, r1) ∩ vars_term (l2 |_ p) = {}" 
      using vars_term_subt_at [OF poss_p] by auto
    from l1_coinc have "l1 ⋅ σ = (l2 |_ p) ⋅ σ2" using eq' by simp
    also have "… = (l2 |_ p) ⋅ (σ2 |s vars_term (l2 |_ p))"
      by (simp add: coincidence_lemma [symmetric])
    also have "… = (l2 |_ p) ⋅ (σ |s vars_term (l2 |_ p))" using disj by (simp add: σ_def [abs_def])
    finally show ?thesis by (simp add: coincidence_lemma [symmetric])
  qed
    
  define μ where "μ = the_mgu l1 (l2 |_ p)"
  have is_mgu:"is_mgu μ {(l1, l2 |_ p)}" by (rule is_mguI, insert unif the_mgu, auto simp: μ_def)
  with unif obtain τ where σ: "σ = μ ∘s τ" by (auto simp: is_mgu_def unifiers_def)
  from unif have mgu: "mgu l1 (l2 |_ p) = Some μ"
    unfolding μ_def the_mgu_def
    using unify_complete and unify_sound by (force split: option.splits simp: is_imgu_def unifiers_def)
      
  have "r1 ⋅ σ = r1 ⋅ (σ |s vars_rule (l1, r1))" using coinc[of r1] by (simp add: vars_rule_def)
  also have "… = r1 ⋅ (σ1 |s vars_rule (l1, r1))" by (simp add: σ_def [abs_def])
  finally have r1_coinc:"r1 ⋅ σ = r1 ⋅ σ1" using coinc[of r1] by (simp add: vars_rule_def)
  from assms(5) subst have nonoriented1:"¬ r1 ⋅ μ ≻ l1 ⋅ μ"
    unfolding l1_coinc[symmetric] r1_coinc[symmetric] σ by auto
      
  note coinc = coincidence_lemma [symmetric]
  have "l2 ⋅ σ = l2 ⋅ (σ |s vars_term l2)" by (simp add: coinc)
  also have "… = l2 ⋅ (σ2 |s vars_term l2)" using disj by (simp add: σ_def [abs_def])
  also have "… = l2 ⋅ σ2" by (simp add: coinc)
  finally have l2_coinc:"l2 ⋅ σ = l2' ⋅ σ2'" unfolding rename by simp
  have disj: "vars_rule (l1, r1) ∩ vars_term r2 = {}"
    using vars_term_subt_at [OF poss_p] and π[unfolded rename]
    by (auto simp: vars_rule_def rule_pt.permute_prod_eqvt r2_def)
  have "r2 ⋅ σ = r2 ⋅ (σ |s vars_term r2)" by (simp add: coinc)
  also have "… = r2 ⋅ (σ2 |s vars_term r2)" using disj by (simp add: σ_def [abs_def])
  also have "… = r2 ⋅ σ2" by (simp add: coinc)
  finally have r2_coinc:"r2 ⋅ σ = r2' ⋅ σ2'" unfolding rename by simp
  from assms(6) subst have nonoriented2:"¬ r2 ⋅ μ ≻ l2 ⋅ μ"
    unfolding l2_coinc[symmetric] r2_coinc[symmetric] σ by auto
  from assms(1) poss_p hole_pos_id_ctxt have "C = ctxt_of_pos_term p (l2 ⋅ σ)"
    unfolding l1_coinc[symmetric] l2_coinc[symmetric] p_def by metis  
  hence cp_left:"C⟨r1 ⋅ σ1⟩ = (ctxt_of_pos_term p (l2 ⋅ μ))⟨r1 ⋅ μ⟩ ⋅ τ"
    unfolding subst_apply_term_ctxt_apply_distrib r1_coinc[symmetric]
      ctxt_of_pos_term_subst[OF poss_imp_subst_poss[OF poss_p], symmetric]
      subst_subst_compose[symmetric] σ[symmetric] by simp
  from subst_subst_compose have cp_right:"r2' ⋅ σ2' = r2 ⋅ μ ⋅ τ" unfolding r2_coinc[symmetric] σ by auto
  let ?s = "(ctxt_of_pos_term p (l2 ⋅ μ))⟨r1 ⋅ μ⟩"
  let ?t = "r2 ⋅ μ"
  from l1_coinc σ have "l1 ⋅ σ1 = l1 ⋅ μ ⋅ τ" by auto
  note facts = rl1 rl2 vars p mgu nonoriented1 nonoriented2 cp_left cp_right this
  show ?thesis unfolding ooverlap_def
    by (rule exI[of _ "(l1,r1)"], rule exI[of _ "(l2,r2)"], rule exI[of _ μ],
        rule exI[of _ τ], rule exI[of _ ?s], rule exI[of _ ?t], insert facts,
        unfold fst_conv snd_conv p_def, auto)
qed

lemma less_imp_less_sk:"s ≻ t ⟹ sk s ≻sk sk t"
  using less_sk[unfolded sk_redord_def] unfolding sk_def by meson

lemma correctness_okb_sk:"GCR (rstep Shω) ∧ (rstep (sk_run.Eh 0))* = (rstep (Rhω ∪ Ehω))*"
  using sk_run.irun_h.correctness_okb[OF sk_run_fair] by auto

abbreviation S_sk_inf ("𝒮sk") where "𝒮sk ≡ sk_run.irun_h.Rinf ∪ E_ord (op ≻sk) sk_run.irun_h.Einf"

lemma S_sk_inf:
  assumes "(l,r) ∈ 𝒮sk"
  shows "∃l' r' τ. l = (sk.h l') ⋅ τ ∧ r = (sk.h r') ⋅ τ ∧ (l',r') ∈ E ∪ R ∧ sk.h l' ⋅ τ ≻sk sk.h r' ⋅ τ"
proof-
  from assms consider "(l,r) ∈ sk_run.irun_h.Rinf" | "(l,r) ∈ E_ord (op ≻sk) sk_run.irun_h.Einf" by auto
  thus ?thesis proof(cases)
    case 1
    then obtain i where lr:"(l, r) ∈ sk.hR (R i)" by auto
    with sk_run.irun_h.Ri_less have "l ≻sk r" by auto
    note rule = sk.inv_hR_rule[OF lr] this
    show ?thesis by (rule exI[of _ "sk.h_inv l"], rule exI[of _ "sk.h_inv r"], rule exI[of _ Var], insert rule, auto)
  next
    case 2
    then obtain l' r' τ where lr:"l = l' ⋅ τ" "r = r' ⋅ τ" "l' ⋅ τ ≻sk r' ⋅ τ" "(l', r') ∈ sk_run.irun_h.Einf"
      unfolding E_ord_def mem_Collect_eq by blast
    then obtain i where lr':"(l', r') ∈ (sk.hR (E i))" unfolding E_ord_def by auto
    let ?l = "sk.h_inv l'" and ?r = "sk.h_inv r'"
    note sk_E = sk.inv_hR_rule[of l' r' "E i"] sk.inv_hR_rule[of r' l' "E i"] lr'
    hence "(?l, ?r) ∈ (E i) ∧ l' = (sk.h ?l) ∧ r' = (sk.h ?r)" unfolding Un_iff converse_iff by fastforce
    hence inE:"(?l, ?r) ∈ E" and lr':"(sk.h ?l) = l'" "(sk.h ?r) = r'" by auto
    show ?thesis by (rule exI[of _ ?l], rule exI[of _ "?r"], rule exI[of _ τ], insert inE lr(3), unfold lr lr', simp)
  qed
qed

lemma peak_cases:
  assumes step:"(s,t) ∈ rstep_r_p_s (Eω) (l,r) ε σ" and irstep:"(t,u) ∈ irstep False Sω"
    and "¬(l⋅σ ≻ r⋅σ)"
  shows "s ∉ NF_trs Sω ∨ (∃w. (s,w) ∈ (rstep (E))= ∧ u ≽ w)" (is "?A ∨ ?B")
proof-
  from step have lr:"(l,r) ∈ Eω" "s = l⋅σ" "t = r⋅σ" unfolding rstep_r_p_s_def by auto
  from irstep[unfolded irstep_def, THEN qrstepE'] obtain D l2 r2 σ2 where
    lr2:"(l2,r2) ∈ Sω" "t = D⟨l2⋅σ2⟩" "u = D⟨r2⋅σ2⟩" "∀u⊲l2 ⋅ σ2. u ∈ NF_terms (lhss Sω)" by blast 
  let ?p = "hole_pos D"
    (* do a CP Lemma-like analysis (standard CP lemma requires variable condition) *)
  show ?thesis proof(cases "?p ∈ funposs r")
    case False (* variable overlap *)
    note hole_D_not_in_funposs_r = this
    from lr(3) lr2(2) have p_rσ:"?p ∈ poss (r ⋅ σ)" by fastforce
    from poss_subst_apply_term[OF this hole_D_not_in_funposs_r] obtain q1 q2 x
      where p[simp]: "?p = q1 <#> q2" and q1: "q1 ∈ poss r"
        and rq1: "r |_ q1 = Var x" and q2: "q2 ∈ poss (σ x)" by auto
    note eq = lr(3)[unfolded lr2(2), simplified]
    hence [simp]: "r ⋅ σ |_ ?p = l2 ⋅ σ2" using subt_at_hole_pos by metis
    with rq1 q2 subt_at_append p_rσ q1 have [simp]: "σ x |_ q2 = l2 ⋅σ2" unfolding p by auto
    define σ' where "σ' y = (if y = x then replace_at (σ x) q2 (r2 ⋅ σ2) else σ y)" for y
    have "(σ x, σ' x) ∈ rstep Sω"
    proof -
      let ?C = "ctxt_of_pos_term q2 (σ x)"
      have "(?C⟨l2 ⋅ σ2⟩, ?C⟨r2 ⋅ σ2⟩) ∈ rstep Sω" using lr2 by blast
      then show ?thesis using q2 by (simp add: σ'_def replace_at_ident)
    qed
    then have *: "⋀x. (σ x, σ' x) ∈ (rstep Sω)*" by (auto simp: σ'_def)
    then have left:"(l ⋅ σ, l ⋅ σ') ∈ (rstep Sω)*" by (rule substs_rsteps)
    moreover
    have right:"(u, r ⋅ σ') ∈ (rstep Sω)*"
    proof -
      have "replace_at (r ⋅ σ) ?p (r2 ⋅ σ2) = replace_at (r ⋅ σ) q1 (σ' x)"
        using q1 and q2 by (simp add: σ'_def ctxt_of_pos_term_append rq1)
      moreover
      have "(replace_at (r ⋅ σ) q1 (σ' x), r ⋅ σ') ∈ (rstep Sω)*"
        by (rule replace_at_subst_rsteps [OF * q1 rq1])
      ultimately
      have "(replace_at (r ⋅ σ) ?p (r2 ⋅ σ2), r ⋅ σ') ∈ (rstep Sω)*" by auto
      thus ?thesis unfolding lr2(3) unfolding eq[symmetric] ctxt_of_pos_term_hole_pos
        by blast 
    qed
    from lr(1) in_mono[OF Ew_subset_Einf] have E_step:"(l ⋅ σ', r ⋅ σ') ∈ rstep (E)" by fast
    from left right Sw_steps_lesseq have leq:"u ≽ r ⋅ σ'" unfolding lr(2) by simp
    show ?thesis proof(cases "l ⋅ σ = l ⋅ σ'")
      case True
      have "?B" unfolding True lr(2) by (rule exI[of _ "r ⋅ σ'"], insert E_step leq, auto)
      thus ?thesis by auto
    next
      case False
      with left[unfolded rtrancl_eq_or_trancl trancl_unfold_left] show ?thesis
        unfolding lr(2) by auto
    qed
  next
    case True (* proper overlap *)
    from lr lr2 have eq:"D⟨l2 ⋅ σ2⟩ = r ⋅ σ" by auto
    note st = lr[unfolded ctxt_apply_term.simps]
    note tu = lr2[unfolded ctxt_apply_term.simps]
    note rl_cases = lr2(1)[unfolded Sw_def Un_iff[of _ "Rω"]]
    have "∃τ l3 r3. l2 = l3⋅τ ∧ r2 = r3⋅τ ∧ (l3, r3) ∈ Rω ∪ Eω"
      by (cases "(l2, r2) ∈ Rω", rule exI[of _ Var], insert rl_cases, auto)
    then obtain τ l3 r3 where lr3:"l2 = l3⋅τ" "r2 = r3⋅τ" "(l3, r3) ∈ Rω ∪ Eω" by auto
    from eq[unfolded this] subst_subst have eq':"D⟨l3 ⋅ (τ ∘s σ2)⟩ = r ⋅ σ" by auto
    from lr(1) have rl0:"(r,l) ∈ Rω ∪ Eω" by auto
    from lr2(1) Sw_less have "l3 ⋅ τ ≻ r3 ⋅ τ" unfolding lr3 by auto
    from subst[OF this] trans irrefl have unoriented':"¬ r3 ⋅ τ ∘s σ2 ≻ l3 ⋅ τ ∘s σ2"
      unfolding subst_subst[symmetric] by fast

    note ooverlap = ooverlap_exists[OF eq' True lr3(3) rl0 unoriented' assms(3)]
    hence ooverlap:"∃r r' μ τ s t. ooverlap {≻} (Rω ∪ Eω) r r' (hole_pos D) μ s t ∧
               D⟨r2 ⋅ σ2⟩ = s ⋅ τ ∧ l ⋅ σ = t ⋅ τ ∧ l2 ⋅ σ2 = (fst r) ⋅ μ ⋅ τ" unfolding lr3 by auto
    then obtain rl1 rl2 μ τ s0 t0 where
      ooverlap:"ooverlap {≻} (Rω ∪ Eω) rl1 rl2 (hole_pos D) μ s0 t0"
      and dec:"D⟨r2 ⋅ σ2⟩ = s0 ⋅ τ" "l ⋅ σ = t0 ⋅ τ" "l2 ⋅ σ2 = (fst rl1) ⋅ μ ⋅ τ" by auto
    { fix u
      assume "u ⊲ fst rl1 ⋅ μ"
      hence  "u ⋅ τ ⊲ fst rl1 ⋅ μ ⋅ τ" by auto
      with lr2(4)[unfolded dec] have nf:"u ∈ NF_trs Sω" unfolding NF_terms_lhss by fast
      have "⋀s t. (s,t) ∈ ordstep {≻} (Rω ∪ Eω) ⟹ (s,t) ∈ rstep Sω"
        unfolding Sw_def rstep_union ordstep.simps by blast
      from nf NF_anti_mono[OF subrelI[OF this]] have "u ∈ NF (ordstep {≻} (Rω ∪ Eω))"
        by blast
    }
    hence prime:"(∀u⊲fst rl1 ⋅ μ. u ∈ NF (ordstep {≻} (Rω ∪ Eω)))" by auto

    with ooverlap have "(s0, t0) ∈ PCP_ext (Rω ∪ Eω)" unfolding PCP_ext_def mem_Collect_eq
      by blast
    with fair have "(s0, t0) ∈ (rstep E)" by auto
    with dec have "(D⟨r2 ⋅ σ2⟩, l ⋅ σ) ∈ (rstep E)" by auto
    with lr(2) lr2(3) have Einf_step:"(s, u) ∈ rstep (E)"
      unfolding rstep_simps(5)[symmetric] by fastforce
    thus ?thesis by auto
  qed
qed

lemma REinf_step_Sinf_step:
  assumes  "sk s ≻sk sk t" and "(s, t) ∈ rstep (R ∪ (E))"
  shows "(sk s, sk t) ∈ rstep 𝒮sk"
proof(cases "(s, t) ∈ rstep(E)")
  case True
  from rstep_imp_rstep_sk[OF this] sk_run.hEinf have "(sk s, sk t) ∈ rstep (Eh)"
    unfolding sk.hR_sym by auto
  then obtain l r σ D where *:"sk s = D⟨l⋅σ⟩" "sk t = D⟨r⋅σ⟩" "(l,r) ∈ Eh" by fast
  note g = ground_sk[of s, unfolded * ground_ctxt_apply] ground_sk[of t, unfolded * ground_ctxt_apply]
  with ground_ctxt_less assms(1) have "l⋅σ ≻sk r⋅σ" unfolding * by fast
  with * have "(l⋅σ, r⋅σ) ∈ E_ord (op ≻sk) Eh" unfolding E_ord_def by fast
  thus ?thesis unfolding * by auto
next
  case False
  with assms have "(s, t) ∈ rstep R" unfolding rstep_union by auto
  from rstep_imp_rstep_sk[OF this] sk_run.hRinf have "(sk s, sk t) ∈ rstep Rh" by auto
  thus ?thesis unfolding rstep_union by auto
qed

lemma sk_step_exists':
  assumes "(sk t, u) ∈ rstep ((sk.hR ℛ) :: (('a,'b) f_ext, 'b) trs)"
  shows "∃u'. (sk t, sk u') ∈ rstep (sk.hR ℛ)"
proof-
  from assms obtain l r D τ where lr:"(l,r) ∈ sk.hR ℛ" and sk_t:"sk t = D⟨l⋅τ⟩" "u = D⟨r⋅τ⟩" by fast
  with sk.inv_hR_rule have **:"l = sk.h (sk.h_inv l)" "r = sk.h (sk.h_inv r)" "(sk.h_inv l,sk.h_inv r) ∈ ℛ" by metis+
  let ?l = "sk.h_inv l" and ?r = "sk.h_inv r"
  from sk_C[OF sk_t(1)] have p:"hole_pos D ∈ poss t" and sk_subt:"sk (t |_ hole_pos D) = l ⋅ τ" by auto
  let ?t = "t |_ hole_pos D"
  from sk_subt ** have 1:"sk ?t =  (sk.h ?l) ⋅ τ" by auto
  from SN_imp_variable_condition[OF SN_R] have "(⋀l r. (l, r) ∈ sk.hR ℛ ⟹ vars_term r ⊆ vars_term l)"
    by (metis case_prodD sk.h_def sk.inv_hR_rule vars_term_map_funs_term2)
  with rstep_ground[OF _ ground_sk assms] have "ground u" by auto
  hence 3:"ground (r ⋅ τ)" using sk_t **  by simp
  from ** R_less have "?l ≻ ?r" by auto
  from sk_run.compat_h[OF this] sk_run.subst_h have 2:"l ⋅ τ ≻sk r ⋅ τ" unfolding **[symmetric] by auto
  from sk_clean_step[OF 1, unfolded **[symmetric], OF 2 _ _ 3, of ?l ?r] ** obtain u' τ' where
    c:"l ⋅ τ = l ⋅ τ'" "r ⋅ τ' = sk u'" "l ⋅ τ' ≻sk r ⋅ τ'" by fastforce
  hence t':"sk t = D⟨sk.h ?l ⋅ τ'⟩" using sk_t ** by auto
  from lr c(3) have "(l ⋅ τ',r ⋅ τ') ∈ rstep (sk.hR ℛ)" unfolding **[symmetric] by blast
  hence step:"(sk (t |_ hole_pos D), sk u') ∈ rstep (sk.hR ℛ)" using sk_subt ** c by auto
  from sk_ctxt_exists[OF sk_t(1)] obtain C v where Cv:"D = (sk.hC C) ⋅c σsk" "l ⋅ τ = sk v" by auto
  from step have "(D⟨sk (t |_ hole_pos D)⟩, D⟨sk u'⟩) ∈ rstep (sk.hR ℛ)" by auto
  hence "(sk t, ((sk.hC C) ⋅c σsk)⟨sk u'⟩) ∈ rstep (sk.hR ℛ)" unfolding sk_subt sk_t using Cv by auto
  hence "(sk t, sk C⟨u'⟩) ∈ rstep (sk.hR ℛ)"
    unfolding sk_def subst_apply_term_ctxt_apply_distrib sk.ctxt_apply_h' by auto
  thus ?thesis by auto
qed

lemma sk_step_exists:
  assumes "(sk t, u) ∈ rstep 𝒮sk"
  shows "∃u'. (sk t, sk u') ∈ rstep 𝒮sk"
proof-
  from assms obtain l r D τ where lr:"(l,r) ∈ 𝒮sk" and sk_t:"sk t = D⟨l⋅τ⟩" "u = D⟨r⋅τ⟩" by fast
  from S_sk_inf[OF this(1)] obtain l' r' μ where
    **:"l = (sk.h l') ⋅ μ" "r = (sk.h r') ⋅ μ" "(l',r') ∈ E ∪ R" "sk.h l' ⋅ μ ≻sk sk.h r' ⋅ μ" by fast+
  from sk_C[OF sk_t(1)] have p:"hole_pos D ∈ poss t" and sk_subt:"sk (t |_ hole_pos D) = l ⋅ τ" by auto
  let ?t = "t |_ hole_pos D"
  from sk_subt ** have 1:"sk ?t =  (sk.h l') ⋅ (μ ∘s τ)" by auto
  from **(4) sk_run.subst_h have 2:"sk.h l' ⋅ (μ ∘s τ) ≻sk sk.h r' ⋅ (μ ∘s τ)" by auto
  from rstep_ground[OF _ ground_sk assms] SN_imp_variable_condition[OF SN_Sinf] have "ground u"
    by (metis SN_imp_variable_condition case_prod_conv sk_run.irun_h.SN_Sinf)
  hence 3:"ground (sk.h r' ⋅ μ ∘s τ)" unfolding sk_t ** subst_subst_compose by simp
  from sk_clean_step[OF 1 2 _ _ 3] subst_subst_compose obtain u' τ' where
    c:"sk.h l' ⋅ μ ⋅ τ = sk.h l' ⋅ τ'" "sk u' = sk.h r' ⋅ τ'" "sk.h l' ⋅ τ' ≻sk sk.h r' ⋅ τ'" by fastforce
  hence t':"sk t = D⟨sk.h l' ⋅ τ'⟩" using sk_t ** by auto
  have "(sk.h l' ⋅ τ', sk.h r' ⋅ τ') ∈ rstep 𝒮sk" proof(cases "(l',r') ∈ E")
    case True
    hence "(sk.h l',sk.h r') ∈ Eh" using sk.rule_h sk_run.hEinf by auto
    with c(3) have "(sk.h l' ⋅ τ',sk.h r' ⋅ τ') ∈ E_ord (op ≻sk) Eh"
      unfolding E_ord_def mem_Collect_eq by blast
    thus ?thesis by auto
  next
    case False
    hence "(l',r') ∈ R" using **(3) by auto
    hence "(sk.h l',sk.h r') ∈ Rh" using sk.rule_h sk_run.hEinf by auto
    then show ?thesis by auto
  qed
  hence step:"(sk (t |_ hole_pos D), sk u') ∈ rstep 𝒮sk" unfolding sk_subt ** c by auto
  from sk_ctxt_exists[OF sk_t(1)] obtain C v where Cv:"D = (sk.hC C) ⋅c σsk" "l ⋅ τ = sk v" by auto
  from step have "(D⟨sk (t |_ hole_pos D)⟩, D⟨sk u'⟩) ∈ rstep 𝒮sk" by auto
  hence "(sk t, ((sk.hC C) ⋅c σsk)⟨sk u'⟩) ∈ rstep 𝒮sk" unfolding sk_subt sk_t using Cv by auto
  hence "(sk t, sk C⟨u'⟩) ∈ rstep 𝒮sk"
    unfolding sk_def subst_apply_term_ctxt_apply_distrib sk.ctxt_apply_h' by auto
  thus ?thesis by auto
qed

lemma hR_less_sk: "sk.hR ℛ ⊆ {≻sk}"
proof
  fix l r :: "(('a, 'b) f_ext, 'b) term"
  assume "(l, r) ∈ sk.hR ℛ"
  with sk.inv_hR_rule have **:"l = sk.h (sk.h_inv l)" "r = sk.h (sk.h_inv r)" "(sk.h_inv l,sk.h_inv r) ∈ ℛ" by metis+
  with R_less have "sk.h_inv l ≻ sk.h_inv r" by auto
  with sk_run.compat_h[of "sk.h_inv l" "sk.h_inv r"] ** show "(l,r) ∈ {≻sk}" by auto
qed

lemma NF_S_NF_R:
  assumes "sk t ∈ NF_trs Shω"
  shows "sk t ∈ NF_trs (sk.hR ℛ)"
proof-
  { fix u :: "(('a, 'b) f_ext, 'b) term"
    assume step:"(sk t, u) ∈ rstep (sk.hR ℛ)"
    from sk_step_exists'[OF step] obtain v where v:"(sk t, sk v) ∈ rstep (sk.hR ℛ)" by auto
    with sk_run.irun_h.rstep_subset_less[OF hR_less_sk] have less:"sk t ≻sk sk v" by auto
    note acyclic = SN_subset[OF sk_run.SN_less_h sk_run.irun_h.rstep_subset_less[OF hR_less_sk],THEN SN_imp_acyclic]
    with r_into_trancl'[OF v] sk_inj have "t ≠ v" unfolding acyclic_def by auto
    with rstep_sk_imp_rstep[OF v] have "(t, v) ∈ rstep ℛ" by auto
    with conv_eq oKBi_conversion_ERw have "(t, v) ∈ (rstep (Rω ∪ Eω))*" by auto

    with rsteps_sk have "(sk t, sk v) ∈ (rstep (sk.hR ((Rω ∪ Eω))))*"
      unfolding conversion_def rstep_simps(5)[symmetric] by auto
    with sk_run.hRw sk_run.hEw have "(sk t, sk v) ∈ (rstep (Rhω ∪ Ehω))*"
      unfolding conversion_def sk.hR_sym using sk.hR_union rstep_simps(5) by metis
    note conv_ground = gterm_conv_GROUND_conv[OF ground_sk ground_sk this]
    with ground_RE_conv_S_conv have gconv:"(sk t, sk v) ∈ (GROUND (rstep Shω))*" by simp
    note GCR = sk_run.irun_h.correctness_okb[OF sk_run_fair, THEN conjunct1]
    with gconv CR_imp_conversionIff_join
    obtain w where w:"(sk t,w) ∈ (GROUND (rstep Shω))*" "(sk v,w) ∈ (GROUND (rstep Shω))*" by blast
    hence v:"(sk t,w) ∈ (rstep Shω)*" "(sk v,w) ∈ (rstep Shω)*" using rtrancl_mono[OF GROUND_subset]
      by auto
    note compat = sk_run.okb_h.compatible_rstep_trancl_imp_less[OF sk_run.irun_h.Sw_less]
    with v(2)[unfolded rtrancl_eq_or_trancl] have geq:"sk v ≽sk w" by auto
    from less sk_run.trans_h geq have "w ≠ sk t" using sk_run.okb_h.irrefl by auto
    with v(1)[unfolded rtrancl_eq_or_trancl] have "∃s. (sk t, s) ∈ rstep Shω"
      unfolding trancl_unfold_left by auto
  }
  with assms show ?thesis by blast
qed

lemma sk_Swinf_step_imp_no_Sinf_NF:
  assumes "(sk s, sk t) ∈ rstep 𝒮sk"
  shows "s ∉ NF_trs (𝒮 R E)"
  using assms
proof(induct "sk t" arbitrary:s t rule: SN_induct [OF sk_run.SN_less_h])
  case 1
  note ind_step = this
  then obtain l r C σ where lr:"(l,r) ∈ 𝒮sk" and st:"sk s = C⟨l⋅σ⟩" "sk t = C⟨r⋅σ⟩" by fast
  from S_sk_inf[OF this(1)] obtain l' r' τ where
    *:"l = (sk.h l') ⋅ τ" "r = (sk.h r') ⋅ τ" "(l',r') ∈ E ∪ R" "sk.h l' ⋅ τ ≻sk sk.h r' ⋅ τ" by fast+
  note ordstep = sk_run.okb_h.ordstep_S_conv[OF sk_run.irun_h.Rinf_less]
    ordstep_imp_ord[OF sk_run.okb_h.ctxt_closed_less]
  from 1 ordstep have sk_less:"sk s ≻sk sk t" by blast
  let  = "τ ∘s σ" and ?l = "sk.h l'" and ?r = "sk.h r'"
  from * st have st:"sk s = C⟨?l⋅?σ⟩" "sk t = C⟨?r⋅?σ⟩" by auto
  show ?case proof(cases C)
    case Hole
    define ρ where "ρ ≡ deskolemize ∘ (τ ∘s σ)"
    note sk_t = st(2)[unfolded Hole ctxt_apply_term.simps]
    from st deskolemize_s have s:"s = l' ⋅ ρ" and t:"t = r' ⋅ ρ" unfolding Hole ρ_def by auto
    from sk_deskolemize_subst_vars[OF sk_t] have sk_vs:"∀x∈vars_term r'. ?σ x = sk (deskolemize (?σ x))"
      unfolding Hole ρ_def by auto
    show ?thesis proof (cases "l'⋅ ρ ≻ r'⋅ ρ")
      case True
      from * consider "(l',r') ∈ E" | "(l',r') ∈ R" by auto
      thus ?thesis proof(cases)
        case 1
        have "(l'⋅ ρ,r'⋅ ρ) ∈ (E_ord (op ≻) E)" unfolding E_ord_def mem_Collect_eq
          by (rule exI[of _ l'], rule exI[of _ r'], rule exI[of _ ρ], insert True 1, auto)
        thus ?thesis unfolding s by auto
      next
        case 2
        thus ?thesis unfolding s by fast
      qed
    next
      case False
      note unoriented = this
      with *(3) subst Rinf_less have lr':"(l', r') ∈ E" by fast
      show ?thesis proof (cases "(l', r') ∈ Eω")
        case True
        show ?thesis proof(cases "sk t ∈ NF_trs 𝒮sk")
          case True
          with sk_run.irun_h.NF_Sw_NF_Sinf have "sk t ∈ NF_trs Shω" by auto
          with NF_S_NF_R have nf:"sk t ∈ NF_trs (sk.hR ℛ)" by auto
          from lr' R_NF_eq[of s t] have "∃u. (s, u) ∈ (rstep ℛ)! ∧ (t, u) ∈ (rstep ℛ)!"
            unfolding s t by fast
          then obtain u where "(s, u) ∈ (rstep ℛ)! ∧ (t, u) ∈ (rstep ℛ)!" by blast
          with normalizability_def have "(s, u) ∈ (rstep ℛ)* ∧ (t, u) ∈ (rstep ℛ)*" by fastforce
          with rsteps_sk have "(sk s, sk u) ∈ (rstep (sk.hR ℛ))* ∧ (sk t, sk u) ∈ (rstep (sk.hR ℛ))*" by fastforce
          with NF_join_imp_reach[OF nf] have steps:"(sk s, sk t) ∈ (rstep (sk.hR ℛ))*" by auto
          from sk_less SN_imp_acyclic[OF sk_run.SN_less_h] sk_inj have "s ≠ t" unfolding acyclic_def by auto
          with rstep_imp_deskolemize_rsteps[OF steps] have "(s, t) ∈ (rstep ℛ)+"
            unfolding deskolemize_sk rtrancl_eq_or_trancl by auto
          with compatible_rstep_trancl_imp_less[OF R_less] have "s ≻ t" by auto
          with lr' unoriented show ?thesis using s t by blast
        next
          case False
          then obtain u where step:"(sk t, u) ∈ rstep 𝒮sk" by blast
          from sk_step_exists[OF step] obtain u' where step:"(sk t, sk u') ∈ rstep 𝒮sk" by auto
          with ordstep have "sk t ≻sk sk u'" by blast
          with sk_run.trans_h[OF sk_less this] ind_step(1)[OF _ step] have "t ∉ NF_trs (𝒮 R E)" by auto
          then obtain v where irstep:"(t, v) ∈ irstep False Sω" by (metis NF_I NF_Sw_NF_Sinf rstep_imp_irstep)

          from True s t have root_step:"(s, t) ∈ rstep_r_p_s (Eω) (l', r') ε ρ" unfolding rstep_r_p_s_def by force
          from unoriented peak_cases[OF root_step irstep] consider
            "s ∉ NF_trs Sω" | "∃w. (s, w) ∈ (rstep (E))= ∧ v ≽ w" by auto
          thus ?thesis proof(cases,simp add: NF_Sw_NF_Sinf)
            case 2
            then obtain w where w:"(s, w) ∈ (rstep (E))=" "v ≽ w" by auto
            from irstep qrstep_subset_rstep compatible_rstep_imp_less[OF Sw_less] have "t ≻ v"
              unfolding irstep_def by blast
            with less_imp_less_sk have tv:"sk t ≻sk sk v" by auto
            from less_imp_less_sk w have "sk v ≻sk sk w ∨ sk v = sk w" by auto
            with tv sk_run.trans_h have tw:"sk t ≻sk sk w" by metis
            with sk_less sk_run.trans_h have sw:"sk s ≻sk sk w" by metis
            with SN_imp_acyclic[OF sk_run.SN_less_h] have "s ≠ w" unfolding acyclic_def by auto
            with w have "(s, w) ∈ rstep (E)" by auto
            with REinf_step_Sinf_step[OF sw] have "(sk s, sk w) ∈ rstep 𝒮sk" by fast
            from ind_step(1)[OF _ this] tw show ?thesis by auto
          qed
        qed
      next
        case False
        from Einf_sym_without_Ew_sym[OF lr' False] obtain j where lr_cases:
          "(l', r') ∈ E j ∧ (l', r') ∉ E (Suc j) ∨ (r', l') ∈ E j ∧ (r', l') ∉ E (Suc j)" by auto
        define u' where "u' ≡ if (l', r') ∈ (E j) ∧ (l', r') ∉ (E (Suc j)) then l' else r'"
        define v' where "v' ≡ if (l', r') ∈ (E j) ∧ (l', r') ∉ (E (Suc j)) then r' else l'"
        from lr_cases have uv:"(u', v') ∈ (E j)" "(u', v') ∉ E (Suc j)" by (auto simp: u'_def v'_def)
        have uv_cases:"(u' = l' ∧ v' = r') ∨ (u' = r' ∧ v' = l')" by (auto simp: u'_def v'_def)
        { fix u v i
          assume "(u, v) ∈ encstep1 (E i) (R i)"
          hence "(u, v) ∈ ostep (E i) (R i)" using encstep1_ostep by simp
          with ostep_mono[of "E i" "E" "R i" "R"] have "(u, v) ∈ ostep E R" by auto
          with ostep_iff_ordstep[OF Rinf_less] have "(u, v) ∈ ordstep {≻} (R ∪ E)" by blast
          with ordstep_subset_S[OF Rinf_less] have "(u, v) ∈ ordstep {≻} (𝒮 R E)" by auto
          with ordstep_S_conv[OF Rinf_less] have "(u, v) ∈ rstep (𝒮 R E)" by auto
        } note encstep1_rstep_Sinf = this

        let ?enc = "encstep1 (E (Suc j)) (R (Suc j)) O (E (Suc j))"
        from uv oKBi_E_supset' [OF irun [of j]]
        have "(u', v') ∈ R (Suc j) ∨ (v', u') ∈ R (Suc j) ∨ u' = v' ∨ (u', v') ∈ ?enc" by auto
        with uv_cases consider "(l', r') ∈ R (Suc j)" | "(r', l') ∈ R (Suc j)" | "l' = r'" |
          "(l', r') ∈ ?enc" | "(r', l') ∈ ?enc" by auto
        thus ?thesis proof(cases)
          case 1
          then show ?thesis unfolding s by fast
        next
          case 2
          with Ri_less sk_run.compat_h have "sk.h r' ≻sk sk.h l'" by auto
          with st sk_run.subst_h[of _ _ "τ ∘s σ"] have "sk t ≻sk sk s" unfolding Hole by auto
          from sk_run.trans_h[OF sk_less this] SN_imp_acyclic[OF sk_run.SN_less_h] show ?thesis
            unfolding acyclic_def by auto
        next
          case 3
          with st sk_less sk_run.subst_h[of _ _ "τ ∘s σ"] have "sk s ≻sk sk s" unfolding Hole by auto
          with SN_imp_acyclic[OF sk_run.SN_less_h] show ?thesis unfolding acyclic_def by auto
        next
          case 4
          then obtain u where "(l', u) ∈ encstep1 (E (Suc j)) (R (Suc j))" by auto
          from encstep1_rstep_Sinf[OF this(1)] show ?thesis unfolding s by auto
        next
          case 5
          then obtain u where u:"(r', u) ∈ encstep1 (E (Suc j)) (R (Suc j))" "(u,l') ∈ (E (Suc j))" by auto
          from u encstep1_less[OF _ Ri_less] have less:"r' ≻ u" by auto
          with subst have "r' ⋅ ρ ≻ u ⋅ ρ" by auto
          with less_imp_less_sk have 1:"sk (r' ⋅ ρ) ≻sk sk (u ⋅ ρ)" by auto
          with sk_less sk_run.trans_h have 3:"sk s ≻sk sk (u ⋅ ρ)" unfolding t by fast
          from u have "(l', u) ∈ E ∪ R" by auto
          hence 4:"(sk.h l', sk.h u) ∈ Eh ∪ Rh" using sk.rule_h[of l' u "E ∪ R"]
            using sk.hR_sym sk.hR_union sk_run.hEinf sk_run.hRinf by metis
          { fix x
            assume x:"x ∈ vars_term (sk.h u)"
            from rstep_subset_less SN_less[THEN SN_subset] less have "SN (rstep {(r',u)})" by force
            from x SN_imp_variable_condition[OF this] sk_vs have "?σ x = sk (deskolemize (?σ x))" by auto
            hence "?σ x = ((sk.hs (deskolemize ∘ ?σ)) ∘s σsk) x" unfolding sk_def subst_compose_def comp_def by simp
          }
          hence sku:"sk (u ⋅ ρ) = sk.h u ⋅ ?σ" unfolding sk_def sk.subst_apply_h ρ_def
            using term_subst_eq_conv[of "sk.h u" ] by force
          have "(sk s, sk (u ⋅ ρ)) ∈ rstep 𝒮sk" proof (cases "(sk.h l', sk.h u) ∈ Eh")
            case True
            with 3 have "(sk s, sk (u ⋅ ρ)) ∈ E_ord (op ≻sk) Eh"
              unfolding st Hole sku ctxt_apply_term.simps E_ord_def by fast
            then show ?thesis by auto
          next
            case False
            with 4 have "(sk.h l', sk.h u) ∈ Rh" by blast
            then show ?thesis unfolding st Hole sku ctxt_apply_term.simps by fast
          qed
          with 1 ind_step(1) show ?thesis unfolding t by auto
        qed
      qed
    qed
  next
    case (More f bef D aft)
    from st have sk_s:"sk s = Fun f (bef @ D⟨?l ⋅ ?σ⟩ # aft)" and sk_t:"sk t = Fun f (bef @ D⟨?r ⋅ ?σ⟩ # aft)"
      unfolding More ctxt_apply_term.simps by auto
    then obtain g where f:"f = FOrig g"  unfolding sk_def σsk_def by (cases s, auto)
    let ?p = "PCons (length bef) ε"
    let ?s = "s |_ ?p" and ?t = "t |_ ?p"
    from sk_C[of s "More f bef Hole aft"] sk_s have p:"?p ∈ poss s" and s:"sk ?s = D⟨?l ⋅ ?σ⟩" by auto
    from sk_C[of t "More f bef Hole aft"] sk_t have t:"sk ?t = D⟨?r ⋅ ?σ⟩" by auto
    from * have step:"(sk ?s, sk ?t) ∈ rstep 𝒮sk" unfolding s t using lr by auto
    from ground_subt_less[OF ground_sk, of t, of "D⟨?r⋅?σ⟩"] have "sk t ≻sk D⟨?r ⋅ ?σ⟩" unfolding st More by auto
    with ind_step(1)[OF _ step] t have "?s ∉ NF_trs (𝒮 R E)" unfolding t by auto
    with NF_subterm[OF _ subt_at_imp_supteq[OF p]] show ?thesis by auto
  qed
qed

lemma NF_Sw_subset_NF_R:"NF_trs Sω ⊆ NF_trs ℛ"
proof-
  { fix s u
    assume su:"(s,u) ∈ rstep ℛ"
    with conv_eq oKBi_conversion_ERw have "(s, u) ∈ (rstep (Rω ∪ Eω))*" by auto
    from rsteps_sk[OF this[unfolded conversion_def rstep_simps(5)[symmetric]]]
    have "(sk s, sk u) ∈ (rstep (sk.hR (Rω ∪ Eω)))*"
      unfolding conversion_def sk.hR_sym rstep_simps(5) by meson
    hence "(sk s, sk u) ∈ (rstep (Rhω ∪ Ehω))*" unfolding sk_run.hRw sk_run.hEw sk.hR_union by auto
    note conv_ground = gterm_conv_GROUND_conv[OF ground_sk ground_sk this]
    with ground_RE_conv_S_conv have gconv:"(sk s, sk u) ∈ (GROUND (rstep Shω))*" by simp
    with correctness_okb_sk[THEN conjunct1] CR_imp_conversionIff_join
    have join:"(sk s, sk u) ∈ (GROUND (rstep Shω))" by auto
    hence join:"(sk s, sk u) ∈ (rstep Shω)" using  GROUND_subset[of "rstep Shω"] join_mono by blast
    then obtain t where t:"(sk s, t) ∈ (rstep Shω)*" "(sk u, t) ∈ (rstep Shω)*" by auto
    { assume eq:"sk s = t"
      with t(2) have geq:"sk u ≽sk sk s" unfolding rtrancl_eq_or_trancl
        using sk_run.okb_h.compatible_rstep_trancl_imp_less sk_run.irun_h.Sw_less by auto
      from compatible_rstep_imp_less[OF R_less su] less_imp_less_sk have "sk s ≻sk sk u" by auto
      with sk_run.trans_h[OF this] geq sk_run.okb_h.irrefl have False by auto
    }
    with t(1) tranclD obtain v' where "(sk s, v') ∈ (rstep Shω)" unfolding rtrancl_eq_or_trancl by metis
    with sk_run.irun_h.Sw_subset_Sinf rstep_mono have "(sk s, v') ∈ (rstep 𝒮sk)" by blast
    with sk_step_exists obtain v where "(sk s, sk v) ∈ (rstep 𝒮sk)" by blast
    from sk_Swinf_step_imp_no_Sinf_NF[OF this] NF_Sw_NF_Sinf have "s ∉ NF_trs Sω" by auto
  }
  thus ?thesis by auto
qed

lemma completeness:
  assumes Ew_irreducible:"∀s t. (s,t) ∈ Eω ⟶ s ∈ NF_trs Sω"
  and Ew_nontrivial:"∀t. (t,t) ∉ Eω"
  and "reduced ℛ" and "reduced Rω"
  shows "Eω = {} ∧ ℛ ≐ Rω"
proof-
  { fix s t
    assume st:"(s,t) ∈ Eω"
    with conv_eq[unfolded oKBi_conversion_ERw] have "(s,t) ∈ (rstep ℛ)*" by auto
    with CR_imp_conversionIff_join[OF CR_R] obtain u where
      u:"(s, u) ∈ (rstep ℛ)*" "(t, u) ∈ (rstep ℛ)*" by auto
    from st Ew_nontrivial sk_inj have alts:"u ≠ s ∨ u ≠ t" by auto
    { assume "s ≠ u"
      with u(1) tranclD obtain v where "(s, v) ∈ rstep ℛ" unfolding rtrancl_eq_or_trancl by metis
      with NF_Sw_subset_NF_R have "s ∉ NF_trs Sω" by auto
      with st Ew_irreducible have False by auto
    } note s_neq_u = this
    { assume "t ≠ u"
      with u(2) tranclD obtain v where "(t, v) ∈ rstep ℛ" unfolding rtrancl_eq_or_trancl by metis
      with NF_Sw_subset_NF_R have "t ∉ NF_trs Sω" by auto
      with st Ew_irreducible have False by auto
    }
    with alts s_neq_u have False by argo
  }
  hence Ew_empty:"Eω = {}" by auto
  have Sw_Rw:"Sω = Rω" using Ew_empty unfolding Sw_def by auto
  note SN_Rw = SN_Sw_step[unfolded Sw_Rw]
  note vars = SN_imp_variable_condition[OF SN_R] SN_imp_variable_condition[OF SN_Rw]
  interpret complete_ars "(rstep ℛ)" by (unfold_locales, insert CR_R SN_R, auto)
  from conv_eq[unfolded oKBi_conversion_ERw] have in_Rconv:"rstep Rω ⊆ (rstep ℛ)*" by fastforce
  from complete_NE_intro1[OF in_Rconv SN_Rw] have NE:" (rstep ℛ)! = (rstep Rω)!"
    using NF_Sw_subset_NF_R unfolding Sw_Rw by auto
  with assms reduced_NE_imp_unique[OF vars] have "ℛ ≐ Rω" by auto
  with Ew_empty show ?thesis by auto
qed


lemma Ew_empty_imp_CR_Rw:
  assumes "Eω = {}"
  shows "CR (rstep Rω)"
proof-
  from assms Sw_def sk_run.hRw have SR[simp]:"Sω = Rω" by auto
  have Eh:"Ehω = {}" unfolding sk_run.hEw[symmetric] assms
    unfolding sk.hR_def map_funs_trs.simps map_funs_rule.simps by simp
  from sk_run.irun_h.correctness_okb[OF sk_run_fair] have GCR:"GCR (rstep Rhω)"
    unfolding sk_run.irun_h.Sw_def Eh by auto
  { fix s t u
    assume "(s,t) ∈ (rstep Rω)*" and "(s,u) ∈ (rstep Rω)*"
    with rsteps_sk[of _ _ "Rω"] have *:"(sk s,sk t) ∈ (rstep Rhω)*" and "(sk s,sk u) ∈ (rstep Rhω)*"
      using sk_run.hRw by auto
    hence "(sk s,sk t) ∈ (rstep Rhω)*" and "(sk s,sk u) ∈ (rstep Rhω)*" by auto
    hence "(sk t, sk u) ∈ (rstep Rhω)*"
      by (metis (no_types, lifting) conversion_def conversion_inv rtrancl_trans)
    with gterm_conv_GROUND_conv ground_sk have *:"(sk t, sk u) ∈ (GROUND(rstep Rhω))*" by blast
    with GCR have "(sk t,sk u) ∈ (GROUND (rstep Rhω))" unfolding CR_iff_conversion_imp_join by auto
    then obtain v where v:"(sk t,v) ∈ (GROUND (rstep Rhω))*" "(sk u,v) ∈ (GROUND (rstep Rhω))*"
      by auto
    let ?v = "deskolemize v"
    from v rtrancl_mono[OF GROUND_subset] have *:"(sk t,v) ∈ (rstep Rhω)*" "(sk u,v) ∈ (rstep Rhω)*" by auto
    note d = rstep_imp_deskolemize_rsteps[of _ _ Rω, unfolded sk_run.hRw] 
    from d[OF *(1)] d[OF *(2)] have "(t,?v) ∈ (rstep Rω)*" and "(u,?v) ∈ (rstep Rω)*" by auto
    with rtrancl_converse have "(t,u) ∈ (rstep Rω)" unfolding join_def by blast
  }
  thus ?thesis unfolding CR_on_def by auto
qed

end
end
end

lemma linear_term_replace:
  assumes "linear_term C⟨t⟩" and "vars_term s ⊆ vars_term t" and "linear_term s"
  shows "linear_term C⟨s⟩"
  using assms(1)
proof(induct C)
  case Hole
  then show ?case using assms by auto
next
  case (More f bef C aft)
  from More(2) have us:"⋀u. u ∈ set (bef @ C⟨t⟩ # aft) ⟹ linear_term u" 
    unfolding ctxt_apply_term.simps linear_term.simps by auto
  hence "linear_term C⟨t⟩" by simp
  from More(1)[OF this] us have lin:"⋀u. u ∈ set (bef @ C⟨s⟩ # aft) ⟹ linear_term u" by auto
  from assms(2) have subset:"vars_term C⟨s⟩ ⊆ vars_term C⟨t⟩" unfolding vars_term_ctxt_apply by auto
  from More(2) have p:"is_partition (map vars_term (bef @ C⟨t⟩ # aft))" 
    unfolding ctxt_apply_term.simps linear_term.simps by auto
  hence "is_partition (map vars_term (bef @ C⟨s⟩ # aft))" proof(induct bef)
    case Nil
    with subset show ?case unfolding append.left_neutral list.map(2) is_partition_Cons by auto
  next
    case (Cons u bef)
    let ?Vt = "⋃set (map vars_term (bef @ C⟨t⟩ # aft))"
    let ?Vs = "⋃set (map vars_term (bef @ C⟨s⟩ # aft))"
    from Cons(2) have "is_partition (map vars_term (bef @ C⟨t⟩ # aft))" and vs:"vars_term u ∩ ?Vt = {}"
      unfolding append.append_Cons list.map(2) is_partition_Cons by auto
    with Cons have p:"is_partition (map vars_term (bef @ C⟨s⟩ # aft))" by auto
    from subset vs have "vars_term u ∩ ?Vs = {}" by auto
    with p show ?case unfolding append.append_Cons list.map(2) is_partition_Cons by blast
  qed
  with lin show ?case unfolding ctxt_apply_term.simps linear_term.simps by auto
qed

lemma linear_term_unique_var_pos:
  assumes "linear_term t" and "x ∈ vars_term t"
  shows "∃!p. p ∈ poss t ∧ t |_ p = Var x"
proof-
  from assms(2) vars_term_poss_subt_at obtain p where p:"p ∈ poss t" "t |_ p = Var x" by metis
  { fix q
    assume q:"q ≠ p" "q ∈ poss t" "t |_ q = Var x"
    with p q have "parallel_pos p q"
      by (metis Position.self_append_conv parallel_pos prefix_pos_diff var_pos_maximal)
    from parallel_pos_prefix[OF this] obtain pq ip iq p' q' where
      *:"p = pq <#> ip <# p'" "q = pq <#> iq <# q'" "ip ≠ iq" by auto
    let ?t = "t |_ pq"
    from p have pq:"pq ∈ poss t" unfolding * by auto
    with p(1) q(2) obtain f ts where t:"?t = Fun f ts" unfolding * by (cases ?t, auto)
    from p(1) q(2) subt_at_append[OF pq] have poss:"ip<#p' ∈ poss ?t" "iq<#q' ∈ poss ?t"
      unfolding * by auto
    from poss have ils:"ip < length ts" "iq < length ts" unfolding t * by auto
    with poss poss_PCons_poss have poss'':"p' ∈ poss (ts ! ip)" "q' ∈ poss (ts ! iq)"
      unfolding * t by auto

    from subt_at_append[OF pq] p(2) have "?t |_ (PCons ip p') = Var x" unfolding * by auto
    with vars_term_subt_at poss'' have x1:"x ∈ vars_term (ts ! ip)" unfolding t by fastforce
    from subt_at_append[OF pq] q(3) have "?t |_ (PCons iq q') = Var x" unfolding * by auto
    with vars_term_subt_at poss'' have "x ∈ vars_term (ts ! iq)" unfolding t by fastforce
    with x1 have "vars_term (ts ! ip) ∩ vars_term (ts ! iq) ≠ {}" by auto

    with *(3) ils nth_map[OF ils(1)] nth_map[OF ils(2)] have "¬ (is_partition (map vars_term ts))"
      unfolding t length_map is_partition_alt is_partition_alt_def by metis
    hence "¬ linear_term ?t" unfolding t linear_term.simps by auto
    hence False using subterm_linear assms subt_at_imp_supteq[OF pq] by auto
  }
  with p show ?thesis by metis
qed


lemma linear_term_subst:
  assumes "linear_term l" and "linear_term r" and "vars_term r ⊆ vars_term l" and "linear_term (l⋅σ)"
  shows "linear_term (r⋅σ)"
  using assms(2) assms(3)
proof(induct r)
  case (Var x)
  hence "l ⊵ Var x" by auto
  with subterm_linear[OF _ assms(4), of "σ x"] show ?case by force
next
  case (Fun f rs)
  note IH = this
  hence p:"is_partition (map vars_term rs)" and lin_ri:"⋀ri. ri ∈ (set rs) ⟹ linear_term ri"
    unfolding linear_term.simps by auto
  { fix riσ
    assume "riσ ∈ set (map (λt. t⋅σ) rs)"
    then obtain ri where ri:"ri ∈ set rs" "riσ = ri⋅σ" by force
    with Fun(3) have "vars_term ri ⊆ vars_term l" by auto
    with ri Fun(1)[OF ri(1) lin_ri[OF ri(1)]] have "linear_term riσ" by simp
  }
  note linear_subterms = this
  { assume "¬ is_partition (map vars_term (map (λt. t⋅σ) rs))"
    from this[unfolded is_partition_alt is_partition_alt_def map_map length_map] obtain i j where
        i:"i < length rs" and j:"j < length rs" and ij:"i ≠ j" and
        inter:"vars_term ((rs ! i)⋅σ) ∩ vars_term ((rs ! j)⋅σ) ≠ {}" by auto
      from inter obtain x where x:"x ∈ vars_term ((rs ! i)⋅σ)" "x ∈ vars_term ((rs ! j)⋅σ)" by auto
      then obtain y z where y:"y ∈ vars_term (rs ! i)" "x ∈ vars_term (σ y)" and
                z:"z ∈ vars_term (rs ! j)" "x ∈ vars_term (σ z)" unfolding vars_term_subst by auto
      from p i j ij y(1) z(1) have yz:"y ≠ z" unfolding is_partition_alt is_partition_alt_def by auto
      from i j y(1) z(1) IH(3) have yl:"y ∈ vars_term l" and zl:"z ∈ vars_term l"
        unfolding term.set(4) by fastforce+
      note vsub = vars_term_poss_subt_at
      from vsub[OF yl] obtain p where p:"p ∈ poss l" "l |_ p = Var y" by auto
      from vsub[OF zl] obtain q where q:"q ∈ poss l" "l |_ q = Var z" by auto
      with p yz have neq:"p ≠ q" by auto
      with p q have par:"parallel_pos p q"
        by (metis Position.self_append_conv parallel_pos prefix_pos_diff var_pos_maximal)
      from vsub[OF y(2)] obtain p' where p':"p' ∈ poss (σ y)" "(σ y) |_ p' = Var x" by auto
      from vsub[OF z(2)] obtain q' where q':"q' ∈ poss (σ z)" "(σ z) |_ q' = Var x" by auto
      from par have neq:"p <#> p' ≠ q <#> q'"
        by (metis less_eq_simps(1) pos_less_eq_append_not_parallel)
      from subt_at_append p p' have x1:"l⋅σ |_ (p <#> p') = Var x" by simp
      from subt_at_append q q' have x2:"l⋅σ |_ (q <#> q') = Var x" by simp
      from p p'(1) have pl:"p <#> p' ∈ poss (l⋅σ)" unfolding poss_append_poss by auto
      from q q'(1) have ql:"q <#> q' ∈ poss (l⋅σ)" unfolding poss_append_poss by auto
      from vars_term_subt_at[OF this] x2 have "x ∈ vars_term (l⋅σ)" by auto
      from linear_term_unique_var_pos[OF assms(4) this] x1 x2 pl ql neq have False by blast
  }
  with linear_subterms show ?case by auto
qed

locale linear_ordered_completion1 = 
  ordered_completion less for less :: "('a, 'b::infinite) term ⇒ ('a, 'b) term ⇒ bool" (infix "≻" 50)
begin

text ‹A linear critical pair as defined by Devie.›
definition "lin_overlap R l r p l' r' σ ≡
    (∃p. p ∙ (l,r) ∈ R) ∧
    (∃p. p ∙ (l',r') ∈ R) ∧
    vars_rule (l,r) ∩ vars_rule (l',r') = {} ∧
    p ∈ funposs l' ∧
    mgu l (l' |_ p) = Some σ ∧
    ((l ≻ r ∧ ¬ (r' ≻ l')) ∨ (l' ≻ r' ∧ ¬ (r ≻ l)))"

definition "LCP R =
  {(replace_at l' p r ⋅ σ, r' ⋅ σ) | l r p l' r' σ. lin_overlap R l r p l' r' σ}"

inductive
  oKBilin :: "('a, 'b) trs × ('a, 'b) trs ⇒ ('a, 'b) trs × ('a, 'b) trs ⇒ bool" (infix "⊢oKBlin" 55)
  where
    deduce: "(s, t) ∈ rstep (R ∪ E) ⟹ (s, u) ∈ rstep (R ∪ E) ⟹ linear_term t ⟹ linear_term u ⟹ (E, R)⊢oKBlin (E ∪ {(t, u)}, R)" |
    orientl: "s ≻ t ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKBlin (E - {(s, t)}, R ∪ {(s, t)})" |
    orientr: "t ≻ s ⟹ (s, t) ∈ E ⟹ (E, R)⊢oKBlin (E - {(s, t)}, R ∪ {(t, s)})" |
    delete: "(s, s) ∈ E ⟹ (E, R) ⊢oKBlin (E - {(s, s)}, R)" |
    compose: "(t, u) ∈ rstep (R - {(s, t)}) ⟹ (s, t) ∈ R ⟹ (E, R) ⊢oKBlin (E, (R - {(s, t)}) ∪ {(s, u)})" |
    simplifyl: "(s, u) ∈ rstep R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKBlin ((E - {(s, t)}) ∪ {(u, t)}, R)" |
    simplifyr: "(t, u) ∈ rstep R ⟹ (s, t) ∈ E ⟹ (E, R) ⊢oKBlin ((E - {(s, t)}) ∪ {(s, u)}, R)" |
    collapse: "(t, u) ∈ encstep2 {} (R - {(t, s)}) ⟹ (t, s) ∈ R ⟹ (E, R) ⊢oKBlin (E ∪ {(u, s)}, R - {(t, s)})"

lemma oKBilin_step_imp_oKBi_step:
  assumes "(E, R) ⊢oKBlin (E', R')"
  shows "(E, R) ⊢oKB (E', R')"
  using assms
proof(cases)
  case (deduce s t u)
  then show ?thesis using oKBi.deduce by auto
next
  case (orientl s t)
  then show ?thesis using oKBi.orientl by auto
next
  case (orientr t s)
  then show ?thesis using oKBi.orientr by auto
next
  case (compose t u s)
  then show ?thesis using oKBi.compose rstep_imp_ostep[OF compose(3)] by auto
next
  case (simplifyl s u t)
  then show ?thesis using oKBi.simplifyl encstep1.rstep[OF simplifyl(3)] by auto
next
  case (simplifyr t u s)
  then show ?thesis using oKBi.simplifyr encstep1.rstep[OF simplifyr(3)] by auto
next
  case (collapse t u s)
  from collapse(3) have "(t, u) ∈ encstep2 E (R - {(t, s)})" using encstep2_mono[of "{}"] by auto
  with collapse show ?thesis using oKBi.collapse by auto
qed (auto intro!:oKBi.delete oKBi.orientl)

lemma oKBilin_less:
  assumes "(E, R)  ⊢oKBlin (E', R')" and "R ⊆ {≻}"
  shows "R' ⊆ {≻}"
  using oKBi_less[OF assms(2) oKBilin_step_imp_oKBi_step] assms by auto

lemma oKBilin_R_supset:
  assumes "(E, R) ⊢oKBlin (E', R')"
  shows "R - R' ⊆ encstep2 {} R' O E' ∪ R' O (rstep R')¯"
  using assms
proof (cases)
  case (compose s t u)
  then have "(s, t) ∈ rstep R'" by auto
  with compose show ?thesis by fast
next
  case (collapse s t u)
  with encstep2_mono [of E "E ∪ {(t, u)}"] show ?thesis by fast
qed auto

lemma linear_term_step:
  assumes "linear_term t" and "linear_trs R" and "R ⊆ {≻}" and "(t,s) ∈ rstep R"
  shows "linear_term s"
proof-
  have "rstep {≻} ⊆ {≻}" using rstep_subset rstep_rstep subst ctxt by fastforce
  hence SN:"SN (rstep {≻})" using SN_less SN_subset by auto
  note vc = SN_imp_variable_condition[OF SN]
  { assume "(t,s) ∈ rstep R"
    then obtain l r C σ where lr:"(l,r) ∈ R" and st:"t = C⟨l⋅σ⟩" "s = C⟨r⋅σ⟩" by auto
    from lr assms(3) vc have vs_rule:"vars_term r ⊆ vars_term l" by auto
    from var_cond_stable[OF this] have vars':"vars_term (r⋅σ) ⊆ vars_term (l⋅σ)" by fast
    from assms(1) subterm_linear[of t "l⋅σ"] have lin_ls:"linear_term (l⋅σ)" unfolding st by auto
    from lr linear_trsE[OF assms(2)] have "linear_term l" "linear_term r" by auto
    from linear_term_subst[OF this vs_rule lin_ls] have "linear_term (r⋅σ)" by auto 
    with linear_term_replace[OF assms(1)[unfolded st] vars'] have "linear_term s" unfolding st by auto
  }
  with assms(4) show ?thesis by auto
qed

lemma linear_trs_remove_linear[simp]: "linear_trs R ⟹ linear_trs (R - R')"
  unfolding linear_trs_def by auto

lemma linear_trs_union_linear[simp]: "linear_trs R ⟹ linear_trs R' ⟹ linear_trs (R ∪ R')"
  unfolding linear_trs_def by auto

lemma oKBilin_step_linearity_preserving:
  assumes "(E, R) ⊢oKBlin (E', R')" and R_less:"R ⊆ {≻}" and lin:"linear_trs R" "linear_trs E"
  shows "linear_trs R' ∧ linear_trs E'"
proof-
  from ordstep_imp_rstep have ostep_rstep:"⋀E R. ostep E R ⊆ rstep (R ∪ E)" unfolding ostep_def by fast
  have lin_add[simp]:"⋀R s t. linear_trs R ⟹ linear_term s ⟹ linear_term t ⟹ linear_trs (R ∪ {(s,t)})"
    unfolding linear_trs_def by auto
  show ?thesis using assms(1)
  proof(cases)
    case (deduce s t u)
    with lin show ?thesis unfolding deduce unfolding linear_trs_def by fast
  next
    case (orientl s t)
    then show ?thesis using assms unfolding linear_trs_def by auto
  next
    case (orientr t s)
    then show ?thesis using assms unfolding linear_trs_def by auto
  next
    case (delete s)
    then show ?thesis using assms unfolding linear_trs_def by auto
  next
    case (compose t u s)
    from compose(4) lin have t:"linear_term t" and s:"linear_term s" unfolding linear_trs_def by auto
    from lin have "linear_trs (R - {(s, t)})" by auto
    from linear_term_step[OF t this _ compose(3)] R_less have "linear_term u" by auto
    from lin lin_add[OF _ s this] show ?thesis unfolding compose by simp
  next
    case (simplifyl s u t)
    from simplifyl(4) lin have t:"linear_term t" and s:"linear_term s" unfolding linear_trs_def by auto
    from lin have "linear_trs (E - {(s, t)})" by auto
    note step = linear_term_step[OF s lin(1) R_less simplifyl(3)]
    from lin lin_add[OF _ this t, OF linear_trs_remove_linear] show ?thesis unfolding simplifyl by auto
  next
    case (simplifyr t u s)
    from simplifyr(4) lin have t:"linear_term t" and s:"linear_term s" unfolding linear_trs_def by auto
    from lin have "linear_trs (E - {(s, t)})" by auto
    note step = linear_term_step[OF t lin(1) R_less simplifyr(3)]
    from lin lin_add[OF _ s this, OF linear_trs_remove_linear] show ?thesis unfolding simplifyr by auto
  next
    case (collapse t u s)
    from collapse(4) lin have t:"linear_term t" and s:"linear_term s" unfolding linear_trs_def by auto
    from lin have "linear_trs (E - {(s, t)})" by auto
    from encstep2_imp_rstep_sym[OF collapse(3)] have "(t, u) ∈ rstep (R - {(t, s)})" by auto
    from linear_term_step[OF t _ _ this] R_less lin have "linear_term u" by auto
    from lin lin_add[OF _ this s] show ?thesis unfolding collapse by auto
  qed
qed

lemma oKBilin_rtrancl_linear:
  assumes "oKBilin** (E, R) (E', R')" and R_less:"R ⊆ {≻}" and lin:"linear_trs R" "linear_trs E"
  shows "R' ⊆ {≻} ∧ linear_trs R' ∧ linear_trs E'"
  using assms 
proof (induct "(E, R)" "(E', R')" arbitrary: E' R')
  case rtrancl_refl
  then show ?case by auto
next
  case (rtrancl_into_rtrancl ER)
  then obtain Ei Ri where ER:"ER = (Ei,Ri)" by force
  from rtrancl_into_rtrancl have "linear_trs Ri" "linear_trs Ei" unfolding ER by auto
  with rtrancl_into_rtrancl oKBilin_step_linearity_preserving oKBilin_less show ?case unfolding ER by metis
qed

lemma oKBi_E_supset_lin:
  assumes "(E, R) ⊢oKBlin (E', R')"
  shows "E - E' ⊆ (rstep R' O E') ∪ (E' O (rstep R')¯) ∪ R' ∪ Id"
  using assms by (cases, insert rstep_mono [of R R'], auto)

end

locale okblin_irun = linear_ordered_completion1 +
  fixes R E
  assumes R0: "R 0 = {}"
    and E0_linear:"linear_trs (E 0)"
    and lin_irun: "⋀i. (E i, R i) ⊢oKBlin (E (Suc i), R (Suc i))"
begin

sublocale okb_irun less R E using R0 lin_irun oKBilin_step_imp_oKBi_step by (unfold_locales, auto)

lemma oKBilin_rtrancl_i: "oKBilin** (E 0, R 0) (E i, R i)"
  using lin_irun by (induct i, auto intro: rtranclp.rtrancl_into_rtrancl)

lemma Ri_linear:"linear_trs (R i)"
  using oKBilin_rtrancl_linear [OF oKBilin_rtrancl_i] and E0_linear and R0 by auto

lemma Ei_linear:"linear_trs (E i)"
  using oKBilin_rtrancl_linear [OF oKBilin_rtrancl_i] and E0_linear and R0 by auto

lemma Rinf_Einf_linear:"linear_trs R ∧ linear_trs E"
  using Ri_linear Ei_linear linear_trs_def by fast

lemma Rw_linear:"linear_trs Rω"
  unfolding Rω_def using Ri_linear linear_trs_def by fastforce

lemma Ew_linear:"linear_trs Eω"
  unfolding Eω_def using Ei_linear linear_trs_def by fastforce

context
  fixes ::"('a,'b) trs"
  assumes R_less:"ℛ ⊆ {≻}" and CR_R:"CR (rstep ℛ)" and "reduced ℛ"
  and conv_eq:"(rstep ℛ)* = (rstep (E 0))*"
begin

lemma SN_R:"SN (rstep ℛ)" 
  by (rule SN_subset [OF SN_less], insert compatible_rstep_imp_less[OF R_less], auto)

abbreviation "mulsucc ≡ s_mul_ext Id {≻}" 
abbreviation "mulsucceq ≡ ns_mul_ext Id {≻}"
abbreviation mulsucc_op (infix "≻mul" 55) where "mulsucc_op s t ≡ (s,t) ∈ mulsucc"
abbreviation mulsucceq_op (infix "≽mul" 55) where "mulsucceq_op s t ≡ (s,t) ∈ mulsucceq"

interpretation mulsucc: SN_order_pair mulsucc mulsucceq using trans_less SN_less trans_Id refl_Id
  by (intro SN_order_pair.mul_ext_SN_order_pair, unfold_locales, auto)

definition "E_ord_lin ℰ = {(s ⋅ σ, t ⋅ σ) | s t σ. (s, t) ∈ ℰ ∧ linear_term (s ⋅ σ) ∧ linear_term (t ⋅ σ) ∧ (s ⋅ σ) ≻ (t ⋅ σ)}"

lemma E_ord_lin_mono: "ℰ ⊆ ℰ' ⟹ E_ord_lin ℰ ⊆ E_ord_lin ℰ'" unfolding E_ord_lin_def by auto

lemma mset_two:
  assumes "s ≻ u"
  shows "{#s, t#} ≻mul {#u, t#}"
proof -
  from assms s_mul_ext_singleton have su: "{#s#} ≻mul {#u#}" by auto
  from ns_mul_ext_refl [OF refl_encompeq] have "({#t#}, {#t#}) ∈ mulsucceq" by auto
  from s_ns_mul_ext_union_compat [OF su this] show ?thesis
    using add.commute add_mset_add_single by metis
qed

lemma mset_two2:
  assumes "s ≻ u"
  shows "{#t, s#} ≻mul {#t, u#}"
  using mset_two [OF assms] add_mset_commute by (metis (no_types, lifting))

lemma Einf_to_Ew:"E ⊆ (rstep R)* O Eω= O ((rstep R)¯)*"
proof-
  { fix s t
    assume "(s,t) ∈ E"
    hence "(s,t) ∈ (rstep R)* O Eω= O ((rstep R)¯)*"
    proof (induct "{#s,t#}" arbitrary: s t rule: SN_induct [OF mulsucc.SN])
      case 1
      note IH = this
      show ?case proof(cases "(s,t) ∈ Eω", fast)
        case False
        with Einf_without_Ew[OF 1(2)] obtain i where i:"(s, t) ∈ E i" "(s, t) ∉ E (Suc i)" by auto
    have S:"𝒮 (R (Suc i)) (E (Suc i)) ⊆ 𝒮 R E" using E_ord_mono by (metis Sup_upper range_eqI sup.mono)
        let ?E = "E (Suc i)" and ?R = "R (Suc i)"
        from i oKBi_E_supset_lin[OF lin_irun] consider
          "(s, t) ∈ rstep ?R O ?E" | "(s, t) ∈ ?E O (rstep ?R)¯" | "(s, t) ∈ (?R)=" by blast
        then show ?thesis proof(cases)
          case 1
          then obtain u where su:"(s,u) ∈ rstep ?R" and ut:"(u,t) ∈ ?E" by auto
          from su rstep_mono[of ?R "R"] have step:"(s,u) ∈ rstep R" by auto
          from compatible_rstep_imp_less[OF Ri_less su] mset_two have "{#s,t#} ≻mul {#u,t#}" by auto
          with ut IH(1)[of u t] have "(u, t) ∈ (rstep R)* O Eω= O ((rstep R)¯)*" by blast
          then obtain v where "(u, v) ∈ (rstep R)*" "(v, t) ∈ Eω= O ((rstep R)¯)*" by auto
          with converse_rtrancl_into_rtrancl[OF step this(1)] show ?thesis by auto
        next
          case 2
          then obtain u where su:"(s,u) ∈ ?E" and tu:"(t,u) ∈ rstep ?R" by fast
          with rstep_mono[of ?R "R"] have step:"(t,u) ∈ rstep R" by auto
          from compatible_rstep_imp_less[OF Ri_less tu] mset_two2 have "{#s,t#} ≻mul {#s,u#}" by auto
          with su IH(1)[of s u] have "(s, u) ∈ (rstep R)* O Eω= O ((rstep R)¯)*" by blast
          then obtain v where "(v, u) ∈ ((rstep R)¯)*" "(s, v) ∈ (rstep R)* O Eω=" by auto
          with rtrancl_into_rtrancl[OF this(1), of t] step show ?thesis by blast
        next
          case 3
          have "?R ⊆ rstep R" unfolding rstep_union using rstep.rule by auto
          with 3 show ?thesis by blast
        qed
      qed
    qed
  }
  thus ?thesis ..
qed

lemma Einf_to_Ew_rstep:"rstep E ⊆ (rstep R)* O (rstep (Eω))= O ((rstep R)¯)*"
proof
  fix s t
  assume "(s,t) ∈ rstep E"
  then obtain l r C σ where lr:"(l,r) ∈ E" "s = C⟨l⋅σ⟩" "t = C⟨r⋅σ⟩" by fast
  with Einf_to_Ew have "(l,r) ∈ (rstep R)* O Eω= O ((rstep R)¯)*" by blast
  then obtain u v where uv:"(l,u) ∈ (rstep R)*" "(u,v) ∈ Eω=" "(v,r) ∈ ((rstep R)¯)*" by auto
  note c = ctxt.closedD[OF ctxt_closed_rsteps subst.closedD[OF subst_closed_rsteps]]
  with uv(1) have 1:"(C⟨l⋅σ⟩,C⟨u⋅σ⟩) ∈ (rstep R)*" by blast
  from uv(2) have 2:"(C⟨u⋅σ⟩,C⟨v⋅σ⟩) ∈ (rstep (Eω))=" by auto
  from c uv(3) have 3:"(C⟨v⋅σ⟩,C⟨r⋅σ⟩) ∈ ((rstep R)¯)*" unfolding reverseTrs by blast
  with 1 2 show "(s,t) ∈ (rstep R)* O (rstep (Eω))= O ((rstep R)¯)*" unfolding lr by auto
qed

lemma rsteps_R_NF_conv:
  assumes "(s, t) ∈ (rstep ℛ)!" and "(s, u) ∈ (rstep (Rω ∪ Eω))*"
  shows "(u, t) ∈ (rstep ℛ)!"
proof-
  from assms conv_eq[unfolded oKBi_conversion_ERw] conversion_sym have us:"(u, s) ∈ (rstep ℛ)*"
    unfolding sym_def by metis
  from assms(1) have "(s, t) ∈ (rstep ℛ)*" unfolding normalizability_def using conversionI' by blast
  with us have "(u, t) ∈ (rstep ℛ)*" unfolding conversion_def by auto
  from CR_NF_conv[OF CR_R _ this] assms show ?thesis unfolding normalizability_def by auto
qed

lemma LCP_exists:
  fixes σ1 :: "'b ⇒ ('a, 'b) term" 
  assumes "C⟨l1⋅σ1⟩ = l2'⋅σ2'" and "hole_pos C ∈ funposs l2'" and "(l1,r1) ∈ RR" and "(l2',r2') ∈ RR"
  and "(l2' ≻ r2' ∧ ¬(r1 ≻ l1)) ∨ (l1 ≻ r1 ∧ ¬(r2' ≻ l2'))"
  shows "∃τ s t. (s,t) ∈ LCP RR ∧ C⟨r1 ⋅ σ1⟩ = s ⋅ τ ∧ r2' ⋅ σ2' = t ⋅ τ"
proof-
  have rl1:"∃p. p ∙ (l1, r1) ∈ RR" by (rule exI[of _ 0], insert assms(3), auto)
  from vars_rule_disjoint obtain π
    where π: "vars_rule (π ∙ (l2', r2')) ∩ vars_rule (l1, r1) = {}" ..
  define l2 and r2 and σ2
    where "l2 = π ∙ l2'" and "r2 = π ∙ r2'" and 2 = (Var ∘ Rep_perm (-π)) ∘s σ2'"
  note rename = l2_def r2_def σ2_def
  from assms(4) have "-π ∙ (l2,r2) ∈ RR" unfolding rename by auto
  hence rl2:"∃p. p ∙ (l2,r2) ∈ RR" by (rule exI[of _ "-π"])
  from π have vars:"vars_rule (l1,r1) ∩ vars_rule (l2, r2) = {}"
    unfolding rename rule_pt.permute_prod_eqvt by auto
  define p where "p ≡ hole_pos C" 
  from assms(2) have p:"p ∈ funposs l2" by (auto simp:rename p_def)
  note poss_p = funposs_imp_poss [OF p]
    
  have r_l1:"l2' = -π ∙ l2" by (auto simp:rename)
  from assms(1) have eq:"l1 ⋅ σ1 = (l2' ⋅ σ2' |_ p)" unfolding p_def by (metis subt_at_hole_pos)
  note subt_l1 = subt_at_subst[OF poss_p]
  from eq[unfolded r_l1 permute_term_subst_apply_term] have eq':"l1 ⋅ σ1 = l2 |_ p ⋅ σ2"
    unfolding subt_at_subst[OF poss_p] rename using ‹l2 ≡ π ∙ l2'› eq r_l1 subt_l1
    by (metis (full_types) subst_subst_compose term_apply_subst_Var_Rep_perm)
      
  note coinc = coincidence_lemma' [of _ "vars_rule (l1, r1)"]
  define σ where "σ x = (if x ∈ vars_rule (l1, r1) then σ1 x else σ2 x)" for x
  have "l1 ⋅ σ = l1 ⋅ (σ |s vars_rule (l1, r1))"
    using coinc[of l1] by (simp add: vars_rule_def)
  also have "… = l1 ⋅ (σ1 |s vars_rule (l1, r1))" by (simp add: σ_def [abs_def])
  finally have l1_coinc:"l1 ⋅ σ = l1 ⋅ σ1" using coinc[of l1] by (simp add: vars_rule_def)
  have disj: "vars_rule (l1, r1) ∩ vars_term l2 = {}"
    using π[unfolded rename] by (auto simp: vars_rule_def rule_pt.permute_prod_eqvt l2_def)
  have unif: "l1 ⋅ σ = (l2 |_ p) ⋅ σ"
  proof -
    from disj have disj: "vars_rule (l1, r1) ∩ vars_term (l2 |_ p) = {}" 
      using vars_term_subt_at [OF poss_p] by auto
    from l1_coinc have "l1 ⋅ σ = (l2 |_ p) ⋅ σ2" using eq' by simp
    also have "… = (l2 |_ p) ⋅ (σ2 |s vars_term (l2 |_ p))"
      by (simp add: coincidence_lemma [symmetric])
    also have "… = (l2 |_ p) ⋅ (σ |s vars_term (l2 |_ p))" using disj by (simp add: σ_def [abs_def])
    finally show ?thesis by (simp add: coincidence_lemma [symmetric])
  qed
    
  define μ where "μ = the_mgu l1 (l2 |_ p)"
  have is_mgu:"is_mgu μ {(l1, l2 |_ p)}" by (rule is_mguI, insert unif the_mgu, auto simp: μ_def)
  with unif obtain τ where σ: "σ = μ ∘s τ" by (auto simp: is_mgu_def unifiers_def)
  from unif have mgu: "mgu l1 (l2 |_ p) = Some μ"
    unfolding μ_def the_mgu_def
    using unify_complete and unify_sound by (force split: option.splits simp: is_imgu_def unifiers_def)
      
  have "r1 ⋅ σ = r1 ⋅ (σ |s vars_rule (l1, r1))" using coinc[of r1] by (simp add: vars_rule_def)
  also have "… = r1 ⋅ (σ1 |s vars_rule (l1, r1))" by (simp add: σ_def [abs_def])
  finally have r1_coinc:"r1 ⋅ σ = r1 ⋅ σ1" using coinc[of r1] by (simp add: vars_rule_def)
      
  note coinc = coincidence_lemma [symmetric]
  have "l2 ⋅ σ = l2 ⋅ (σ |s vars_term l2)" by (simp add: coinc)
  also have "… = l2 ⋅ (σ2 |s vars_term l2)" using disj by (simp add: σ_def [abs_def])
  also have "… = l2 ⋅ σ2" by (simp add: coinc)
  finally have l2_coinc:"l2 ⋅ σ = l2' ⋅ σ2'" unfolding rename by simp
  have disj: "vars_rule (l1, r1) ∩ vars_term r2 = {}"
    using vars_term_subt_at [OF poss_p] and π[unfolded rename]
    by (auto simp: vars_rule_def rule_pt.permute_prod_eqvt r2_def)
  have "r2 ⋅ σ = r2 ⋅ (σ |s vars_term r2)" by (simp add: coinc)
  also have "… = r2 ⋅ (σ2 |s vars_term r2)" using disj by (simp add: σ_def [abs_def])
  also have "… = r2 ⋅ σ2" by (simp add: coinc)
  finally have r2_coinc:"r2 ⋅ σ = r2' ⋅ σ2'" unfolding rename by simp

  from assms(5) have oriented:"(l1 ≻ r1 ∧ ¬ r2 ≻ l2) ∨ (l2 ≻ r2 ∧ ¬ r1 ≻ l1)" unfolding l2_def r2_def
    unfolding permute_term_subst_apply_term
    by (metis subst term_apply_subst_Var_Rep_perm term_pt.permute_minus_cancel(2))

  from assms(1) poss_p hole_pos_id_ctxt have "C = ctxt_of_pos_term p (l2 ⋅ σ)"
    unfolding l1_coinc[symmetric] l2_coinc[symmetric] p_def by metis  
  hence cp_left:"C⟨r1 ⋅ σ1⟩ = (ctxt_of_pos_term p (l2 ⋅ μ))⟨r1 ⋅ μ⟩ ⋅ τ"
    unfolding subst_apply_term_ctxt_apply_distrib r1_coinc[symmetric]
      ctxt_of_pos_term_subst[OF poss_imp_subst_poss[OF poss_p], symmetric]
      subst_subst_compose[symmetric] σ[symmetric] by simp
  from subst_subst_compose have cp_right:"r2' ⋅ σ2' = r2 ⋅ μ ⋅ τ" unfolding r2_coinc[symmetric] σ by auto
  let ?s = "(ctxt_of_pos_term p (l2 ⋅ μ))⟨r1 ⋅ μ⟩"
  let ?t = "r2 ⋅ μ"
  from l1_coinc σ have eq:"l1 ⋅ σ1 = l1 ⋅ μ ⋅ τ" by auto
  note facts = rl1 rl2 vars p mgu oriented
  have o:"lin_overlap RR l1 r1 p l2 r2 μ" unfolding lin_overlap_def overlap_def fst_conv using facts
    unfolding the_mgu_def by simp
  show ?thesis unfolding LCP_def by (rule exI[of _ τ], rule exI[of _ ?s], rule exI[of _ ?t],
   unfold mem_Collect_eq, rule conjI, rule exI[of _ l1], rule exI[of _ r1], rule exI[of _ p],
   rule exI[of _ l2], rule exI[of _ r2], rule exI[of _ μ], insert o cp_left cp_right eq,
   unfold subst_apply_term_ctxt_apply_distrib, insert ctxt_of_pos_term_subst[OF poss_p, of μ], auto)
qed


context
  assumes fair:"∀(s, t)∈ LCP (Rω ∪ Eω). (s, t) ∈ (rstep E)"
begin

lemma PCP_LCP:
  assumes R_less:"RR ⊆ {≻}"
  shows "PCP RR ⊆ LCP RR"
proof
  fix s t
  assume "(s, t) ∈ PCP RR"
  then obtain μ l1 r1 p l2 r2 where o: "overlap RR RR (l1, r1) p (l2, r2)"
    and "the_mgu l1 (l2 |_ p) = μ"
    and s: "s = replace_at l2 p r1 ⋅ μ"
    and t: "t = r2 ⋅ μ"
    and NF: "∀u ⊲ l1 ⋅ μ. u ∈ NF (rstep RR)"
    by (auto simp: PCP_def the_mgu_def split: option.splits)
  then have mgu: "mgu l1 (l2 |_ p) = Some μ"
    by (auto simp: the_mgu_def unifiers_def overlap_def simp del: mgu.simps split: option.splits dest!: mgu_complete)
  have p: "p ∈ funposs l2" using o by (auto simp: overlap_def)
  from o have rules:"∃p. p ∙ (l1, r1) ∈ RR" "∃p. p ∙ (l2, r2) ∈ RR" "vars_rule (l1, r1) ∩ vars_rule (l2, r2) = {}"
    unfolding overlap_def by auto
  { fix l r p
    assume "p ∙ (l,r) ∈ RR" 
    hence "p ∙ l ≻ p ∙ r" using R_less unfolding rule_pt.permute_prod_eqvt by auto
    hence lr:"l ≻ r" using subst [of "p ∙ l" "p ∙ r" "sop (-p)"] by auto
    from trans[OF this] SN_imp_acyclic[OF SN_less] trans have "¬(r ≻ l)" unfolding acyclic_def by blast
    with lr have "l ≻ r ∧ ¬(r ≻ l)" by auto
  }
  with rules R_less have oriented:"l1 ≻ r1" "¬(r2 ≻ l2)" by auto
  from rules p mgu oriented
  have overlap:"lin_overlap RR l1 r1 p l2 r2 μ"
    unfolding lin_overlap_def fst_conv snd_conv s t by auto
  then show "(s, t) ∈ LCP RR" unfolding LCP_def s t by fast
qed

lemma Ew_empty_CR_Rw_linear:
  assumes "Eω = {}"
  shows "CR (rstep Rω)"
proof-
  from fair PCP_LCP[OF Rw_less] have fair:"PCP Rω ⊆ (rstep E)" unfolding assms by auto
  interpret okb_irun_nonfailing by (unfold_locales, insert fair assms, auto)
  from Ew_empty_implies_CR_Rw show ?thesis by auto
qed

lemma linear_peak_cases:
  assumes step:"(s,t) ∈ rstep_r_p_s (Rω¯ ∪ Eω) (l1,r1) p1 σ1" and Rstep:"(t,u) ∈ rstep Rω" and "¬ l1 ≻ r1"
  shows "(s,u) ∈ (rstep R)* O (rstep (Eω))= O ((rstep R)¯)*"
proof-
  from step obtain C where lr:"(l1,r1) ∈ Rω¯ ∪ Eω" "s = C⟨l1⋅σ1⟩" "t = C⟨r1⋅σ1⟩" "p1 = hole_pos C"
    unfolding rstep_r_p_s_def mem_Collect_eq split fst_conv snd_conv 
    by (metis (mono_tags, lifting) hole_pos_ctxt_of_pos_term)
  from lr Rw_subset_Rinf have lr':"(l1, r1) ∈ R¯ ∪ Eω" by auto
  from Rstep obtain D l2 r2 σ2 where lr2:"(l2,r2) ∈ Rω" "t = D⟨l2⋅σ2⟩" "u = D⟨r2⋅σ2⟩" by blast 
  let ?p2 = "hole_pos D"
  note eq = lr(3)[unfolded lr2(2), simplified]
  hence eq'[simp]: "C⟨r1 ⋅ σ1⟩ |_ ?p2 = l2 ⋅ σ2" using subt_at_hole_pos by metis
  (* do a CP Lemma-like analysis (standard CP lemma requires variable condition) *)
  consider "?p2 ≥ p1" | "p1 ≥ ?p2" | "parallel_pos p1 ?p2" using position_cases[of p1 ?p2] by auto
  thus ?thesis proof(cases)
    case 1
    with prefix_exists obtain q where q:"?p2 = p1 <#> q" by metis
    from lr(3) q have DC:"D = C ∘c (ctxt_of_pos_term q (r1 ⋅ σ1))" unfolding lr2 lr(4)
      by (metis ctxt_of_pos_term_append ctxt_of_pos_term_hole_pos hole_pos_poss subt_at_hole_pos)
    from eq[unfolded DC] have eq'':"(ctxt_of_pos_term q (r1 ⋅ σ1))⟨l2 ⋅ σ2⟩ = r1 ⋅ σ1"
      unfolding ctxt_ctxt_compose by auto
    show ?thesis proof(cases "q ∈ funposs r1")
      case False (* variable overlap *)
      note q_not_in_funposs_r1 = this
      from lr lr2(2) q have p_rσ:"q ∈ poss (r1 ⋅ σ1)" by (metis hole_pos_poss hole_pos_poss_conv)
      from poss_subst_apply_term[OF this q_not_in_funposs_r1] obtain q1 q2 x
        where p[simp]: "q = q1 <#> q2" and q1: "q1 ∈ poss r1"
          and rq1: "r1 |_ q1 = Var x" and q2: "q2 ∈ poss (σ1 x)" by auto
      from eq' have [simp]: "r1 ⋅ σ1 |_ q = l2 ⋅ σ2" using subt_at_hole_pos subt_at_append[OF hole_pos_poss]
        unfolding q lr(4) by metis
      with rq1 q2 subt_at_append p_rσ q1 have at_q2[simp]: 1 x |_ q2 = l2 ⋅σ2" unfolding p by auto

      let ?u = "replace_at (σ1 x) q2 (r2 ⋅ σ2)"
      define σ1' where 1' y = (if y = x then ?u else σ1 y)" for y
      have "(σ1 x, σ1' x) ∈ rstep Rω"
      proof -
        let ?C = "ctxt_of_pos_term q21 x)"
        have "(?C⟨l2 ⋅ σ2⟩, ?C⟨r2 ⋅ σ2⟩) ∈ rstep Rω" using lr2 by blast
        then show ?thesis using q2 by (simp add: σ1'_def replace_at_ident)
      qed
      then have *: "⋀x. (σ1 x, σ1' x) ∈ (rstep Rω)=" by (auto simp: σ1'_def)

      then have "(l1 ⋅ σ1, l1 ⋅ σ1') ∈ (rstep Rω)=" proof(cases "x ∈ vars_term l1")
        case True
        from lr(1) Ew_linear Rw_linear have lin:"linear_term l1" unfolding linear_trs_def by auto
        from linear_term_unique_var_pos[OF lin True] obtain ql where
          ql:"ql ∈ poss l1 ∧ l1 |_ ql = Var x" by auto
        let ?C = "ctxt_of_pos_term ql (l1⋅σ1)"
        from linear_term_replace_in_subst[OF lin, of ql] ql σ1'_def have 1:"l1⋅σ1' = ?C⟨?u⟩" by metis
        have 2:"(?C⟨(ctxt_of_pos_term q21 x))⟨l2 ⋅ σ2⟩⟩,?C⟨?u⟩) ∈ (rstep Rω)" using lr2 by auto
        from ql q2 at_q2 have "?C⟨(ctxt_of_pos_term q21 x))⟨l2 ⋅ σ2⟩⟩ = l1 ⋅ σ1"
          by (simp add: replace_at_ident)
        with 1 2 show ?thesis by auto 
      next
        case False
        with term_subst_eq_conv show ?thesis unfolding σ1'_def by fastforce
      qed
      hence "(C⟨l1 ⋅ σ1⟩, C⟨l1 ⋅ σ1'⟩) ∈ (rstep Rω)=" by auto
      with rstep_mono[OF Rw_subset_Rinf] have left:"(C⟨l1 ⋅ σ1⟩, C⟨l1 ⋅ σ1'⟩) ∈ (rstep R)=" by auto
      from lr(1) Rw_subset_Rinf have right':"(C⟨r1 ⋅ σ1'⟩, C⟨l1 ⋅ σ1'⟩) ∈ rstep (R ∪ Eω)" by fastforce
      from lr(1) Ew_linear Rw_linear have lin:"linear_term r1" unfolding linear_trs_def by auto
      from linear_term_replace_in_subst[OF lin q1 rq1, of σ1] have "(ctxt_of_pos_term q1 (r1 ⋅ σ1))⟨?u⟩ = r1 ⋅ σ1'"
        using σ1'_def by auto
      with p rq1 q1 q2 have "(ctxt_of_pos_term q (r1 ⋅ σ1))⟨r2 ⋅ σ2⟩ = r1 ⋅ σ1'"
        by (simp add: ctxt_of_pos_term_append)
      hence u:"C⟨r1 ⋅ σ1'⟩ = u" unfolding lr2 DC ctxt_ctxt_compose ctxt_eq by auto
      from right' have "(C⟨l1 ⋅ σ1'⟩, C⟨r1 ⋅ σ1'⟩) ∈ rstep (R¯ ∪ Eω)" unfolding reverseTrs
        using converse_iff[of _ _ "rstep (Eω)"] symcl_converse by fastforce
      with left have "(s,u) ∈ (rstep R)= O (rstep (R¯ ∪ Eω))="
        unfolding u lr(2)[symmetric] by auto
      thus ?thesis unfolding rstep_union[of "R¯"] unfolding relcomp_distrib by auto
    next
      case True (* proper overlap *)
      note st = lr[unfolded ctxt_apply_term.simps]
      note tu = lr2[unfolded ctxt_apply_term.simps]
      let ?C = "ctxt_of_pos_term q (r1 ⋅ σ1)"
      from hole_pos_ctxt_of_pos_term[OF poss_imp_subst_poss, of q] True have qC:"hole_pos ?C = q"
        using funposs_imp_poss by auto
      let ?C = "ctxt_of_pos_term q (r1 ⋅ σ1)"
      from Rw_less lr2 assms(3) have "l2 ≻ r2 ∧ ¬ l1 ≻ r1" by auto
      with LCP_exists[OF eq'', unfolded qC, OF True, of _ "Rω ∪ Eω"] lr(1) lr2(1) have
        "∃τ' s t. (s, t) ∈ LCP (Rω ∪ Eω) ∧ (ctxt_of_pos_term q (r1 ⋅ σ1))⟨r2 ⋅ σ2⟩ = s ⋅ τ' ∧ l1 ⋅ σ1 = t ⋅ τ'" by auto
      then obtain ρ s0 t0 where lcp:"(s0, t0) ∈ LCP (Rω ∪ Eω)" and
        id:"(ctxt_of_pos_term q (r1 ⋅ σ1))⟨r2 ⋅ σ2⟩ = s0 ⋅ ρ" "l1 ⋅ σ1 = t0 ⋅ ρ" by auto

      with fair have "(s0, t0) ∈ rstep (E)" using rstep_simps(5) by fast
      from rstep.ctxt[OF rstep.subst[OF this], of C ρ] id have "(D⟨r2 ⋅ σ2⟩, C⟨l1 ⋅ σ1⟩) ∈ rstep (E)"
        unfolding DC ctxt_ctxt_compose lr by auto
      with lr(2) lr2(3) have Einf_step:"(s, u) ∈ (rstep E)" unfolding rstep_simps(5) by fastforce
      with Einf_to_Ew_rstep have su:"(s, u) ∈ ((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)" by blast
      have "((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)¯ = ((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)"
        unfolding converse_relcomp rtrancl_converse converse_converse using O_assoc reverseTrs by blast
      with su show ?thesis by blast
    qed
  next
    case 2
    with prefix_exists obtain q where q:"p1 = ?p2 <#> q" by metis
    from lr(3) q subt_at_hole_pos have DC:"C = D ∘c (ctxt_of_pos_term q (l2 ⋅ σ2))" unfolding lr2 lr(4)
      using ctxt_of_pos_term_append[OF hole_pos_poss] ctxt_of_pos_term_hole_pos by metis
    from eq[unfolded DC] have eq'':"(ctxt_of_pos_term q (l2 ⋅ σ2))⟨r1 ⋅ σ1⟩ = l2 ⋅ σ2"
      unfolding ctxt_ctxt_compose by auto
    show ?thesis proof(cases "q ∈ funposs l2")
      case False (* variable overlap *)
      note q_not_in_funposs_l2 = this
      from lr lr2(2) q have q_l2σ:"q ∈ poss (l2 ⋅ σ2)" by (metis hole_pos_poss hole_pos_poss_conv)
      from poss_subst_apply_term[OF this q_not_in_funposs_l2] obtain q1 q2 x
        where p[simp]: "q = q1 <#> q2" and q1: "q1 ∈ poss l2"
          and l2q1: "l2 |_ q1 = Var x" and q2: "q2 ∈ poss (σ2 x)" by auto
      from eq have [simp]: "l2 ⋅ σ2 |_ q = r1 ⋅ σ1"
        unfolding DC q ctxt_ctxt_compose ctxt_eq using replace_at_subt_at[OF q_l2σ] by metis
      with l2q1 q2 subt_at_append q_l2σ q1 have at_q2[simp]: 2 x |_ q2 = r1 ⋅σ1" unfolding p by auto

      let ?u = "replace_at (σ2 x) q2 (l1 ⋅ σ1)"
      define σ2' where 2' y = (if y = x then ?u else σ2 y)" for y
      have "(σ2 x, σ2' x) ∈ rstep (R ∪ Eω)"
      proof -
        let ?C = "ctxt_of_pos_term q22 x)"
        from rstepI[OF lr', of _ ?C σ1] have "(?C⟨r1 ⋅ σ1⟩, ?C⟨l1 ⋅ σ1⟩) ∈ rstep (R ∪ Eω)"
          unfolding rstep_union reverseTrs[of "R", symmetric] by fastforce
        then show ?thesis using q2 by (simp add: σ2'_def replace_at_ident)
      qed
      then have *: "⋀x. (σ2 x, σ2' x) ∈ (rstep (R ∪ Eω))=" by (auto simp: σ2'_def)

      then have "(r2 ⋅ σ2, r2 ⋅ σ2') ∈ (rstep (R ∪ Eω))=" proof(cases "x ∈ vars_term r2")
        case True
        from lr2(1) Rw_linear have lin:"linear_term r2" unfolding linear_trs_def by auto
        from linear_term_unique_var_pos[OF lin True] obtain ql where
          ql:"ql ∈ poss r2" "r2 |_ ql = Var x" by auto
        let ?C = "ctxt_of_pos_term ql (r2⋅σ2)"
        from linear_term_replace_in_subst[OF lin ql] σ2'_def have 1:"r2⋅σ2' = ?C⟨?u⟩" by metis
        have 2:"(?C⟨(ctxt_of_pos_term q22 x))⟨r1 ⋅ σ1⟩⟩,?C⟨?u⟩) ∈ rstep (R ∪ Eω)" using lr' by auto
        from ql q2 at_q2 have "?C⟨(ctxt_of_pos_term q22 x))⟨r1 ⋅ σ1⟩⟩ = r2 ⋅ σ2"
          by (simp add: replace_at_ident)
        with 1 2 show ?thesis by auto 
      next
        case False
        with term_subst_eq_conv show ?thesis unfolding σ2'_def by fastforce
      qed
      hence right:"(D⟨r2 ⋅ σ2⟩, D⟨r2 ⋅ σ2'⟩) ∈ (rstep (R ∪ Eω))=" by auto
      hence right:"(D⟨r2 ⋅ σ2'⟩, D⟨r2 ⋅ σ2⟩) ∈ (rstep (R¯ ∪ Eω))="
        unfolding rstep_union[of _ " Eω"] Un_iff reverseTrs[symmetric] converse_iff by auto
      from lr2(1) have "(D⟨l2 ⋅ σ2'⟩, D⟨r2 ⋅ σ2'⟩) ∈ rstep Rω" by fast
      with rstep_mono[OF Rw_subset_Rinf] have left:"(D⟨l2 ⋅ σ2'⟩, D⟨r2 ⋅ σ2'⟩) ∈ rstep R" by auto
      from lr2(1) Rw_linear have lin:"linear_term l2" unfolding linear_trs_def by auto
      from linear_term_replace_in_subst[OF lin q1 l2q1, of σ2] have "(ctxt_of_pos_term q1 (l2 ⋅ σ2))⟨?u⟩ = l2 ⋅ σ2'"
        using σ2'_def by simp
      with p l2q1 q1 q2 have "(ctxt_of_pos_term q (l2 ⋅ σ2))⟨l1 ⋅ σ1⟩ = l2 ⋅ σ2'"
        by (simp add: ctxt_of_pos_term_append)
      with lr have s:"s = D⟨l2 ⋅ σ2'⟩" unfolding lr DC ctxt_ctxt_compose ctxt_eq by auto
      from right left have "(s,u) ∈ (rstep R) O (rstep (R¯ ∪ Eω))="
        unfolding s lr2(3)[symmetric] using converse_iff by auto
      thus ?thesis unfolding rstep_union relcomp_distrib by auto
    next
      case True (* proper overlap *)
      note st = lr[unfolded ctxt_apply_term.simps]
      note tu = lr2[unfolded ctxt_apply_term.simps]
      let ?C = "ctxt_of_pos_term q (l2 ⋅ σ2)"
      from hole_pos_ctxt_of_pos_term[OF poss_imp_subst_poss, OF funposs_imp_poss, OF True] have qC:"hole_pos ?C = q"
        using funposs_imp_poss by auto
      let ?C = "ctxt_of_pos_term q (r1 ⋅ σ1)"
      from Rw_less lr2 assms(3) have "l2 ≻ r2 ∧ ¬ l1 ≻ r1" by auto
      with LCP_exists[OF eq'', unfolded qC, OF True, of l1 "Rω ∪ Eω"] lr(1) lr2(1) have
        "∃τ' s t. (s, t) ∈ LCP (Rω ∪ Eω) ∧ (ctxt_of_pos_term q (l2 ⋅ σ2))⟨l1 ⋅ σ1⟩ = s ⋅ τ' ∧ r2 ⋅ σ2 = t ⋅ τ'" by auto
      then obtain ρ s0 t0 where lcp:"(s0, t0) ∈ LCP (Rω ∪ Eω)" and
        id:"(ctxt_of_pos_term q (l2 ⋅ σ2))⟨l1 ⋅ σ1⟩ = s0 ⋅ ρ" "r2 ⋅ σ2 = t0 ⋅ ρ" by auto

      with fair have "(s0, t0) ∈ rstep (E)" using rstep_simps(5) by fast
      from rstep.ctxt[OF rstep.subst[OF this], of D ρ] have "(C⟨l1 ⋅ σ1⟩, D⟨r2 ⋅ σ2⟩) ∈ rstep (E)"
        unfolding id[symmetric] DC ctxt_ctxt_compose by auto
      then have Einf_step:"(s, u) ∈ (rstep E)" unfolding lr(2) lr2(3) rstep_simps(5) by simp
      with Einf_to_Ew_rstep have su:"(s, u) ∈ ((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)" by blast
      have "((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)¯ = ((rstep R)* O (rstep (Eω))= O ((rstep R)¯)*)"
        unfolding converse_relcomp rtrancl_converse converse_converse using O_assoc reverseTrs by blast
      with su show ?thesis by blast
    qed
  next
    case 3
    from lr' lr have ts:"(t, s) ∈ rstep_r_p_s (R ∪ Eω) (r1, l1) p1 σ1"
      unfolding rstep_r_p_s_def mem_Collect_eq by auto
    from lr2 Rw_subset_Rinf have tu:"(t, u) ∈ rstep_r_p_s R (l2, r2) ?p2 σ2"
      unfolding rstep_r_p_s_def mem_Collect_eq by auto
    let ?v = "(ctxt_of_pos_term p1 u)⟨l1 ⋅ σ1⟩"
    from parallel_steps[OF ts tu 3] have sv:"(s,?v) ∈ rstep R" and "(u,?v) ∈ rstep (R ∪ Eω)"
      using rstep_r_p_s_imp_rstep by auto
    from this(2) reverseTrs converse_iff have "(?v,u) ∈ rstep (R¯ ∪ Eω)" by blast
    with sv have "(s,u) ∈ (rstep R O rstep (R¯ ∪ Eω))" by auto
    thus ?thesis unfolding rstep_union[of "R¯"] relcomp_distrib Un_iff using r_into_rtrancl by auto
  qed
qed

lemma Rinf_Rw_msteps:"(l,r) ∈ R ⟹ ∃y S. (l,y) ∈ rstep Rω ∧ (∀s ∈# S. l ≻ s) ∧ (y,r) ∈ (mstep S (Rω ∪ Eω))*" 
proof (induct "(l,r)" arbitrary: l r rule: SN_induct [OF SN_lexless])
  case 1
  note IH = this
  show ?case proof(cases "(l,r) ∈ Rω")
    case True
    show ?thesis by (rule exI[of _ r],rule exI[of _ "{#}"], insert rstep.rule[OF True], auto)
  next
    case False
    with Rinf_without_Rw 1(2) obtain j where j:"(l, r) ∈ R j" "(l, r) ∉ R (Suc j)" by auto
    let ?E = "E (Suc j)" and ?R = "R (Suc j)"
    from oKBilin_R_supset[OF lin_irun, of j] j consider
      "(l, r) ∈ encstep2 {} ?R O ?E" | "(l, r) ∈ ?R O (rstep ?R)¯" by auto
    then show ?thesis proof(cases)
      case 1
      then obtain u where u:"(l, u) ∈ encstep2 {} ?R" "(u,r) ∈ ?E" by blast
      from u(1) obtain l' r' C σ where lr':
        "(l',r') ∈ ?R" "l = C⟨l' ⋅ σ⟩" "u = C⟨r' ⋅ σ⟩" "l' ⊲⋅ l" unfolding encstep2.simps by auto
      then have l_l0: "l ⋅≻ l'" unfolding lessencp_def by auto
      hence "((l,r), (l',r')) ∈ lexless" by simp
      from IH(1)[OF this] lr'(1) obtain v S' where v:
        "(l', v) ∈ rstep Rω" "∀s∈#S'. l' ≻ s" "(v, r') ∈ (mstep S' (Rω ∪ Eω))*" by auto
      let ?v = "C⟨v⋅σ⟩"
      from v(1) have step:"(l, ?v) ∈ rstep Rω" unfolding lr' by auto
      define S where "S ≡ {#C⟨u ⋅ σ⟩. u ∈# S'#}"
      define T where "T ≡ {#C⟨u ⋅ σ⟩. u ∈# S'#} + {#u,r#}"
      have ST:"∀s∈#S. ∃t∈#T. t ≽ s" unfolding S_def T_def by auto
      from v(2) ctxt[OF subst, of _ _ C σ] have lS:"∀s∈#S. l ≻ s" unfolding S_def lr' by auto
      from msteps_subst_ctxt[OF v(3)] have "(?v, u) ∈ (mstep S (Rω ∪ Eω))*"
        unfolding lr' S_def by auto
      with msteps_succeq_mono[OF ST] have m1:"(?v, u) ∈ (mstep T (Rω ∪ Eω))*" by blast
      from u(2) have "(u,r) ∈ mstep T (E ∪ R)" unfolding mstep_def mem_Collect_eq T_def by auto
      with ERi_subset_ERw[of T] have "(u, r) ∈ (mstep T (Rω ∪ Eω))*" by blast
      with transD[OF conversion_trans] m1 have msteps:"(?v, r) ∈ (mstep T (Rω ∪ Eω))*" by metis
      from ostep_imp_less[OF Ri_less] Ri_less IH(2) encstep2_ostep[OF u(1)] have "l ≻ u ∧ l ≻ r"
        using Rinf_less by fast
      with lS have lT:"∀t∈#T. l ≻ t" unfolding T_def S_def by auto
      show ?thesis by (rule exI[of _ ?v], rule exI[of _ T], insert step lT msteps, auto)
    next
      case 2
      then obtain u where u:"(l, u) ∈ R (Suc j)" "(r,u) ∈ rstep (R (Suc j))" by blast
      with compatible_rstep_imp_less[OF Ri_less] have "r ≻ u" by auto
      hence "((l,r), (l,u)) ∈ lexless" unfolding lex_two.simps mem_Collect_eq lessencp_def by fast
      from IH(1)[OF this] u(1) obtain y S 
        where yS:"(l, y) ∈ rstep Rω" "∀t∈#S. l ≻ t" "(y, u) ∈ (mstep S (Rω ∪ Eω))*" by blast
      define T where "T ≡ S + {#r,u#}"
      from Ri_less compatible_rstep_imp_less[OF Ri_less] IH(2) u(2) have gt:"l ≻ r ∧ l ≻ u"
        using trans[of l r u] by auto
      with yS(2) have T:"∀t∈#T. l ≻ t" unfolding T_def by simp
      from insert u(2) gt have "(r,u) ∈ mstep T ?R" unfolding mstep_def mem_Collect_eq T_def
        by auto
      hence "(r,u) ∈ mstep T (E ∪ R)" using mstep_mono[of ?R "E ∪ R"] by fast
      with ERi_subset_ERw[of T] have ru:"(r, u) ∈ (mstep T (Rω ∪ Eω))*" by blast
      have ST:"∀s∈#S. ∃t∈#T. t ≽ s" unfolding T_def by auto
      from msteps_succeq_mono[OF ST] yS(3) have "(y, u) ∈ (mstep T (Rω ∪ Eω))*" by blast
      with ru transD[OF conversion_trans] have msteps:"(y, r) ∈ (mstep T (Rω ∪ Eω))*"
        unfolding conversion_inv[of r] by metis
      show ?thesis by (rule exI[of _ y], rule exI[of _ T], insert step yS T msteps, auto)
    qed
  qed
qed

lemma Rinf_step_Rw_msteps:
  assumes "(s,t) ∈ rstep R"
  shows "∃u S. (s,u) ∈ rstep Rω ∧ (∀t∈#S. s ≻ t) ∧ (u,t) ∈ (mstep S (Rω ∪ Eω))*" 
proof-
  from assms obtain C σ l r where lr:"(l,r) ∈ R" "s = C⟨l⋅σ⟩" "t = C⟨r⋅σ⟩" by fast
  from Rinf_Rw_msteps[OF lr(1)] obtain u S where uS:
    "(l, u) ∈ rstep Rω" "∀t∈#S. l ≻ t" "(u, r) ∈ (mstep S (Rω ∪ Eω))*" by auto
  hence step:"(s, C⟨u⋅σ⟩) ∈ rstep Rω" unfolding lr by auto
  define T where "T ≡ {#C⟨u ⋅ σ⟩. u ∈# S#}"
  from uS(2) ctxt[OF subst, of _ _ C σ] have cover:"∀t∈#T. s ≻ t" unfolding T_def lr by auto
  from msteps_subst_ctxt[OF uS(3), of C σ] have "(C⟨u ⋅ σ⟩, t) ∈ (mstep T (Rω ∪ Eω))*"
    unfolding T_def lr by argo
  with step cover show ?thesis by auto
qed

lemma Rinf_step_ERw_msteps_src:
  assumes "(s,t) ∈ rstep R"
  shows "(s,t) ∈ (mstep {#s#} (Rω ∪ Eω))*"
proof-
  from assms Rinf_step_Rw_msteps obtain u S where
    rstep0:"(s,u) ∈ rstep Rω" and S:"∀t∈#S. s ≻ t" and ut:"(u,t) ∈ (mstep S (Rω ∪ Eω))*" by blast
  from rstep0 compatible_rstep_imp_less Rw_less have "s ≻ u" by auto
  with rstep0 have su:"(s,u) ∈ mstep {#s#} (Rω ∪ Eω)" unfolding mstep_def mem_Collect_eq by auto
  from mstep_succeq_mono[of S "{#s#}"] S have "mstep S (Rω ∪ Eω) ⊆ mstep {#s#} (Rω ∪ Eω)" by auto
  from conversion_mono[OF this] ut have "(u, t) ∈ (mstep {#s#} (Rω ∪ Eω))*" by auto
  with r_into_rtrancl[OF su, THEN conversionI'] show ?thesis unfolding conversion_def
    using rtrancl_trans by auto
qed

lemma Rinf_steps_ERw_msteps_src:
  assumes "(s,t) ∈ (rstep R)*"
  shows "(s,t) ∈ (mstep {#s#} (Rω ∪ Eω))*"
  using assms
proof(induct rule:converse_rtrancl_induct)
case base
  then show ?case unfolding mstep_def by auto
next
  case (step s u)
  note lift = mstep_succeq_mono[of "{#u#}" "{#s#}",THEN conversion_mono, of "Rω ∪ Eω"]
  from step(1) compatible_rstep_imp_less Rinf_less have "s ≻ u" by auto
  with lift step(3) have ut:"(u, t) ∈ (mstep {#s#} (Rω ∪ Eω))*" by auto
  from Rinf_step_ERw_msteps_src[OF step(1)] have "(s,u) ∈ (mstep {#s#} (Rω ∪ Eω))*" by auto
  with ut conversion_O_conversion show ?case by force
qed

lemma mstep_conv_subset_add_mset: "(mstep S RR)* ⊆ (mstep (S' + S) RR)*"
  using msteps_succeq_mono[of S "S' + S"] by auto

lemma Rinf_steps_Rw_msteps:
  assumes "(s,t) ∈ (rstep R)+"
  shows "∃u S. (s,u) ∈ rstep Rω ∧ (∀t∈#S. s ≻ t) ∧ (u,t) ∈ (mstep S (Rω ∪ Eω))*" 
proof-
  from tranclD[OF assms] obtain u where u:"(s,u) ∈ rstep R" "(u,t) ∈ (rstep R)*" by auto
  from Rinf_step_Rw_msteps[OF u(1)] obtain v S where
    v:"(s,v) ∈ rstep Rω" "∀t∈#S. s ≻ t" "(v,u) ∈ (mstep S (Rω ∪ Eω))*" by auto
  from Rinf_steps_ERw_msteps_src[OF u(2)] mstep_conv_subset_add_mset[of "{#u#}" _ S]
    have ut:"(u, t) ∈ (mstep ({#u#} + S) (Rω ∪ Eω))*" unfolding add.commute[of S] by blast
  from v(3) mstep_conv_subset_add_mset[of S _ "{#u#}"] have "(v, u) ∈ (mstep ({#u#} + S) (Rω ∪ Eω))*" by blast
  with ut have vt:"(v, t) ∈ (mstep ({#u#} + S) (Rω ∪ Eω))*" by fastforce
  from compatible_rstep_imp_less[OF Rinf_less u(1)] v(2) have "∀t∈#{#u#} +S. s ≻ t" by auto
  with vt v(1) show ?thesis by blast
qed

context
  assumes Ew_unorientable:"Eω ∩ {≻} = {}" and Ew_nontrivial:"Eω ∩ Id = {}"
begin

definition mstep1 where "mstep1 M RR = {(x,y) |x y x'. (x,y) ∈ rstep RR ∧ M = {#x'#} ∧ x' ≽ x ∧ x' ≽ y}"
definition mstep2 where "mstep2 M RR = {(x,y) |x y x' y'. (x,y) ∈ rstep RR ∧ M = {#x',y'#} ∧ x' ≽ x ∧ y' ≽ y}"

lemma mstep1I:"(x,y) ∈ rstep RR ⟹ x' ≽ x ⟹ x' ≽ y ⟹ (x,y) ∈ mstep1 {#x'#} RR"
  unfolding mstep1_def by force

lemma mstep2I:"(x,y) ∈ rstep RR ⟹ x' ≽ x ⟹ y' ≽ y ⟹ (x,y) ∈ mstep2 {#x',y'#} RR"
  unfolding mstep2_def by force

lemma mstep1_dom:"(x, y) ∈ mstep1 M Rω ⟹ M ≻mul {#y#}"
proof-
  assume a:"(x, y) ∈ mstep1 M Rω"
  with compatible_rstep_imp_less[OF Rw_less] have xy:"x ≻ y" unfolding mstep1_def by auto
  from a obtain x' where M:"M = {#x'#}" and "x' ≽ x" unfolding mstep1_def by auto
  with trans xy have "x' ≻ y" by auto
  then show "M ≻mul {#y#}" unfolding M using ns_mul_ext_singleton2 by fastforce
qed

lemma mstep2_dom:"(x, y) ∈ (mstep2 M RR) ⟹ M ≻mul {#y#}"
proof-
  assume a:"(x, y) ∈ (mstep2 M RR)"
  from a obtain x' y' x'' y'' where M:"M = {#x',y'#}" and "x'' ∈# M" and "y'' ∈# M" and "x'' ≽ x"  and "y'' ≽ y"
    unfolding mstep2_def by auto
  then show "M ≻mul {#y#}" unfolding M using s_mul_ext_singleton
    by (smt add_mset_add_single add_mset_eq_singleton_iff all_s_s_mul_ext case_prodI empty_not_add_mset mem_Collect_eq mulsucc.rtrancl_NS multi_member_split rtrancl_eq_or_trancl s_ns_mul_ext_union_compat sup2E)
qed

lemma mstep1_rstep[simp]:"(x, y) ∈ mstep1 M RR ⟹ (x, y) ∈ rstep RR" unfolding mstep1_def by auto

lemma mstep2_rstep[simp]:"(x, y) ∈ mstep2 M RR ⟹ (x, y) ∈ rstep RR" unfolding mstep2_def by auto

lemma mstep1_subset_mstep: "mstep1 M Rω ⊆ mstep M Rω"
proof
  fix x y
  assume a:"(x, y) ∈ mstep1 M Rω"
  then obtain x' where M:"M = {#x'#}" "x' ≽ x" "x' ≽ y" unfolding mstep1_def by auto
  with a M show "(x, y) ∈ mstep M Rω" unfolding mstep_def mem_Collect_eq split by auto
qed

lemma mstep2_subset_mstep: "mstep2 M RR ⊆ mstep M RR"
proof
  fix x y
  assume a:"(x, y) ∈ mstep2 M RR"
  thus "(x, y) ∈ mstep M RR" unfolding mstep_def mstep2_def mem_Collect_eq split by force
qed

lemma mstep1_conv_mstep: "mstep1 {#z#} RR = mstep {#z#} RR"
proof-
  { fix x y
    assume a:"(x, y) ∈ mstep {#z#} RR"
    hence "(x, y) ∈ mstep1 {#z#} RR" unfolding mstep1_def mstep_def mem_Collect_eq split by auto
  } note A = this
  { fix x y
    assume a:"(x, y) ∈ mstep1 {#z#} RR"
    with a have "(x, y) ∈ mstep {#z#} RR" unfolding mstep1_def mstep_def mem_Collect_eq split by auto
  }
  with A show ?thesis by auto
qed

lemma mstep2_converse:"(mstep2 M RR) = mstep2 M (RR)"
unfolding mstep2_def mem_Collect_eq using add_mset_commute by fast

lemma mstep1_union:"mstep1 M (R1 ∪ R2) = mstep1 M R1 ∪ mstep1 M R2"
  unfolding mstep1_def by auto

interpretation A:ars_mod_labeled_sn "λM. mstep1 M Rω" "λM. mstep2 M Eω" UNIV UNIV "op ≻mul"
  using mulsucc.SN by unfold_locales auto

lemma succ_irrefl:"¬ s ≻ s"
  using SN_imp_acyclic[OF SN_less] unfolding  acyclic_irrefl irrefl_def by auto

lemma pdm:
  assumes "A.mod_peak s M w N t"
  shows "(s, t) ∈ (⋃K∈A.downset2 M N. (mstep1 K Rω ∪ mstep2 K Eω))*"
proof-
  let ?J = "(rstep R)* O (rstep (Eω))= O ((rstep R)¯)*"
  from compatible_rstep_trancl_imp_less[OF Rinf_less] have succeq:"⋀x y.(x,y) ∈ (rstep R)* ⟹ x ≽ y"
    unfolding rtrancl_eq_or_trancl by blast
  from assms mstep_Un have "∃s' t' M' N'. (w,s') ∈ (mstep1 M' Rω ∪ mstep2 M' (Eω)) ∧ (w,t') ∈ mstep1 N' Rω ∧
    ((s,t) = (s',t') ∨ (s,t) = (t',s')) ∧ (M' = M ∨ M' = N)"
    unfolding A.mod_peak_def mstep2_converse by blast
  then obtain s' t' M' N' where *:"(w,s') ∈ (mstep1 M' Rω ∪ mstep2 M' (Eω))" "(w,t') ∈ mstep1 N' Rω"
    "((s,t) = (s',t') ∨ (s,t) = (t',s'))" "M' = M ∨ M' = N" unfolding A.mod_peak_def by force
  from * have rstep:"(w, s') ∈ rstep (Rω ∪ Eω)" "(w, t') ∈ rstep Rω"
    using mstep1_rstep mstep2_rstep reverseTrs by blast+
  with reverseTrs have "(s', w) ∈ rstep (Rω¯ ∪ Eω)" by fast
  then obtain l r p σ where 1:"(s', w) ∈ rstep_r_p_s (Rω¯ ∪ Eω) (l,r) p σ"
    unfolding rstep_iff_rstep_r_p_s by fast
  hence lr:"(l,r) ∈ Rω¯ ∪ Eω" unfolding rstep_r_p_s_def mem_Collect_eq split by meson
  hence ngt_R:"(l,r) ∈ Rω¯ ⟹ ¬ l ≻ r" using SN_imp_acyclic[OF SN_less] trans[of r l r] Rw_less
    unfolding converse_iff acyclic_irrefl irrefl_def by (cases "l ≻ r", auto)
  with lr Ew_unorientable have "¬ l ≻ r" by auto
  with linear_peak_cases[OF 1 rstep(2)] have valley:"(s', t') ∈ ?J" by auto
  have J_inv:"?J¯ = ?J" using O_assoc symcl_converse
    unfolding converse_relcomp rtrancl_converse converse_converse converse_Un rstep_simps(5) by blast
  from *(3) have valley:"(s, t) ∈ ?J" by (rule disjE, insert valley J_inv converse_iff) blast+
  then obtain s'' t'' where **:"(s,s'') ∈ (rstep R)*" "(s'',t'') ∈ (rstep (Eω))=" "(t,t'') ∈ (rstep R)*"
    unfolding rtrancl_converse by auto
  let ?M = "λM. (mstep M (Rω ∪ Eω))*"
  let ?M1 = "λM. (mstep1 M (Rω ∪ Eω))*"
  from Rinf_steps_ERw_msteps_src **(1) **(3) have "(s, s'') ∈ ?M {#s#}" "(t, t'') ∈ ?M {#t#}" by auto
  with mstep1_conv_mstep have Rsteps:"(s, s'') ∈ ?M1 {#s#}" "(t, t'') ∈ ?M1 {#t#}" by auto
  note a = assms[unfolded A.mod_peak_def]
  from a mstep1_dom[of w s M] mstep2_dom have cover1:"M ≻mul {#s#} ∨ N ≻mul {#s#}" by blast
  from a mstep1_dom[of w t] mstep2_dom[of w t N] have cover2:"M ≻mul {#t#} ∨ N ≻mul {#t#}" by blast
  from *(2) compatible_rstep_imp_less[OF Rw_less] have wt:"w ≻ t'" unfolding mstep1_def by auto
  have "M' ≻mul {#s',t'#}" proof(cases "(w, s') ∈ mstep1 M' Rω")
    case True
    then have "∃u v. u ∈# M' ∧ u ≽ w  ∧ v ∈# M' ∧ v ≽ s'" unfolding mstep1_def by auto
    then obtain u v where uv:"u ∈# M'" "u ≽ w" "v ∈# M'" "v ≽ s'" unfolding mstep1_def mstep2_def by blast
    from True compatible_rstep_imp_less[OF Rw_less] have "w ≻ s'" unfolding mstep1_def by auto
    with uv(2) wt trans[of u w] have "u ≻ t' ∧ u ≻ s'" by auto
    with uv(1) have u:"∀b. b ∈# {#s', t'#} ⟶ (∃a. a ∈# M' ∧ a ≻ b)" by auto
    from multi_member_split[OF uv(1)] uv(1) empty_not_add_mset have "M' ≠ {#}" by auto
    from all_s_s_mul_ext[OF this, of "{#s',t'#}"] u show ?thesis by auto
  next
    case False
    with *(1) have step:"(w, s') ∈ mstep2 M' (Eω)" by auto
    then obtain u v where uv:"M' = {#u,v#}" "u ≽ w" "v ≽ s'" unfolding mstep2_def by auto
    from uv(3) ns_mul_ext_singleton2[of v s'] have ns:"{#v#} ≽mul {#s'#}" by (cases "v = s'", auto)
    from uv trans wt have "u ≻ t'" by auto
    hence "{#u#} ≻mul {# t'#}" by auto
    with uv(3) s_ns_mul_ext_union_compat[OF this ns] show ?thesis
      unfolding uv add_mset_commute[of u v "{#}"] by auto
  qed
  with *(3) have "M' ≻mul {#s,t#}" using add_mset_commute[of s' t' "{#}"] by fastforce  
  with *(4) have cover3:"(M ≻mul {#s,t#} ∨ N ≻mul {#s,t#})" by auto
  let ?D = "⋃K∈A.downset2 M N. (mstep1 K Rω ∪ mstep2 K Eω)"
  from ** succeq have "s ≽ s''" "t ≽ t''" by auto
  with **(2) have "(s'', t'') ∈ (mstep2 {#s,t#} (Eω))=" unfolding mstep2_def mem_Collect_eq by auto
  with cover3 have st:"(s'', t'') ∈ ?D*" using mstep2_converse by fast

  { fix u S
    assume u:"u = s ∨ u = t" and S:"∀t ∈# S. u ≻ t"
    have D1:"mstep S Rω ⊆ ?D" proof
      fix x y
      assume step:"(x, y) ∈ mstep S Rω"
      with u S trans[of u] have "u ≻ x" "u ≻ y" unfolding mstep_def mem_Collect_eq split by fast+
      with cover1 have "(x, y) ∈ mstep1 {#u#} Rω" using step unfolding mstep_def by (auto intro!: mstep1I)
      with cover1 cover2 u show "(x,y) ∈ ?D" by auto
    qed
    have "mstep S Eω ⊆ ?D" proof
      fix x y
      assume step:"(x, y) ∈ mstep S Eω"
      with S u trans[of u] have s:"u ≻ x" "u ≻ y" unfolding mstep_def mem_Collect_eq split by fast+
      have dec:"{#u#} ≻mul {#x,y#}" by (auto intro!:all_s_s_mul_ext simp:s)
      have "(x, y) ∈ mstep2 {#x,y#} Eω" using step unfolding mstep_def by (auto intro!: mstep2I)
      with u mulsucc.trans_S_point[OF _ dec] cover1 cover2 show "(x,y) ∈ ?D" by blast
    qed
    with D1 have "mstep S (Rω ∪ Eω) ⊆ ?D" by auto
    from conversion_mono[OF this] have "(mstep S (Rω ∪ Eω))* ⊆ ?D*" by auto
  } note conv_dec = this

  have ss:"(s, s'') ∈ ?D*" proof(cases "s=s''", simp)
    assume "s ≠ s''"
    with **(1) have "(s, s'') ∈ (rstep R)+" unfolding rtrancl_eq_or_trancl by auto
    from Rinf_steps_Rw_msteps[OF this] obtain s1 S where
      S:"(s, s1) ∈ rstep Rω" "∀t ∈# S. s ≻ t" "(s1, s'') ∈ (mstep S (Rω ∪ Eω))*" by auto
    have "(s, s1) ∈ mstep1 {#s#} Rω" using compatible_rstep_imp_less[OF Rw_less] S(1) by (auto intro!:mstep1I)
    with cover1 have 1:"(s, s1) ∈ ?D*" by auto
    from conv_dec[OF _ S(2)] S(3) have "(s1, s'') ∈ ?D*" by auto
    with conversion_trans'[OF 1] show "(s, s'') ∈ ?D*" by blast
  qed

  have "(t, t'') ∈ ?D*" proof(cases "t=t''", simp)
    assume "t ≠ t''"
    with **(3) have "(t, t'') ∈ (rstep R)+" unfolding rtrancl_eq_or_trancl by auto
    from Rinf_steps_Rw_msteps[OF this] obtain t1 T where
      S:"(t, t1) ∈ rstep Rω" "∀u ∈# T. t ≻ u" "(t1, t'') ∈ (mstep T (Rω ∪ Eω))*" by auto
    have "(t, t1) ∈ mstep1 {#t#} Rω" using compatible_rstep_imp_less[OF Rw_less] S(1) by (auto intro!:mstep1I)
    with cover2 have 1:"(t, t1) ∈ ?D*" by auto
    from conv_dec[OF _ S(2)] S(3) have "(t1, t'') ∈ ?D*" by auto
    with conversion_trans'[OF 1] show "(t, t'') ∈ ?D*" by blast
  qed
  hence tt:"(t'', t) ∈ ?D*" using conversion_sym[unfolded sym_def, rule_format, of t t''] by blast
  note t = conversion_trans[THEN transD]
  from t[OF t, OF ss st ] tt show ?thesis by auto
qed

lemma CRm: "CRm (rstep Rω) (rstep Eω)"
proof-
  from A.CRm pdm have CRm:"CRm (⋃a. mstep1 a Rω) (⋃a. mstep2 a Eω)" by force
  have id:"(⋃a. mstep1 a Rω) = rstep Rω"
    unfolding mstep1_def using compatible_rstep_imp_less[OF Rw_less] by auto
  have "(⋃a. mstep2 a Eω) = rstep Eω" unfolding mstep2_def by auto
  with id CRm show ?thesis by auto
qed

lemma NF_R_subset_NF_REw:"NF_trs ℛ ⊆ NF_trs Rω ∩ NF (rstep (Eω))"
proof-
  { fix t u
    assume NF:"t ∈ NF_trs ℛ" and "(t,u) ∈ rstep (Rω ∪ (Eω))"
    then obtain l r C σ where lr:"(l,r) ∈ Rω ∪ Eω" "t = C⟨l⋅σ⟩" "u = C⟨r⋅σ⟩" by blast
    from conv_eq[unfolded oKBi_conversion_ERw] lr(1) have "(l,r) ∈ (rstep ℛ)*" by auto
    with CR_R[unfolded CR_iff_conversion_imp_join] have "(l,r) ∈ (rstep ℛ)" by auto
    then obtain v where v:"(l,v) ∈ (rstep ℛ)*" "(r,v) ∈ (rstep ℛ)*" by auto
    from NF[unfolded lr, THEN NF_subterm,THEN NF_instance] have "l ∈ NF_trs ℛ" by auto
    with v(1)[unfolded rtrancl_eq_or_trancl] tranclD[of l v] have "v=l" by force
    from v(2)[unfolded this] compatible_rstep_trancl_imp_less[OF R_less] have succeq:"r ≽ l"
      unfolding rtrancl_eq_or_trancl by blast
    have False proof(cases "(l,r) ∈ Eω")
      case True
      with Ew_nontrivial have neq:"l ≠ r" by auto
      with succeq have "r ≻ l" by blast
      with True Ew_unorientable show False by auto
    next
      case False
      with lr have "(l,r) ∈ Rω" by auto
      with compatible_rstep_imp_less[OF Rw_less] have "l ≻ r" by auto
      with SN_imp_acyclic[OF SN_less] succeq trans[OF this] show False
        unfolding acyclic_def by (cases "l=r", auto)
    qed
  }
  thus ?thesis using NF_trs_union[of Rω Eω] by blast
qed

lemma NF_Rw_subset_NF_R:"NF_trs Rω ⊆ NF_trs ℛ"
proof-
  { fix s t
    assume "(s,t) ∈ rstep ℛ"
    then obtain l r C σ where st:"s = C⟨l⋅σ⟩" "t = C⟨r⋅σ⟩" and rule:"(l,r) ∈ ℛ" by auto
    from conv_eq[unfolded oKBi_conversion_ERw] have "(l,r) ∈ (rstep (Rω ∪ Eω))*"
      using rstep.rule[OF rule] by auto
    with CRm[unfolded CRm_def] have "(l,r) ∈ (rstep Rω)* O (rstep Eω)* O ((rstep Rω)¯)*"
      unfolding rstep_union by auto
    then obtain v u where vu:"(l,u) ∈ (rstep Rω)*" "(u,v) ∈ (rstep Eω)*" "(r,v) ∈ (rstep Rω)*"
      unfolding rtrancl_converse using converse_iff by auto
    from rule ‹reduced ℛ› NF_R_subset_NF_REw have NF:"r ∈ NF_trs Rω" "r ∈ NF_trs (Eω)"
      unfolding reduced_def right_reduced_def by auto
    with vu(3) tranclD[of r v "rstep Rω"] have "r = v" unfolding rtrancl_eq_or_trancl by auto
    from vu(2) have "(v,u) ∈ (rstep Eω)*" using conversion_sym[unfolded sym_def] by metis
    with NF_no_trancl_step[OF NF(2), rule_format, of u] have "v = u"
      unfolding conversion_def rtrancl_eq_or_trancl ‹r = v› rstep_simps(5) by auto
    from vu have lr:"(l,r) ∈ (rstep Rω)*" using ‹r = v› ‹v = u› by auto
    from R_less rule SN_imp_acyclic[OF SN_less, unfolded acyclic_def] have "l ≠ r" by auto
    hence "∃t. (l,t) ∈ rstep Rω"
      using lr[unfolded rtrancl_eq_or_trancl] tranclD[of l r "rstep Rω"] by blast
    hence "∃t. (s,t) ∈ rstep Rω" unfolding st by blast
  }
  thus ?thesis by auto
qed

lemma linear_completeness:
  assumes Ew_irreducible:"∀s t. (s,t) ∈ Eω ⟶ s ∈ NF_trs Rω"
  and Ew_nontrivial:"∀t. (t,t) ∉ Eω"
  and "reduced ℛ" and "reduced Rω"
  shows "Eω = {} ∧ ℛ ≐ Rω"
proof-
  interpret complete_ars "(rstep ℛ)" by (unfold_locales, insert CR_R SN_R, auto)
  have SN_Rw:"SN (rstep Rω)" 
    by (rule SN_subset [OF SN_less], insert compatible_rstep_imp_less[OF Rw_less], auto)
  from conv_eq[unfolded oKBi_conversion_ERw] have "rstep Rω ⊆ (rstep ℛ)*" by fastforce
  from complete_NE_intro1[OF this SN_Rw NF_Rw_subset_NF_R] have
    complete:"complete_ars (rstep Rω)" and norm:"(rstep ℛ)! = (rstep Rω)!" by auto
  note vars = SN_imp_variable_condition[OF SN_R] SN_imp_variable_condition[OF SN_Rw]
  from reduced_NE_imp_unique[OF vars ‹reduced ℛ› _ norm] ‹reduced Rω have litsim:"ℛ ≐ Rω" by auto
  { fix s t
    assume st:"(s,t) ∈ Eω"
    with conv_eq[unfolded oKBi_conversion_ERw] have "(s,t) ∈ (rstep ℛ)*" by auto
    with CR_imp_conversionIff_join[OF CR_R] litsim_rstep_eq[OF litsim] obtain w where
      w:"(s, w) ∈ (rstep Rω)*" "(t, w) ∈ (rstep Rω)*" by auto
    from st Ew_irreducible[rule_format] have nfs:"s ∈ NF_trs Rω" "t ∈ NF_trs Rω" by auto
    note nfs = nfs[unfolded NF_def] w[unfolded rtrancl_eq_or_trancl] tranclD
    hence "s = w ∧ t = w" by fast
    with st Ew_nontrivial have False by auto
  }
  with litsim show ?thesis by auto
qed
end
end
end
end
end