Theory Signature_Groebner

(* Author: Alexander Maletzky *)

section ‹Signature-Based Algorithms for Computing Gr\"obner Bases›

theory Signature_Groebner
  imports More_MPoly Groebner_Bases.Syzygy Polynomials.Quasi_PM_Power_Products
begin

text ‹First, we develop the whole theory for elements of the module $K[X]^r$, i.\,e. objects of type
  @{typ "'t 0 'b"}. Later, we transfer all algorithms defined on such objects to algorithms
  efficiently operating on sig-poly-pairs, i.\,e. objects of type @{typ "'t × ('a 0 'b)"}.›

subsection ‹More Preliminaries›

lemma (in gd_term) lt_spoly_less_lcs:
  assumes "p  0" and "q  0" and "spoly p q  0"
  shows "lt (spoly p q) t term_of_pair (lcs (lp p) (lp q), component_of_term (lt p))"
proof -
  let ?l = "lcs (lp p) (lp q)"
  let ?p = "monom_mult (1 / lc p) (?l - lp p) p"
  let ?q = "monom_mult (1 / lc q) (?l - lp q) q"
  from assms(3) have eq1: "component_of_term (lt p) = component_of_term (lt q)"
    and eq2: "spoly p q = ?p - ?q"
    by (simp_all add: spoly_def Let_def lc_def split: if_split_asm)
  from p  0 have "lc p  0" by (rule lc_not_0)
  with assms(1) have "lt ?p = (?l - lp p)  lt p" and "lc ?p = 1" by (simp_all add: lt_monom_mult)
  from this(1) have lt_p: "lt ?p = term_of_pair (?l, component_of_term (lt p))"
    by (simp add: splus_def adds_minus adds_lcs)
  from q  0 have "lc q  0" by (rule lc_not_0)
  with assms(2) have "lt ?q = (?l - lp q)  lt q" and "lc ?q = 1" by (simp_all add: lt_monom_mult)
  from this(1) have lt_q: "lt ?q = term_of_pair (?l, component_of_term (lt p))"
    by (simp add: eq1 splus_def adds_minus adds_lcs_2)
  from assms(3) have "?p - ?q  0" by (simp add: eq2)
  moreover have "lt ?q = lt ?p" by (simp only: lt_p lt_q)
  moreover have "lc ?q = lc ?p" by (simp only: lc ?p = 1 lc ?q = 1)
  ultimately have "lt (?p - ?q) t lt ?p" by (rule lt_minus_lessI)
  thus ?thesis by (simp only: eq2 lt_p)
qed

subsection ‹Module Polynomials›

locale qpm_inf_term =
    gd_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict
      for pair_of_term::"'t  ('a::quasi_pm_powerprod × nat)"
      and term_of_pair::"('a × nat)  't"
      and ord::"'a  'a  bool" (infixl "" 50)
      and ord_strict (infixl "" 50)
      and ord_term::"'t  't  bool" (infixl "t" 50)
      and ord_term_strict::"'t  't  bool" (infixl "t" 50)
begin

lemma in_idealE_rep_dgrad_p_set:
  assumes "hom_grading d" and "B  punit.dgrad_p_set d m" and "p  punit.dgrad_p_set d m" and "p  ideal B"
  obtains r where "keys r  B" and "Poly_Mapping.range r  punit.dgrad_p_set d m" and "p = ideal.rep r"
proof -
  from assms obtain A q where "finite A" and "A  B" and 0: "b. q b  punit.dgrad_p_set d m"
    and p: "p = (aA. q a * a)" by (rule punit.in_pmdlE_dgrad_p_set[simplified], blast)
  define r where "r = Abs_poly_mapping (λk. q k when k  A)"
  have 1: "lookup r = (λk. q k when k  A)" unfolding r_def
    by (rule Abs_poly_mapping_inverse, simp add: finite A)
  have 2: "keys r  A" by (auto simp: in_keys_iff 1)
  show ?thesis
  proof
    show "Poly_Mapping.range r  punit.dgrad_p_set d m"
    proof
      fix f
      assume "f  Poly_Mapping.range r"
      then obtain b where "b  keys r" and f: "f = lookup r b" by (rule poly_mapping_rangeE)
      from this(1) 2 have "b  A" ..
      hence "f = q b" by (simp add: f 1)
      show "f  punit.dgrad_p_set d m" unfolding f = q b by (rule 0)
    qed
  next
    have "p = (aA. lookup r a * a)" unfolding p by (rule sum.cong, simp_all add: 1)
    also from finite A 2 have "... = (akeys r. lookup r a * a)"
    proof (rule sum.mono_neutral_right)
      show "aA - keys r. lookup r a * a = 0"
        by (simp add: in_keys_iff)
    qed
    finally show "p = ideal.rep r" by (simp only: ideal.rep_def)
  next
    from 2 A  B show "keys r  B" by (rule subset_trans)
  qed
qed

context fixes fs :: "('a 0 'b::field) list"
begin

definition sig_inv_set' :: "nat  ('t 0 'b) set"
  where "sig_inv_set' j = {r. keys (vectorize_poly r)  {0..<j}}"

abbreviation "sig_inv_set  sig_inv_set' (length fs)"

definition rep_list :: "('t 0 'b)  ('a 0 'b)"
  where "rep_list r = ideal.rep (pm_of_idx_pm fs (vectorize_poly r))"

lemma sig_inv_setI: "keys (vectorize_poly r)  {0..<j}  r  sig_inv_set' j"
  by (simp add: sig_inv_set'_def)

lemma sig_inv_setD: "r  sig_inv_set' j  keys (vectorize_poly r)  {0..<j}"
  by (simp add: sig_inv_set'_def)

lemma sig_inv_setI':
  assumes "v. v  keys r  component_of_term v < j"
  shows "r  sig_inv_set' j"
proof (rule sig_inv_setI, rule)
  fix k
  assume "k  keys (vectorize_poly r)"
  then obtain v where "v  keys r" and k: "k = component_of_term v" unfolding keys_vectorize_poly ..
  from this(1) have "k < j" unfolding k by (rule assms)
  thus "k  {0..<j}" by simp
qed

lemma sig_inv_setD':
  assumes "r  sig_inv_set' j" and "v  keys r"
  shows "component_of_term v < j"
proof -
  from assms(2) have "component_of_term v  component_of_term ` keys r" by (rule imageI)
  also have "... = keys (vectorize_poly r)" by (simp only: keys_vectorize_poly)
  also from assms(1) have "...  {0..<j}" by (rule sig_inv_setD)
  finally show ?thesis by simp
qed

corollary sig_inv_setD_lt:
  assumes "r  sig_inv_set' j" and "r  0"
  shows "component_of_term (lt r) < j"
  by (rule sig_inv_setD', fact, rule lt_in_keys, fact)

lemma sig_inv_set_mono:
  assumes "i  j"
  shows "sig_inv_set' i  sig_inv_set' j"
proof
  fix r
  assume "r  sig_inv_set' i"
  hence "keys (vectorize_poly r)  {0..<i}" by (rule sig_inv_setD)
  also from assms have "...  {0..<j}" by fastforce
  finally show "r  sig_inv_set' j" by (rule sig_inv_setI)
qed

lemma sig_inv_set_zero: "0  sig_inv_set' j"
  by (rule sig_inv_setI', simp)

lemma sig_inv_set_closed_uminus: "r  sig_inv_set' j  - r  sig_inv_set' j"
  by (auto dest!: sig_inv_setD' intro!: sig_inv_setI' simp: keys_uminus)

lemma sig_inv_set_closed_plus:
  assumes "r  sig_inv_set' j" and "s  sig_inv_set' j"
  shows "r + s  sig_inv_set' j"
proof (rule sig_inv_setI')
  fix v
  assume "v  keys (r + s)"
  hence "v  keys r  keys s" using Poly_Mapping.keys_add ..
  thus "component_of_term v < j"
  proof
    assume "v  keys r"
    with assms(1) show ?thesis by (rule sig_inv_setD')
  next
    assume "v  keys s"
    with assms(2) show ?thesis by (rule sig_inv_setD')
  qed
qed

lemma sig_inv_set_closed_minus:
  assumes "r  sig_inv_set' j" and "s  sig_inv_set' j"
  shows "r - s  sig_inv_set' j"
proof (rule sig_inv_setI')
  fix v
  assume "v  keys (r - s)"
  hence "v  keys r  keys s" using keys_minus ..
  thus "component_of_term v < j"
  proof
    assume "v  keys r"
    with assms(1) show ?thesis by (rule sig_inv_setD')
  next
    assume "v  keys s"
    with assms(2) show ?thesis by (rule sig_inv_setD')
  qed
qed

lemma sig_inv_set_closed_monom_mult:
  assumes "r  sig_inv_set' j"
  shows "monom_mult c t r  sig_inv_set' j"
proof (rule sig_inv_setI')
  fix v
  assume "v  keys (monom_mult c t r)"
  hence "v  (⊕) t ` keys r" using keys_monom_mult_subset ..
  then obtain u where "u  keys r" and v: "v = t  u" ..
  from assms this(1) have "component_of_term u < j" by (rule sig_inv_setD')
  thus "component_of_term v < j" by (simp add: v term_simps)
qed

lemma sig_inv_set_closed_mult_scalar:
  assumes "r  sig_inv_set' j"
  shows "p  r  sig_inv_set' j"
proof (rule sig_inv_setI')
  fix v
  assume "v  keys (p  r)"
  then obtain t u where "u  keys r" and v: "v = t  u" by (rule in_keys_mult_scalarE)
  from assms this(1) have "component_of_term u < j" by (rule sig_inv_setD')
  thus "component_of_term v < j" by (simp add: v term_simps)
qed

lemma rep_list_zero: "rep_list 0 = 0"
  by (simp add: rep_list_def vectorize_zero)

lemma rep_list_uminus: "rep_list (- r) = - rep_list r"
  by (simp add: rep_list_def vectorize_uminus pm_of_idx_pm_uminus)

lemma rep_list_plus: "rep_list (r + s) = rep_list r + rep_list s"
  by (simp add: rep_list_def vectorize_plus pm_of_idx_pm_plus ideal.rep_plus)

lemma rep_list_minus: "rep_list (r - s) = rep_list r - rep_list s"
  by (simp add: rep_list_def vectorize_minus pm_of_idx_pm_minus ideal.rep_minus)

lemma vectorize_mult_scalar:
  "vectorize_poly (p  q) = MPoly_Type_Class.punit.monom_mult p 0 (vectorize_poly q)"
  by (rule poly_mapping_eqI, simp add: lookup_vectorize_poly MPoly_Type_Class.punit.lookup_monom_mult_zero proj_mult_scalar)

lemma rep_list_mult_scalar: "rep_list (c  r) = c * rep_list r"
  by (simp add: rep_list_def vectorize_mult_scalar pm_of_idx_pm_monom_mult punit.rep_mult_scalar[simplified])

lemma rep_list_monom_mult: "rep_list (monom_mult c t r) = punit.monom_mult c t (rep_list r)"
  unfolding mult_scalar_monomial[symmetric] times_monomial_left[symmetric] by (rule rep_list_mult_scalar)

lemma rep_list_monomial:
  assumes "distinct fs"
  shows "rep_list (monomial c u) =
            (punit.monom_mult c (pp_of_term u) (fs ! (component_of_term u))
              when component_of_term u < length fs)"
  by (simp add: rep_list_def vectorize_monomial pm_of_idx_pm_monomial[OF assms] when_def times_monomial_left)

lemma rep_list_in_ideal_sig_inv_set:
  assumes "r  sig_inv_set' j"
  shows "rep_list r  ideal (set (take j fs))"
proof -
  let ?fs = "take j fs"
  from assms have "keys (vectorize_poly r)  {0..<j}" by (rule sig_inv_setD)
  hence eq: "pm_of_idx_pm fs (vectorize_poly r) = pm_of_idx_pm ?fs (vectorize_poly r)"
    by (simp only: pm_of_idx_pm_take)
  have "rep_list r  ideal (keys (pm_of_idx_pm fs (vectorize_poly r)))"
    unfolding rep_list_def by (rule ideal.rep_in_span)
  also have "... = ideal (keys (pm_of_idx_pm ?fs (vectorize_poly r)))" by (simp only: eq)
  also from keys_pm_of_idx_pm_subset have "...  ideal (set ?fs)" by (rule ideal.span_mono)
  finally show ?thesis .
qed

corollary rep_list_subset_ideal_sig_inv_set:
  "B  sig_inv_set' j  rep_list ` B  ideal (set (take j fs))"
  by (auto dest: rep_list_in_ideal_sig_inv_set)

lemma rep_list_in_ideal: "rep_list r  ideal (set fs)"
proof -
  have "rep_list r  ideal (keys (pm_of_idx_pm fs (vectorize_poly r)))"
    unfolding rep_list_def by (rule ideal.rep_in_span)
  also from keys_pm_of_idx_pm_subset have "...  ideal (set fs)" by (rule ideal.span_mono)
  finally show ?thesis .
qed

corollary rep_list_subset_ideal: "rep_list ` B  ideal (set fs)"
  by (auto intro: rep_list_in_ideal)

lemma in_idealE_rep_list:
  assumes "p  ideal (set fs)"
  obtains r where "p = rep_list r" and "r  sig_inv_set"
proof -
  from assms obtain r0 where r0: "keys r0  set fs" and p: "p = ideal.rep r0"
    by (rule ideal.spanE_rep)
  show ?thesis
  proof
    show "p = rep_list (atomize_poly (idx_pm_of_pm fs r0))"
      by (simp add: rep_list_def vectorize_atomize_poly pm_of_idx_pm_of_pm[OF r0] p)
  next
    show "atomize_poly (idx_pm_of_pm fs r0)  sig_inv_set"
      by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_idx_pm_of_pm_subset)
  qed
qed

lemma keys_rep_list_subset:
  assumes "t  keys (rep_list r)"
  obtains v s where "v  keys r" and "s  Keys (set fs)" and "t = pp_of_term v + s"
proof -
  from assms obtain v0 s where v0: "v0  Keys (Poly_Mapping.range (pm_of_idx_pm fs (vectorize_poly r)))"
    and s: "s  Keys (keys (pm_of_idx_pm fs (vectorize_poly r)))" and t: "t = v0 + s"
    unfolding rep_list_def by (rule punit.keys_rep_subset[simplified])
  note s
  also from keys_pm_of_idx_pm_subset have "Keys (keys (pm_of_idx_pm fs (vectorize_poly r)))  Keys (set fs)"
    by (rule Keys_mono)
  finally have "s  Keys (set fs)" .
  note v0
  also from range_pm_of_idx_pm_subset'
  have "Keys (Poly_Mapping.range (pm_of_idx_pm fs (vectorize_poly r)))  Keys (Poly_Mapping.range (vectorize_poly r))"
    by (rule Keys_mono)
  also have "... = pp_of_term ` keys r" by (fact Keys_range_vectorize_poly)
  finally obtain v where "v  keys r" and "v0 = pp_of_term v" ..
  from this(2) have "t = pp_of_term v + s" by (simp only: t)
  with v  keys r s  Keys (set fs) show ?thesis ..
qed

lemma dgrad_p_set_le_rep_list:
  assumes "dickson_grading d" and "dgrad_set_le d (pp_of_term ` keys r) (Keys (set fs))"
  shows "punit.dgrad_p_set_le d {rep_list r} (set fs)"
proof (simp add: punit.dgrad_p_set_le_def Keys_insert, rule dgrad_set_leI)
  fix t
  assume "t  keys (rep_list r)"
  then obtain v s1 where "v  keys r" and "s1  Keys (set fs)" and t: "t = pp_of_term v + s1"
    by (rule keys_rep_list_subset)
  from this(1) have "pp_of_term v  pp_of_term ` keys r" by fastforce
  with assms(2) obtain s2 where "s2  Keys (set fs)" and "d (pp_of_term v)  d s2"
    by (rule dgrad_set_leE)
  from assms(1) have "d t = ord_class.max (d (pp_of_term v)) (d s1)" unfolding t
    by (rule dickson_gradingD1)
  hence "d t = d (pp_of_term v)  d t = d s1" by (simp add: ord_class.max_def)
  thus "sKeys (set fs). d t  d s"
  proof
    assume "d t = d (pp_of_term v)"
    with d (pp_of_term v)  d s2 have "d t  d s2" by simp
    with s2  Keys (set fs) show ?thesis ..
  next
    assume "d t = d s1"
    hence "d t  d s1" by simp
    with s1  Keys (set fs) show ?thesis ..
  qed
qed

corollary dgrad_p_set_le_rep_list_image:
  assumes "dickson_grading d" and "dgrad_set_le d (pp_of_term ` Keys F) (Keys (set fs))"
  shows "punit.dgrad_p_set_le d (rep_list ` F) (set fs)"
proof (rule punit.dgrad_p_set_leI, elim imageE, simp)
  fix f
  assume "f  F"
  have "pp_of_term ` keys f  pp_of_term ` Keys F" by (rule image_mono, rule keys_subset_Keys, fact)
  hence "dgrad_set_le d (pp_of_term ` keys f) (pp_of_term ` Keys F)" by (rule dgrad_set_le_subset)
  hence "dgrad_set_le d (pp_of_term ` keys f) (Keys (set fs))" using assms(2) by (rule dgrad_set_le_trans)
  with assms(1) show "punit.dgrad_p_set_le d {rep_list f} (set fs)" by (rule dgrad_p_set_le_rep_list)
qed
term Max

definition dgrad_max :: "('a  nat)  nat"
  where "dgrad_max d = (Max (d ` (insert 0 (Keys (set fs)))))"

abbreviation "dgrad_max_set d  dgrad_p_set d (dgrad_max d)"
abbreviation "punit_dgrad_max_set d  punit.dgrad_p_set d (dgrad_max d)"

lemma dgrad_max_0: "d 0  dgrad_max d"
proof -
  from finite_Keys have "finite (d ` insert 0 (Keys (set fs)))" by auto
  moreover have "d 0  d ` insert 0 (Keys (set fs))" by blast
  ultimately show ?thesis unfolding dgrad_max_def by (rule Max_ge)
qed

lemma dgrad_max_1: "set fs  punit_dgrad_max_set d"
proof (cases "Keys (set fs) = {}")
  case True
  show ?thesis
  proof (rule, rule punit.dgrad_p_setI[simplified])
    fix f v
    assume "f  set fs" and "v  keys f"
    with True show "d v  dgrad_max d" by (auto simp: Keys_def)
  qed
next
  case False
  show ?thesis
  proof (rule subset_trans)
    from finite_set show "set fs  punit.dgrad_p_set d (Max (d ` (Keys (set fs))))"
      by (rule punit.dgrad_p_set_exhaust_expl[simplified])
  next
    from finite_set have "finite (Keys (set fs))" by (rule finite_Keys)
    hence "finite (d ` Keys (set fs))" by (rule finite_imageI)
    moreover from False have 2: "d ` Keys (set fs)  {}" by simp
    ultimately have "dgrad_max d = ord_class.max (d 0) (Max (d ` Keys (set fs)))"
      by (simp add: dgrad_max_def)
    hence "Max (d ` (Keys (set fs)))  dgrad_max d" by simp
    thus "punit.dgrad_p_set d (Max (d ` (Keys (set fs))))  punit_dgrad_max_set d"
      by (rule punit.dgrad_p_set_subset)
  qed
qed

lemma dgrad_max_2:
  assumes "dickson_grading d" and "r  dgrad_max_set d"
  shows "rep_list r  punit_dgrad_max_set d"
proof (rule punit.dgrad_p_setI[simplified])
  fix t
  assume "t  keys (rep_list r)"
  then obtain v s where "v  keys r" and "s  Keys (set fs)" and t: "t = pp_of_term v + s"
    by (rule keys_rep_list_subset)
  from assms(2) v  keys r have "d (pp_of_term v)  dgrad_max d" by (rule dgrad_p_setD)
  moreover have "d s  dgrad_max d" by (simp add: s  Keys (set fs) dgrad_max_def finite_Keys)
  ultimately show "d t  dgrad_max d" by (simp add: t dickson_gradingD1[OF assms(1)])
qed

corollary dgrad_max_3:
  assumes "dickson_grading d" and "F  dgrad_max_set d"
  shows "rep_list ` F  punit_dgrad_max_set d"
proof (rule, elim imageE, simp)
  fix f
  assume "f  F"
  hence "f  dgrad_p_set d (dgrad_max d)" using assms(2) ..
  with assms(1) show "rep_list f  punit.dgrad_p_set d (dgrad_max d)" by (rule dgrad_max_2)
qed

lemma punit_dgrad_max_set_subset_dgrad_p_set:
  assumes "dickson_grading d" and "set fs  punit.dgrad_p_set d m" and "¬ set fs  {0}"
  shows "punit_dgrad_max_set d  punit.dgrad_p_set d m"
proof (rule punit.dgrad_p_set_subset)
  show "dgrad_max d  m" unfolding dgrad_max_def
  proof (rule Max.boundedI)
    show "finite (d ` insert 0 (Keys (set fs)))" by (simp add: finite_Keys)
  next
    show "d ` insert 0 (Keys (set fs))  {}" by simp
  next
    fix a
    assume "a  d ` insert 0 (Keys (set fs))"
    then obtain t where "t  insert 0 (Keys (set fs))" and "a = d t" ..
    from this(1) show "a  m" unfolding a = d t
    proof
      assume "t = 0"
      from assms(3) obtain f where "f  set fs" and "f  0" by auto
      from this(1) assms(2) have "f  punit.dgrad_p_set d m" ..
      from f  0 have "keys f  {}" by simp
      then obtain s where "s  keys f" by blast
      have "d s = d (t + s)" by (simp add: t = 0)
      also from assms(1) have "... = ord_class.max (d t) (d s)" by (rule dickson_gradingD1)
      finally have "d t  d s" by (simp add: max_def)
      also from f  punit.dgrad_p_set d m s  keys f have "...  m"
        by (rule punit.dgrad_p_setD[simplified])
      finally show "d t  m" .
    next
      assume "t  Keys (set fs)"
      then obtain f where "f  set fs" and "t  keys f" by (rule in_KeysE)
      from this(1) assms(2) have "f  punit.dgrad_p_set d m" ..
      thus "d t  m" using t  keys f by (rule punit.dgrad_p_setD[simplified])
    qed
  qed
qed

definition dgrad_sig_set' :: "nat  ('a  nat)  ('t 0 'b) set"
  where "dgrad_sig_set' j d = dgrad_max_set d  sig_inv_set' j"

abbreviation "dgrad_sig_set  dgrad_sig_set' (length fs)"

lemma dgrad_sig_set_set_mono: "i  j  dgrad_sig_set' i d  dgrad_sig_set' j d"
  by (auto simp: dgrad_sig_set'_def dest: sig_inv_set_mono)

lemma dgrad_sig_set_closed_uminus: "r  dgrad_sig_set' j d  - r  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_uminus sig_inv_set_closed_uminus)

lemma dgrad_sig_set_closed_plus:
  "r  dgrad_sig_set' j d  s  dgrad_sig_set' j d  r + s  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_plus sig_inv_set_closed_plus)

lemma dgrad_sig_set_closed_minus:
  "r  dgrad_sig_set' j d  s  dgrad_sig_set' j d  r - s  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_minus sig_inv_set_closed_minus)

lemma dgrad_sig_set_closed_monom_mult:
  assumes "dickson_grading d" and "d t  dgrad_max d"
  shows "p  dgrad_sig_set' j d  monom_mult c t p  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: assms dgrad_p_set_closed_monom_mult sig_inv_set_closed_monom_mult)

lemma dgrad_sig_set_closed_monom_mult_zero:
  "p  dgrad_sig_set' j d  monom_mult c 0 p  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_monom_mult_zero sig_inv_set_closed_monom_mult)

lemma dgrad_sig_set_closed_mult_scalar:
  "dickson_grading d  p  punit_dgrad_max_set d  r  dgrad_sig_set' j d  p  r  dgrad_sig_set' j d"
  unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_mult_scalar sig_inv_set_closed_mult_scalar)

lemma dgrad_sig_set_closed_monomial:
  assumes "d (pp_of_term u)  dgrad_max d" and "component_of_term u < j"
  shows "monomial c u  dgrad_sig_set' j d"
proof (simp add: dgrad_sig_set'_def, rule)
  show "monomial c u  dgrad_max_set d"
  proof (rule dgrad_p_setI)
    fix v
    assume "v  keys (monomial c u)"
    also have "...  {u}" by simp
    finally show "d (pp_of_term v)  dgrad_max d" using assms(1) by simp
  qed
next
  show "monomial c u  sig_inv_set' j"
  proof (rule sig_inv_setI')
    fix v
    assume "v  keys (monomial c u)"
    also have "...  {u}" by simp
    finally show "component_of_term v < j" using assms(2) by simp
  qed
qed

lemma rep_list_in_ideal_dgrad_sig_set:
  "r  dgrad_sig_set' j d  rep_list r  ideal (set (take j fs))"
  by (auto simp: dgrad_sig_set'_def dest: rep_list_in_ideal_sig_inv_set)

lemma in_idealE_rep_list_dgrad_sig_set_take:
  assumes "hom_grading d" and "p  punit_dgrad_max_set d" and "p  ideal (set (take j fs))"
  obtains r where "r  dgrad_sig_set d" and "r  dgrad_sig_set' j d" and "p = rep_list r"
proof -
  let ?fs = "take j fs"
  from set_take_subset dgrad_max_1 have "set ?fs  punit_dgrad_max_set d"
    by (rule subset_trans)
  with assms(1) obtain r0 where r0: "keys r0  set ?fs"
    and 1: "Poly_Mapping.range r0  punit_dgrad_max_set d" and p: "p = ideal.rep r0"
    using assms(2, 3) by (rule in_idealE_rep_dgrad_p_set)
  define q where "q = idx_pm_of_pm ?fs r0"
  have "keys q  {0..<length ?fs}" unfolding q_def by (rule keys_idx_pm_of_pm_subset)
  also have "...  {0..<j}" by fastforce
  finally have keys_q: "keys q  {0..<j}" .
  have *: "atomize_poly q  dgrad_max_set d"
  proof
    fix v
    assume "v  keys (atomize_poly q)"
    then obtain i where i: "i  keys q"
      and v_in: "v  (λt. term_of_pair (t, i)) ` keys (lookup q i)"
      unfolding keys_atomize_poly ..
    from i keys_idx_pm_of_pm_subset[of ?fs r0] have "i < length ?fs" by (auto simp: q_def)
    from v_in obtain t where "t  keys (lookup q i)" and v: "v = term_of_pair (t, i)" ..
    from this(1) i < length ?fs have t: "t  keys (lookup r0 (?fs ! i))"
      by (simp add: lookup_idx_pm_of_pm q_def)
    hence "lookup r0 (?fs ! i)  0" by fastforce
    hence "lookup r0 (?fs ! i)  Poly_Mapping.range r0" by (simp add: in_keys_iff)
    hence "lookup r0 (?fs ! i)  punit_dgrad_max_set d" using 1 ..
    hence "d t  dgrad_max d" using t by (rule punit.dgrad_p_setD[simplified])
    thus "d (pp_of_term v)  dgrad_max d" by (simp add: v pp_of_term_of_pair)
  qed
  show ?thesis
  proof
    have "atomize_poly q  sig_inv_set' j"
      by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_q)
    with * show "atomize_poly q  dgrad_sig_set' j d" unfolding dgrad_sig_set'_def ..
  next
    from keys q  {0..<length ?fs} have keys_q': "keys q  {0..<length fs}" by auto
    have "atomize_poly q  sig_inv_set"
      by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_q')
    with * show "atomize_poly q  dgrad_sig_set d" unfolding dgrad_sig_set'_def ..
  next
    from keys_q have "pm_of_idx_pm fs q = pm_of_idx_pm ?fs q" by (simp only: pm_of_idx_pm_take)
    thus "p = rep_list (atomize_poly q)"
      by (simp add: rep_list_def vectorize_atomize_poly pm_of_idx_pm_of_pm[OF r0] p q_def)
  qed
qed

corollary in_idealE_rep_list_dgrad_sig_set:
  assumes "hom_grading d" and "p  punit_dgrad_max_set d" and "p  ideal (set fs)"
  obtains r where "r  dgrad_sig_set d" and "p = rep_list r"
proof -
  from assms(3) have "p  ideal (set (take (length fs) fs))" by simp
  with assms(1, 2) obtain r where "r  dgrad_sig_set d" and "p = rep_list r"
    by (rule in_idealE_rep_list_dgrad_sig_set_take)
  thus ?thesis ..
qed

lemma dgrad_sig_setD_lp:
  assumes "p  dgrad_sig_set' j d"
  shows "d (lp p)  dgrad_max d"
proof (cases "p = 0")
  case True
  show ?thesis by (simp add: True min_term_def pp_of_term_of_pair dgrad_max_0)
next
  case False
  from assms have "p  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  thus ?thesis using False by (rule dgrad_p_setD_lp)
qed

lemma dgrad_sig_setD_lt:
  assumes "p  dgrad_sig_set' j d" and "p  0"
  shows "component_of_term (lt p) < j"
proof -
  from assms have "p  sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
  thus ?thesis using assms(2) by (rule sig_inv_setD_lt)
qed

lemma dgrad_sig_setD_rep_list_lt:
  assumes "dickson_grading d" and "p  dgrad_sig_set' j d"
  shows "d (punit.lt (rep_list p))  dgrad_max d"
proof (cases "rep_list p = 0")
  case True
  show ?thesis by (simp add: True dgrad_max_0)
next
  case False
  from assms(2) have "p  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  with assms(1) have "rep_list p  punit_dgrad_max_set d" by (rule dgrad_max_2)
  thus ?thesis using False by (rule punit.dgrad_p_setD_lp[simplified])
qed

definition spp_of :: "('t 0 'b)  ('t × ('a 0 'b))"
  where "spp_of r = (lt r, rep_list r)"

text ‹``spp'' stands for ``sig-poly-pair''.›

lemma fst_spp_of: "fst (spp_of r) = lt r"
  by (simp add: spp_of_def)

lemma snd_spp_of: "snd (spp_of r) = rep_list r"
  by (simp add: spp_of_def)

subsubsection ‹Signature Reduction›

lemma term_is_le_rel_canc_left:
  assumes "ord_term_lin.is_le_rel rel"
  shows "rel (t  u) (t  v)  rel u v"
  using assms
  by (rule ord_term_lin.is_le_relE,
      auto simp: splus_left_canc dest: ord_term_canc ord_term_strict_canc splus_mono splus_mono_strict)

lemma term_is_le_rel_minus:
  assumes "ord_term_lin.is_le_rel rel" and "s adds t"
  shows "rel ((t - s)  u) v  rel (t  u) (s  v)"
proof -
  from assms(2) have eq: "s + (t - s) = t" unfolding add.commute[of s] by (rule adds_minus)
  from assms(1) have "rel ((t - s)  u) v = rel (s  ((t - s)  u)) (s  v)"
    by (simp only: term_is_le_rel_canc_left)
  also have "... = rel (t  u) (s  v)" by (simp only: splus_assoc[symmetric] eq)
  finally show ?thesis .
qed

lemma term_is_le_rel_minus_minus:
  assumes "ord_term_lin.is_le_rel rel" and "a adds t" and "b adds t"
  shows "rel ((t - a)  u) ((t - b)  v)  rel (b  u) (a  v)"
proof -
  from assms(2) have eq1: "a + (t - a) = t" unfolding add.commute[of a] by (rule adds_minus)
  from assms(3) have eq2: "b + (t - b) = t" unfolding add.commute[of b] by (rule adds_minus)
  from assms(1) have "rel ((t - a)  u) ((t - b)  v) = rel ((a + b)  ((t - a)  u)) ((a + b)  ((t - b)  v))"
    by (simp only: term_is_le_rel_canc_left)
  also have "... = rel ((t + b)  u) ((t + a)  v)" unfolding splus_assoc[symmetric]
    by (metis (no_types, lifting) add.assoc add.commute eq1 eq2)
  also from assms(1) have "... = rel (b  u) (a  v)" by (simp only: splus_assoc term_is_le_rel_canc_left)
  finally show ?thesis .
qed

lemma pp_is_le_rel_canc_right:
  assumes "ordered_powerprod_lin.is_le_rel rel"
  shows "rel (s + u) (t + u)  rel s t"
  using assms
  by (rule ordered_powerprod_lin.is_le_relE, auto dest: ord_canc ord_strict_canc plus_monotone plus_monotone_strict)

lemma pp_is_le_rel_canc_left: "ordered_powerprod_lin.is_le_rel rel  rel (t + u) (t + v)  rel u v"
  by (simp add: add.commute[of t] pp_is_le_rel_canc_right)

definition sig_red_single :: "('t  't  bool)  ('a  'a  bool)  ('t 0 'b)  ('t 0 'b)  ('t 0 'b)  'a  bool"
  where "sig_red_single sing_reg top_tail p q f t 
                (rep_list f  0  lookup (rep_list p) (t + punit.lt (rep_list f))  0 
                 q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f 
                 ord_term_lin.is_le_rel sing_reg  ordered_powerprod_lin.is_le_rel top_tail 
                 sing_reg (t  lt f) (lt p)  top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p)))"

text ‹The first two parameters of @{const sig_red_single}, sing_reg› and top_tail›, specify whether
  the reduction is a singular/regular/arbitrary top/tail/arbitrary signature-reduction.
   If sing_reg› is @{const HOL.eq}, the reduction is singular.
   If sing_reg› is @{term "(≺t)"}, the reduction is regular.
   If sing_reg› is @{term "(≼t)"}, the reduction is an arbitrary signature-reduction.
   If top_tail› is @{const HOL.eq}, it is a top reduction.
   If top_tail› is @{term "(≺)"}, it is a tail reduction.
   If top_tail› is @{term "(≼)"}, the reduction is an arbitrary signature-reduction.›

definition sig_red :: "('t  't  bool)  ('a  'a  bool)  ('t 0 'b) set  ('t 0 'b)  ('t 0 'b)  bool"
  where "sig_red sing_reg top_tail F p q  (fF. t. sig_red_single sing_reg top_tail p q f t)"

definition is_sig_red :: "('t  't  bool)  ('a  'a  bool)  ('t 0 'b) set  ('t 0 'b)  bool"
  where "is_sig_red sing_reg top_tail F p  (q. sig_red sing_reg top_tail F p q)"

lemma sig_red_singleI:
  assumes "rep_list f  0" and "t + punit.lt (rep_list f)  keys (rep_list p)"
    and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    and "ord_term_lin.is_le_rel sing_reg" and "ordered_powerprod_lin.is_le_rel top_tail"
    and "sing_reg (t  lt f) (lt p)"
    and "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
  shows "sig_red_single sing_reg top_tail p q f t"
  unfolding sig_red_single_def using assms by blast

lemma sig_red_singleD1:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "rep_list f  0"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD2:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "t + punit.lt (rep_list f)  keys (rep_list p)"
  using assms unfolding sig_red_single_def by (simp add: in_keys_iff)

lemma sig_red_singleD3:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD4:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "ord_term_lin.is_le_rel sing_reg"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD5:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "ordered_powerprod_lin.is_le_rel top_tail"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD6:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "sing_reg (t  lt f) (lt p)"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD7:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
  using assms unfolding sig_red_single_def by blast

lemma sig_red_singleD8:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "t  lt f t lt p"
proof -
  from assms have "ord_term_lin.is_le_rel sing_reg" and "sing_reg (t  lt f) (lt p)"
    by (rule sig_red_singleD4, rule sig_red_singleD6)
  thus ?thesis by (rule ord_term_lin.is_le_rel_le)
qed

lemma sig_red_singleD9:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "t + punit.lt (rep_list f)  punit.lt (rep_list p)"
proof -
  from assms have "ordered_powerprod_lin.is_le_rel top_tail"
    and "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
    by (rule sig_red_singleD5, rule sig_red_singleD7)
  thus ?thesis by (rule ordered_powerprod_lin.is_le_rel_le)
qed

lemmas sig_red_singleD = sig_red_singleD1 sig_red_singleD2 sig_red_singleD3 sig_red_singleD4
                         sig_red_singleD5 sig_red_singleD6 sig_red_singleD7 sig_red_singleD8 sig_red_singleD9

lemma sig_red_single_red_single:
  "sig_red_single sing_reg top_tail p q f t  punit.red_single (rep_list p) (rep_list q) (rep_list f) t"
  by (simp add: sig_red_single_def punit.red_single_def rep_list_minus rep_list_monom_mult)

lemma sig_red_single_regular_lt:
  assumes "sig_red_single (≺t) top_tail p q f t"
  shows "lt q = lt p"
proof -
  let ?f = "monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
  from assms have lt: "t  lt f t lt p" and q: "q = p - ?f"
    by (rule sig_red_singleD6, rule sig_red_singleD3)
  from lt_monom_mult_le lt have "lt ?f t lt p" by (rule ord_term_lin.order.strict_trans1)
  thus ?thesis unfolding q by (rule lt_minus_eqI_2)
qed

lemma sig_red_single_regular_lc:
  assumes "sig_red_single (≺t) top_tail p q f t"
  shows "lc q = lc p"
proof -
  from assms have "lt q = lt p" by (rule sig_red_single_regular_lt)
  from assms have lt: "t  lt f t lt p"
    and q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    (is "_ = _ - ?f") by (rule sig_red_singleD6, rule sig_red_singleD3)
  from lt_monom_mult_le lt have "lt ?f t lt p" by (rule ord_term_lin.order.strict_trans1)
  hence "lookup ?f (lt p) = 0" using lt_max ord_term_lin.leD by blast
  thus ?thesis unfolding lc_def lt q = lt p by (simp add: q lookup_minus)
qed

lemma sig_red_single_lt:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "lt q t lt p"
proof -
  from assms have lt: "t  lt f t lt p"
    and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    by (rule sig_red_singleD8, rule sig_red_singleD3)
  from this(2) have q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    (is "_ = _ + ?f") by (simp add: monom_mult_uminus_left)
  from lt_monom_mult_le lt have 1: "lt ?f t lt p" by (rule ord_term_lin.order.trans)
  have "lt q t ord_term_lin.max (lt p) (lt ?f)" unfolding q by (fact lt_plus_le_max)
  also from 1 have "ord_term_lin.max (lt p) (lt ?f) = lt p" by (rule ord_term_lin.max.absorb1)
  finally show ?thesis .
qed

lemma sig_red_single_lt_rep_list:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "punit.lt (rep_list q)  punit.lt (rep_list p)"
proof -
  from assms have "punit.red_single (rep_list p) (rep_list q) (rep_list f) t"
    by (rule sig_red_single_red_single)
  hence "punit.ord_strict_p (rep_list q) (rep_list p)" by (rule punit.red_single_ord)
  hence "punit.ord_p (rep_list q) (rep_list p)" by simp
  thus ?thesis by (rule punit.ord_p_lt)
qed

lemma sig_red_single_tail_lt_in_keys_rep_list:
  assumes "sig_red_single sing_reg (≺) p q f t"
  shows "punit.lt (rep_list p)  keys (rep_list q)"
proof -
  from assms have "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    by (rule sig_red_singleD3)
  hence q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    by (simp add: monom_mult_uminus_left)
  show ?thesis unfolding q rep_list_plus rep_list_monom_mult
  proof (rule in_keys_plusI1)
    from assms have "t + punit.lt (rep_list f)  keys (rep_list p)" by (rule sig_red_singleD2)
    hence "rep_list p  0" by auto
    thus "punit.lt (rep_list p)  keys (rep_list p)" by (rule punit.lt_in_keys)
  next
    show "punit.lt (rep_list p) 
      keys (punit.monom_mult (- lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t (rep_list f))"
        (is "_  keys ?f")
    proof
      assume "punit.lt (rep_list p)  keys ?f"
      hence "punit.lt (rep_list p)  punit.lt ?f" by (rule punit.lt_max_keys)
      also have "...  t + punit.lt (rep_list f)" by (fact punit.lt_monom_mult_le[simplified])
      also from assms have "...  punit.lt (rep_list p)" by (rule sig_red_singleD7)
      finally show False by simp
    qed
  qed
qed

corollary sig_red_single_tail_lt_rep_list:
  assumes "sig_red_single sing_reg (≺) p q f t"
  shows "punit.lt (rep_list q) = punit.lt (rep_list p)"
proof (rule ordered_powerprod_lin.order_antisym)
  from assms show "punit.lt (rep_list q)  punit.lt (rep_list p)" by (rule sig_red_single_lt_rep_list)
next
  from assms have "punit.lt (rep_list p)  keys (rep_list q)" by (rule sig_red_single_tail_lt_in_keys_rep_list)
  thus "punit.lt (rep_list p)  punit.lt (rep_list q)" by (rule punit.lt_max_keys)
qed

lemma sig_red_single_tail_lc_rep_list:
  assumes "sig_red_single sing_reg (≺) p q f t"
  shows "punit.lc (rep_list q) = punit.lc (rep_list p)"
proof -
  from assms have *: "punit.lt (rep_list q) = punit.lt (rep_list p)"
    by (rule sig_red_single_tail_lt_rep_list)
  from assms have lt: "t + punit.lt (rep_list f)  punit.lt (rep_list p)"
    and q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    (is "_ = _ - ?f") by (rule sig_red_singleD7, rule sig_red_singleD3)
  from punit.lt_monom_mult_le[simplified] lt have "punit.lt (rep_list ?f)  punit.lt (rep_list p)"
    unfolding rep_list_monom_mult by (rule ordered_powerprod_lin.order.strict_trans1)
  hence "lookup (rep_list ?f) (punit.lt (rep_list p)) = 0"
    using punit.lt_max ordered_powerprod_lin.leD by blast
  thus ?thesis unfolding punit.lc_def * by (simp add: q lookup_minus rep_list_minus punit.lc_def)
qed

lemma sig_red_single_top_lt_rep_list:
  assumes "sig_red_single sing_reg (=) p q f t" and "rep_list q  0"
  shows "punit.lt (rep_list q)  punit.lt (rep_list p)"
proof -
  from assms(1) have "rep_list f  0" and in_keys: "t + punit.lt (rep_list f)  keys (rep_list p)"
    and lt: "t + punit.lt (rep_list f) = punit.lt (rep_list p)"
    and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    by (rule sig_red_singleD)+
  from this(4) have q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
    (is "_ = _ + monom_mult ?c _ _") by (simp add: monom_mult_uminus_left)
  from rep_list f  0 have "punit.lc (rep_list f)  0" by (rule punit.lc_not_0)
  from assms(2) have *: "rep_list p + punit.monom_mult ?c t (rep_list f)  0"
    by (simp add: q rep_list_plus rep_list_monom_mult)
  from in_keys have "lookup (rep_list p) (t + punit.lt (rep_list f))  0"
    by (simp add: in_keys_iff)
  moreover from rep_list f  0 have "punit.lc (rep_list f)  0" by (rule punit.lc_not_0)
  ultimately have "?c  0" by simp
  hence "punit.lt (punit.monom_mult ?c t (rep_list f)) = t + punit.lt (rep_list f)"
    using rep_list f  0 by (rule lp_monom_mult)
  hence "punit.lt (punit.monom_mult ?c t (rep_list f)) = punit.lt (rep_list p)" by (simp only: lt)
  moreover have "punit.lc (punit.monom_mult ?c t (rep_list f)) = - punit.lc (rep_list p)"
    by (simp add: lt punit.lc_def[symmetric] punit.lc (rep_list f)  0)
  ultimately show ?thesis unfolding rep_list_plus rep_list_monom_mult q by (rule punit.lt_plus_lessI[OF *])
qed

lemma sig_red_single_monom_mult:
  assumes "sig_red_single sing_reg top_tail p q f t" and "c  0"
  shows "sig_red_single sing_reg top_tail (monom_mult c s p) (monom_mult c s q) f (s + t)"
proof -
  from assms(1) have a: "ord_term_lin.is_le_rel sing_reg" and b: "ordered_powerprod_lin.is_le_rel top_tail"
    by (rule sig_red_singleD4, rule sig_red_singleD5)
  have eq1: "(s + t)  lt f = s  (t  lt f)" by (simp only: splus_assoc)
  from assms(1) have 1: "t + punit.lt (rep_list f)  keys (rep_list p)" by (rule sig_red_singleD2)
  hence "rep_list p  0" by auto
  hence "p  0" by (auto simp: rep_list_zero)
  with assms(2) have eq2: "lt (monom_mult c s p) = s  lt p" by (rule lt_monom_mult)
  show ?thesis
  proof (rule sig_red_singleI)
    from assms(1) show "rep_list f  0" by (rule sig_red_singleD1)
  next
    show "s + t + punit.lt (rep_list f)  keys (rep_list (monom_mult c s p))"
      by (auto simp: rep_list_monom_mult punit.keys_monom_mult[OF assms(2)] ac_simps intro: 1)
  next
    from assms(1) have q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
      by (rule sig_red_singleD3)
    show "monom_mult c s q =
          monom_mult c s p -
            monom_mult (lookup (rep_list (monom_mult c s p)) (s + t + punit.lt (rep_list f)) / punit.lc (rep_list f)) (s + t) f"
      by (simp add: q monom_mult_dist_right_minus ac_simps rep_list_monom_mult
          punit.lookup_monom_mult_plus[simplified] monom_mult_assoc)
  next
    from assms(1) have "sing_reg (t  lt f) (lt p)" by (rule sig_red_singleD6)
    thus "sing_reg ((s + t)  lt f) (lt (monom_mult c s p))"
      by (simp only: eq1 eq2 term_is_le_rel_canc_left[OF a])
  next
    from assms(1) have "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
      by (rule sig_red_singleD7)
    thus "top_tail (s + t + punit.lt (rep_list f)) (punit.lt (rep_list (monom_mult c s p)))"
      by (simp add: rep_list_monom_mult punit.lt_monom_mult[OF assms(2) rep_list p  0] add.assoc pp_is_le_rel_canc_left[OF b])
  qed (fact a, fact b)
qed

lemma sig_red_single_sing_reg_cases:
  "sig_red_single (≼t) top_tail p q f t = (sig_red_single (=) top_tail p q f t  sig_red_single (≺t) top_tail p q f t)"
  by (auto simp: sig_red_single_def)

corollary sig_red_single_sing_regI:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "sig_red_single (≼t) top_tail p q f t"
proof -
  from assms have "ord_term_lin.is_le_rel sing_reg" by (rule sig_red_singleD)
  with assms show ?thesis unfolding ord_term_lin.is_le_rel_def
    by (auto simp: sig_red_single_sing_reg_cases)
qed

lemma sig_red_single_top_tail_cases:
  "sig_red_single sing_reg (≼) p q f t = (sig_red_single sing_reg (=) p q f t  sig_red_single sing_reg (≺) p q f t)"
  by (auto simp: sig_red_single_def)

corollary sig_red_single_top_tailI:
  assumes "sig_red_single sing_reg top_tail p q f t"
  shows "sig_red_single sing_reg (≼) p q f t"
proof -
  from assms have "ordered_powerprod_lin.is_le_rel top_tail" by (rule sig_red_singleD)
  with assms show ?thesis unfolding ordered_powerprod_lin.is_le_rel_def
    by (auto simp: sig_red_single_top_tail_cases)
qed

lemma dgrad_max_set_closed_sig_red_single:
  assumes "dickson_grading d" and "p  dgrad_max_set d" and "f  dgrad_max_set d"
    and "sig_red_single sing_red top_tail p q f t"
  shows "q  dgrad_max_set d"
proof -
  let ?f = "monom_mult (lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
  from assms(4) have t: "t + punit.lt (rep_list f)  keys (rep_list p)" and q: "q = p - ?f"
    by (rule sig_red_singleD2, rule sig_red_singleD3)
  from assms(1, 2) have "rep_list p  punit_dgrad_max_set d" by (rule dgrad_max_2)
  show ?thesis unfolding q using assms(2)
  proof (rule dgrad_p_set_closed_minus)
    from assms(1) _ assms(3) show "?f  dgrad_max_set d"
    proof (rule dgrad_p_set_closed_monom_mult)
      from assms(1) have "d t  d (t + punit.lt (rep_list f))" by (simp add: dickson_gradingD1)
      also from rep_list p  punit_dgrad_max_set d t have "...  dgrad_max d"
        by (rule punit.dgrad_p_setD[simplified])
      finally show "d t  dgrad_max d" .
    qed
  qed
qed

lemma sig_inv_set_closed_sig_red_single:
  assumes "p  sig_inv_set" and "f  sig_inv_set" and "sig_red_single sing_red top_tail p q f t"
  shows "q  sig_inv_set"
proof -
  let ?f = "monom_mult (lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
  from assms(3) have t: "t + punit.lt (rep_list f)  keys (rep_list p)" and q: "q = p - ?f"
    by (rule sig_red_singleD2, rule sig_red_singleD3)
  show ?thesis unfolding q using assms(1)
  proof (rule sig_inv_set_closed_minus)
    from assms(2) show "?f  sig_inv_set" by (rule sig_inv_set_closed_monom_mult)
  qed
qed

corollary dgrad_sig_set_closed_sig_red_single:
  assumes "dickson_grading d" and "p  dgrad_sig_set d" and "f  dgrad_sig_set d"
    and "sig_red_single sing_red top_tail p q f t"
  shows "q  dgrad_sig_set d"
  using assms unfolding dgrad_sig_set'_def
  by (auto intro: dgrad_max_set_closed_sig_red_single sig_inv_set_closed_sig_red_single)

lemma sig_red_regular_lt: "sig_red (≺t) top_tail F p q  lt q = lt p"
  by (auto simp: sig_red_def intro: sig_red_single_regular_lt)

lemma sig_red_regular_lc: "sig_red (≺t) top_tail F p q  lc q = lc p"
  by (auto simp: sig_red_def intro: sig_red_single_regular_lc)

lemma sig_red_lt: "sig_red sing_reg top_tail F p q  lt q t lt p"
  by (auto simp: sig_red_def intro: sig_red_single_lt)

lemma sig_red_tail_lt_rep_list: "sig_red sing_reg (≺) F p q  punit.lt (rep_list q) = punit.lt (rep_list p)"
  by (auto simp: sig_red_def intro: sig_red_single_tail_lt_rep_list)

lemma sig_red_tail_lc_rep_list: "sig_red sing_reg (≺) F p q  punit.lc (rep_list q) = punit.lc (rep_list p)"
  by (auto simp: sig_red_def intro: sig_red_single_tail_lc_rep_list)

lemma sig_red_top_lt_rep_list:
  "sig_red sing_reg (=) F p q  rep_list q  0  punit.lt (rep_list q)  punit.lt (rep_list p)"
  by (auto simp: sig_red_def intro: sig_red_single_top_lt_rep_list)

lemma sig_red_lt_rep_list: "sig_red sing_reg top_tail F p q  punit.lt (rep_list q)  punit.lt (rep_list p)"
  by (auto simp: sig_red_def intro: sig_red_single_lt_rep_list)

lemma sig_red_red: "sig_red sing_reg top_tail F p q  punit.red (rep_list ` F) (rep_list p) (rep_list q)"
  by (auto simp: sig_red_def punit.red_def dest: sig_red_single_red_single)

lemma sig_red_monom_mult:
  "sig_red sing_reg top_tail F p q  c  0  sig_red sing_reg top_tail F (monom_mult c s p) (monom_mult c s q)"
  by (auto simp: sig_red_def punit.red_def dest: sig_red_single_monom_mult)

lemma sig_red_sing_reg_cases:
  "sig_red (≼t) top_tail F p q = (sig_red (=) top_tail F p q  sig_red (≺t) top_tail F p q)"
  by (auto simp: sig_red_def sig_red_single_sing_reg_cases)

corollary sig_red_sing_regI: "sig_red sing_reg top_tail F p q  sig_red (≼t) top_tail F p q"
  by (auto simp: sig_red_def intro: sig_red_single_sing_regI)

lemma sig_red_top_tail_cases:
  "sig_red sing_reg (≼) F p q = (sig_red sing_reg (=) F p q  sig_red sing_reg (≺) F p q)"
  by (auto simp: sig_red_def sig_red_single_top_tail_cases)

corollary sig_red_top_tailI: "sig_red sing_reg top_tail F p q  sig_red sing_reg (≼) F p q"
  by (auto simp: sig_red_def intro: sig_red_single_top_tailI)

lemma sig_red_wf_dgrad_max_set:
  assumes "dickson_grading d" and "F  dgrad_max_set d"
  shows "wfP (sig_red sing_reg top_tail F)¯¯"
proof -
  from assms have "rep_list ` F  punit_dgrad_max_set d" by (rule dgrad_max_3)
  with assms(1) have "wfP (punit.red (rep_list ` F))¯¯" by (rule punit.red_wf_dgrad_p_set)
  hence *: "f. i. (punit.red (rep_list ` F))¯¯ (f (Suc i)) (f i)"
    by (simp add: wf_iff_no_infinite_down_chain[to_pred])
  show ?thesis unfolding wf_iff_no_infinite_down_chain[to_pred]
  proof (rule, elim exE)
    fix seq
    assume "i. (sig_red sing_reg top_tail F)¯¯ (seq (Suc i)) (seq i)"
    hence "sig_red sing_reg top_tail F (seq i) (seq (Suc i))" for i by simp
    hence "punit.red (rep_list ` F) ((rep_list  seq) i) ((rep_list  seq) (Suc i))" for i
      by (auto intro: sig_red_red)
    hence "i. (punit.red (rep_list ` F))¯¯ ((rep_list  seq) (Suc i)) ((rep_list  seq) i)" by simp
    hence "f. i. (punit.red (rep_list ` F))¯¯ (f (Suc i)) (f i)" by blast
    with * show False ..
  qed
qed

lemma dgrad_sig_set_closed_sig_red:
  assumes "dickson_grading d" and "F  dgrad_sig_set d" and "p  dgrad_sig_set d"
    and "sig_red sing_red top_tail F p q"
  shows "q  dgrad_sig_set d"
  using assms by (auto simp: sig_red_def intro: dgrad_sig_set_closed_sig_red_single)

lemma sig_red_mono: "sig_red sing_reg top_tail F p q  F  F'  sig_red sing_reg top_tail F' p q"
  by (auto simp: sig_red_def)

lemma sig_red_Un:
  "sig_red sing_reg top_tail (A  B) p q  (sig_red sing_reg top_tail A p q  sig_red sing_reg top_tail B p q)"
  by (auto simp: sig_red_def)

lemma sig_red_subset:
  assumes "sig_red sing_reg top_tail F p q" and "sing_reg = (≼t)  sing_reg = (≺t)"
  shows "sig_red sing_reg top_tail {fF. sing_reg (lt f) (lt p)} p q"
proof -
  from assms(1) obtain f t where "f  F" and *: "sig_red_single sing_reg top_tail p q f t"
    unfolding sig_red_def by blast
  have "lt f = 0  lt f" by (simp only: term_simps)
  also from zero_min have "... t t  lt f" by (rule splus_mono_left)
  finally have 1: "lt f t t  lt f" .
  from * have 2: "sing_reg (t  lt f) (lt p)" by (rule sig_red_singleD6)
  from assms(2) have "sing_reg (lt f) (lt p)"
  proof
    assume "sing_reg = (≼t)"
    with 1 2 show ?thesis by simp
  next
    assume "sing_reg = (≺t)"
    with 1 2 show ?thesis by simp
  qed
  with f  F have "f  {fF. sing_reg (lt f) (lt p)}" by simp
  thus ?thesis using * unfolding sig_red_def by blast
qed

lemma sig_red_regular_rtrancl_lt:
  assumes "(sig_red (≺t) top_tail F)** p q"
  shows "lt q = lt p"
  using assms by (induct, auto dest: sig_red_regular_lt)

lemma sig_red_regular_rtrancl_lc:
  assumes "(sig_red (≺t) top_tail F)** p q"
  shows "lc q = lc p"
  using assms by (induct, auto dest: sig_red_regular_lc)

lemma sig_red_rtrancl_lt:
  assumes "(sig_red sing_reg top_tail F)** p q"
  shows "lt q t lt p"
  using assms by (induct, auto dest: sig_red_lt)

lemma sig_red_tail_rtrancl_lt_rep_list:
  assumes "(sig_red sing_reg (≺) F)** p q"
  shows "punit.lt (rep_list q) = punit.lt (rep_list p)"
  using assms by (induct, auto dest: sig_red_tail_lt_rep_list)

lemma sig_red_tail_rtrancl_lc_rep_list:
  assumes "(sig_red sing_reg (≺) F)** p q"
  shows "punit.lc (rep_list q) = punit.lc (rep_list p)"
  using assms by (induct, auto dest: sig_red_tail_lc_rep_list)

lemma sig_red_rtrancl_lt_rep_list:
  assumes "(sig_red sing_reg top_tail F)** p q"
  shows "punit.lt (rep_list q)  punit.lt (rep_list p)"
  using assms by (induct, auto dest: sig_red_lt_rep_list)

lemma sig_red_red_rtrancl:
  assumes "(sig_red sing_reg top_tail F)** p q"
  shows "(punit.red (rep_list ` F))** (rep_list p) (rep_list q)"
  using assms by (induct, auto dest: sig_red_red)

lemma sig_red_rtrancl_monom_mult:
  assumes "(sig_red sing_reg top_tail F)** p q"
  shows "(sig_red sing_reg top_tail F)** (monom_mult c s p) (monom_mult c s q)"
proof (cases "c = 0")
  case True
  thus ?thesis by simp
next
  case False
  from assms(1) show ?thesis
  proof induct
    case base
    show ?case ..
  next
    case (step y z)
    from step(2) False have "sig_red sing_reg top_tail F (monom_mult c s y) (monom_mult c s z)"
      by (rule sig_red_monom_mult)
    with step(3) show ?case ..
  qed
qed

lemma sig_red_rtrancl_sing_regI: "(sig_red sing_reg top_tail F)** p q  (sig_red (≼t) top_tail F)** p q"
  by (induct rule: rtranclp_induct, auto dest: sig_red_sing_regI)

lemma sig_red_rtrancl_top_tailI: "(sig_red sing_reg top_tail F)** p q  (sig_red sing_reg (≼) F)** p q"
  by (induct rule: rtranclp_induct, auto dest: sig_red_top_tailI)

lemma dgrad_sig_set_closed_sig_red_rtrancl:
  assumes "dickson_grading d" and "F  dgrad_sig_set d" and "p  dgrad_sig_set d"
    and "(sig_red sing_red top_tail F)** p q"
  shows "q  dgrad_sig_set d"
  using assms(4, 1, 2, 3) by (induct, auto intro: dgrad_sig_set_closed_sig_red)

lemma sig_red_rtrancl_mono:
  assumes "(sig_red sing_reg top_tail F)** p q" and "F  F'"
  shows "(sig_red sing_reg top_tail F')** p q"
  using assms(1) by (induct rule: rtranclp_induct, auto dest: sig_red_mono[OF _ assms(2)])

lemma sig_red_rtrancl_subset:
  assumes "(sig_red sing_reg top_tail F)** p q" and "sing_reg = (≼t)  sing_reg = (≺t)"
  shows "(sig_red sing_reg top_tail {fF. sing_reg (lt f) (lt p)})** p q"
  using assms(1)
proof (induct rule: rtranclp_induct)
  case base
  show ?case by (fact rtranclp.rtrancl_refl)
next
  case (step y z)
  from step(2) assms(2) have "sig_red sing_reg top_tail {f  F. sing_reg (lt f) (lt y)} y z"
    by (rule sig_red_subset)
  moreover have "{f  F. sing_reg (lt f) (lt y)}  {f  F. sing_reg (lt f) (lt p)}"
  proof
    fix f
    assume "f  {f  F. sing_reg (lt f) (lt y)}"
    hence "f  F" and 1: "sing_reg (lt f) (lt y)" by simp_all
    from step(1) have 2: "lt y t lt p" by (rule sig_red_rtrancl_lt)
    from assms(2) have "sing_reg (lt f) (lt p)"
    proof
      assume "sing_reg = (≼t)"
      with 1 2 show ?thesis by simp
    next
      assume "sing_reg = (≺t)"
      with 1 2 show ?thesis by simp
    qed
    with f  F show "f  {f  F. sing_reg (lt f) (lt p)}" by simp
  qed
  ultimately have "sig_red sing_reg top_tail {f  F. sing_reg (lt f) (lt p)} y z"
    by (rule sig_red_mono)
  with step(3) show ?case ..
qed

lemma is_sig_red_is_red: "is_sig_red sing_reg top_tail F p  punit.is_red (rep_list ` F) (rep_list p)"
  by (auto simp: is_sig_red_def punit.is_red_alt dest: sig_red_red)

lemma is_sig_red_monom_mult:
  assumes "is_sig_red sing_reg top_tail F p" and "c  0"
  shows "is_sig_red sing_reg top_tail F (monom_mult c s p)"
proof -
  from assms(1) obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
  hence "sig_red sing_reg top_tail F (monom_mult c s p) (monom_mult c s q)"
    using assms(2) by (rule sig_red_monom_mult)
  thus ?thesis unfolding is_sig_red_def ..
qed

lemma is_sig_red_sing_reg_cases:
  "is_sig_red (≼t) top_tail F p = (is_sig_red (=) top_tail F p  is_sig_red (≺t) top_tail F p)"
  by (auto simp: is_sig_red_def sig_red_sing_reg_cases)

corollary is_sig_red_sing_regI: "is_sig_red sing_reg top_tail F p  is_sig_red (≼t) top_tail F p"
  by (auto simp: is_sig_red_def intro: sig_red_sing_regI)

lemma is_sig_red_top_tail_cases:
  "is_sig_red sing_reg (≼) F p = (is_sig_red sing_reg (=) F p  is_sig_red sing_reg (≺) F p)"
  by (auto simp: is_sig_red_def sig_red_top_tail_cases)

corollary is_sig_red_top_tailI: "is_sig_red sing_reg top_tail F p  is_sig_red sing_reg (≼) F p"
  by (auto simp: is_sig_red_def intro: sig_red_top_tailI)

lemma is_sig_red_singletonI:
  assumes "is_sig_red sing_reg top_tail F r"
  obtains f where "f  F" and "is_sig_red sing_reg top_tail {f} r"
proof -
  from assms obtain r' where "sig_red sing_reg top_tail F r r'" unfolding is_sig_red_def ..
  then obtain f t where "f  F" and t: "sig_red_single sing_reg top_tail r r' f t"
    by (auto simp: sig_red_def)
  have "is_sig_red sing_reg top_tail {f} r" unfolding is_sig_red_def sig_red_def
  proof (intro exI bexI)
    show "f  {f}" by simp
  qed fact
  with f  F show ?thesis ..
qed

lemma is_sig_red_singletonD:
  assumes "is_sig_red sing_reg top_tail {f} r" and "f  F"
  shows "is_sig_red sing_reg top_tail F r"
proof -
  from assms(1) obtain r' where "sig_red sing_reg top_tail {f} r r'" unfolding is_sig_red_def ..
  then obtain t where "sig_red_single sing_reg top_tail r r' f t" by (auto simp: sig_red_def)
  show ?thesis unfolding is_sig_red_def sig_red_def by (intro exI bexI, fact+)
qed

lemma is_sig_redD1:
  assumes "is_sig_red sing_reg top_tail F p"
  shows "ord_term_lin.is_le_rel sing_reg"
proof -
  from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
  then obtain f s where "f  F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
  from this(2) show ?thesis by (rule sig_red_singleD)
qed

lemma is_sig_redD2:
  assumes "is_sig_red sing_reg top_tail F p"
  shows "ordered_powerprod_lin.is_le_rel top_tail"
proof -
  from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
  then obtain f s where "f  F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
  from this(2) show ?thesis by (rule sig_red_singleD)
qed

lemma is_sig_red_addsI:
  assumes "f  F" and "t  keys (rep_list p)" and "rep_list f  0" and "punit.lt (rep_list f) adds t"
    and "ord_term_lin.is_le_rel sing_reg" and "ordered_powerprod_lin.is_le_rel top_tail"
    and "sing_reg (t  lt f) (punit.lt (rep_list f)  lt p)" and "top_tail t (punit.lt (rep_list p))"
  shows "is_sig_red sing_reg top_tail F p"
  unfolding is_sig_red_def
proof
  let ?q = "p - monom_mult ((lookup (rep_list p) t) / punit.lc (rep_list f)) (t - punit.lt (rep_list f)) f"
  show "sig_red sing_reg top_tail F p ?q" unfolding sig_red_def
  proof (intro bexI exI)
    from assms(4) have eq: "(t - punit.lt (rep_list f)) + punit.lt (rep_list f) = t" by (rule adds_minus)
    from assms(4, 5, 7) have "sing_reg ((t - punit.lt (rep_list f))  lt f) (lt p)"
      by (simp only: term_is_le_rel_minus)
    thus "sig_red_single sing_reg top_tail p ?q f (t - punit.lt (rep_list f))"
      by (simp add: assms eq sig_red_singleI)
  qed fact
qed

lemma is_sig_red_addsE:
  assumes "is_sig_red sing_reg top_tail F p"
  obtains f t where "f  F" and "t  keys (rep_list p)" and "rep_list f  0"
    and "punit.lt (rep_list f) adds t"
    and "sing_reg (t  lt f) (punit.lt (rep_list f)  lt p)"
    and "top_tail t (punit.lt (rep_list p))"
proof -
  from assms have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
  from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
  then obtain f s where "f  F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
  from this(2) have 1: "rep_list f  0" and 2: "s + punit.lt (rep_list f)  keys (rep_list p)"
    and 3: "sing_reg (s  lt f) (lt p)" and 4: "top_tail (s + punit.lt (rep_list f)) (punit.lt (rep_list p))"
    by (rule sig_red_singleD)+
  note f  F 2 1
  moreover have "punit.lt (rep_list f) adds s + punit.lt (rep_list f)" by simp
  moreover from 3 have "sing_reg ((s + punit.lt (rep_list f))  lt f) (punit.lt (rep_list f)  lt p)"
    by (simp add: add.commute[of s] splus_assoc term_is_le_rel_canc_left[OF *])
  moreover from 4 have "top_tail (s + punit.lt (rep_list f)) (punit.lt (rep_list p))" by simp
  ultimately show ?thesis ..
qed

lemma is_sig_red_top_addsI:
  assumes "f  F" and "rep_list f  0" and "rep_list p  0"
    and "punit.lt (rep_list f) adds punit.lt (rep_list p)" and "ord_term_lin.is_le_rel sing_reg"
    and "sing_reg (punit.lt (rep_list p)  lt f) (punit.lt (rep_list f)  lt p)"
  shows "is_sig_red sing_reg (=) F p"
proof -
  note assms(1)
  moreover from assms(3) have "punit.lt (rep_list p)  keys (rep_list p)" by (rule punit.lt_in_keys)
  moreover note assms(2, 4, 5) ordered_powerprod_lin.is_le_relI(1) assms(6) refl
  ultimately show ?thesis by (rule is_sig_red_addsI)
qed

lemma is_sig_red_top_addsE:
  assumes "is_sig_red sing_reg (=) F p"
  obtains f where "f  F" and "rep_list f  0" and "rep_list p  0"
    and "punit.lt (rep_list f) adds punit.lt (rep_list p)"
    and "sing_reg (punit.lt (rep_list p)  lt f) (punit.lt (rep_list f)  lt p)"
proof -
  from assms obtain f t where 1: "f  F" and 2: "t  keys (rep_list p)" and 3: "rep_list f  0"
    and 4: "punit.lt (rep_list f) adds t"
    and 5: "sing_reg (t  lt f) (punit.lt (rep_list f)  lt p)"
    and t: "t = punit.lt (rep_list p)" by (rule is_sig_red_addsE)
  note 1 3
  moreover from 2 have "rep_list p  0" by auto
  moreover from 4 have "punit.lt (rep_list f) adds punit.lt (rep_list p)" by (simp only: t)
  moreover from 5 have "sing_reg (punit.lt (rep_list p)  lt f) (punit.lt (rep_list f)  lt p)"
    by (simp only: t)
  ultimately show ?thesis ..
qed

lemma is_sig_red_top_plusE:
  assumes "is_sig_red sing_reg (=) F p" and "is_sig_red sing_reg (=) F q"
    and "lt p t lt (p + q)" and "lt q t lt (p + q)" and "sing_reg = (≼t)  sing_reg = (≺t)"
  assumes 1: "is_sig_red sing_reg (=) F (p + q)  thesis"
  assumes 2: "punit.lt (rep_list p) = punit.lt (rep_list q)  punit.lc (rep_list p) + punit.lc (rep_list q) = 0  thesis"
  shows thesis
proof -
  from assms(1) obtain f1 where "f1  F" and "rep_list f1  0" and "rep_list p  0"
    and a: "punit.lt (rep_list f1) adds punit.lt (rep_list p)"
    and b: "sing_reg (punit.lt (rep_list p)  lt f1) (punit.lt (rep_list f1)  lt p)"
    by (rule is_sig_red_top_addsE)
  from assms(2) obtain f2 where "f2  F" and "rep_list f2  0" and "rep_list q  0"
    and c: "punit.lt (rep_list f2) adds punit.lt (rep_list q)"
    and d: "sing_reg (punit.lt (rep_list q)  lt f2) (punit.lt (rep_list f2)  lt q)"
    by (rule is_sig_red_top_addsE)
  show ?thesis
  proof (cases "punit.lt (rep_list p) = punit.lt (rep_list q)  punit.lc (rep_list p) + punit.lc (rep_list q) = 0")
    case True
    hence "punit.lt (rep_list p) = punit.lt (rep_list q)" and "punit.lc (rep_list p) + punit.lc (rep_list q) = 0"
      by simp_all
    thus ?thesis by (rule 2)
  next
    case False
    hence disj: "punit.lt (rep_list p)  punit.lt (rep_list q)  punit.lc (rep_list p) + punit.lc (rep_list q)  0"
      by simp
    from assms(5) have "ord_term_lin.is_le_rel sing_reg" by (simp add: ord_term_lin.is_le_rel_def)
    have "rep_list (p + q)  0" unfolding rep_list_plus
    proof
      assume eq: "rep_list p + rep_list q = 0"
      have eq2: "punit.lt (rep_list p) = punit.lt (rep_list q)"
      proof (rule ordered_powerprod_lin.linorder_cases)
        assume *: "punit.lt (rep_list p)  punit.lt (rep_list q)"
        hence "punit.lt (rep_list p + rep_list q) = punit.lt (rep_list q)" by (rule punit.lt_plus_eqI)
        with * zero_min[of "punit.lt (rep_list p)"] show ?thesis by (simp add: eq)
      next
        assume *: "punit.lt (rep_list q)  punit.lt (rep_list p)"
        hence "punit.lt (rep_list p + rep_list q) = punit.lt (rep_list p)" by (rule punit.lt_plus_eqI_2)
        with * zero_min[of "punit.lt (rep_list q)"] show ?thesis by (simp add: eq)
      qed
      with disj have "punit.lc (rep_list p) + punit.lc (rep_list q)  0" by simp
      thus False by (simp add: punit.lc_def eq2 lookup_add[symmetric] eq)
    qed
    have "punit.lt (rep_list (p + q)) = ordered_powerprod_lin.max (punit.lt (rep_list p)) (punit.lt (rep_list q))"
      unfolding rep_list_plus
    proof (rule punit.lt_plus_eq_maxI)
      assume "punit.lt (rep_list p) = punit.lt (rep_list q)"
      with disj show "punit.lc (rep_list p) + punit.lc (rep_list q)  0" by simp
    qed
    hence "punit.lt (rep_list (p + q)) = punit.lt (rep_list p)  punit.lt (rep_list (p + q)) = punit.lt (rep_list q)"
      by (simp add: ordered_powerprod_lin.max_def)
    thus ?thesis
    proof
      assume eq: "punit.lt (rep_list (p + q)) = punit.lt (rep_list p)"
      show ?thesis
      proof (rule 1, rule is_sig_red_top_addsI)
        from a show "punit.lt (rep_list f1) adds punit.lt (rep_list (p + q))" by (simp only: eq)
      next
        from b have "sing_reg (punit.lt (rep_list (p + q))  lt f1) (punit.lt (rep_list f1)  lt p)"
          by (simp only: eq)
        moreover from assms(3) have "... t punit.lt (rep_list f1)  lt (p + q)" by (rule splus_mono)
        ultimately show "sing_reg (punit.lt (rep_list (p + q))  lt f1) (punit.lt (rep_list f1)  lt (p + q))"
          using assms(5) by auto
      qed fact+
    next
      assume eq: "punit.lt (rep_list (p + q)) = punit.lt (rep_list q)"
      show ?thesis
      proof (rule 1, rule is_sig_red_top_addsI)
        from c show "punit.lt (rep_list f2) adds punit.lt (rep_list (p + q))" by (simp only: eq)
      next
        from d have "sing_reg (punit.lt (rep_list (p + q))  lt f2) (punit.lt (rep_list f2)  lt q)"
          by (simp only: eq)
        moreover from assms(4) have "... t punit.lt (rep_list f2)  lt (p + q)" by (rule splus_mono)
        ultimately show "sing_reg (punit.lt (rep_list (p + q))  lt f2) (punit.lt (rep_list f2)  lt (p + q))"
          using assms(5) by auto
      qed fact+
    qed
  qed
qed

lemma is_sig_red_singleton_monom_multD:
  assumes "is_sig_red sing_reg top_tail {monom_mult c t f} p"
  shows "is_sig_red sing_reg top_tail {f} p"
proof -
  let ?f = "monom_mult c t f"
  from assms obtain s where "s  keys (rep_list p)" and 2: "rep_list ?f  0"
    and 3: "punit.lt (rep_list ?f) adds s"
    and 4: "sing_reg (s  lt ?f) (punit.lt (rep_list ?f)  lt p)"
    and "top_tail s (punit.lt (rep_list p))"
    by (auto elim: is_sig_red_addsE)
  from 2 have "c  0" and "rep_list f  0"
    by (simp_all add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
  hence "f  0" by (auto simp: rep_list_zero)
  with c  0 have eq1: "lt ?f = t  lt f" by (simp add: lt_monom_mult)
  from c  0 rep_list f  0 have eq2: "punit.lt (rep_list ?f) = t + punit.lt (rep_list f)"
    by (simp add: rep_list_monom_mult punit.lt_monom_mult)
  from assms have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
  show ?thesis
  proof (rule is_sig_red_addsI)
    show "f  {f}" by simp
  next
    have "punit.lt (rep_list f) adds t + punit.lt (rep_list f)" by (rule adds_triv_right)
    also from 3 have "... adds s" by (simp only: eq2)
    finally show "punit.lt (rep_list f) adds s" .
  next
    from 4 have "sing_reg (t  (s  lt f)) (t  (punit.lt (rep_list f)  lt p))"
      by (simp add: eq1 eq2 splus_assoc splus_left_commute)
    with * show "sing_reg (s  lt f) (punit.lt (rep_list f)  lt p)"
      by (simp add: term_is_le_rel_canc_left)
  next
    from assms show "ordered_powerprod_lin.is_le_rel top_tail" by (rule is_sig_redD2)
  qed fact+
qed

lemma is_sig_red_top_singleton_monom_multI:
  assumes "is_sig_red sing_reg (=) {f} p" and "c  0"
    and "t adds punit.lt (rep_list p) - punit.lt (rep_list f)"
  shows "is_sig_red sing_reg (=) {monom_mult c t f} p"
proof -
  let ?f = "monom_mult c t f"
  from assms have 2: "rep_list f  0" and "rep_list p  0"
    and 3: "punit.lt (rep_list f) adds punit.lt (rep_list p)"
    and 4: "sing_reg (punit.lt (rep_list p)  lt f) (punit.lt (rep_list f)  lt p)"
    by (auto elim: is_sig_red_top_addsE)
  hence "f  0" by (auto simp: rep_list_zero)
  with c  0 have eq1: "lt ?f = t  lt f" by (simp add: lt_monom_mult)
  from c  0 rep_list f  0 have eq2: "punit.lt (rep_list ?f) = t + punit.lt (rep_list f)"
    by (simp add: rep_list_monom_mult punit.lt_monom_mult)
  from assms(1) have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
  show ?thesis
  proof (rule is_sig_red_top_addsI)
    show "?f  {?f}" by simp
  next
    from c  0 rep_list f  0 show "rep_list ?f  0"
      by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
  next
    from assms(3) have "t + punit.lt (rep_list f) adds
                        (punit.lt (rep_list p) - punit.lt (rep_list f)) + punit.lt (rep_list f)"
      by (simp only: adds_canc)
    also from 3 have "... = punit.lt (rep_list p)" by (rule adds_minus)
    finally show "punit.lt (rep_list ?f) adds punit.lt (rep_list p)" by (simp only: eq2)
  next
    from 4 * show "sing_reg (punit.lt (rep_list p)  lt ?f) (punit.lt (rep_list ?f)  lt p)"
      by (simp add: eq1 eq2 term_is_le_rel_canc_left splus_assoc splus_left_commute)
  qed fact+
qed

lemma is_sig_red_cong':
  assumes "is_sig_red sing_reg top_tail F p" and "lt p = lt q" and "rep_list p = rep_list q"
  shows "is_sig_red sing_reg top_tail F q"
proof -
  from assms(1) have 1: "ord_term_lin.is_le_rel sing_reg" and 2: "ordered_powerprod_lin.is_le_rel top_tail"
    by (rule is_sig_redD1, rule is_sig_redD2)
  from assms(1) obtain f t where "f  F" and "t  keys (rep_list p)" and "rep_list f  0"
    and "punit.lt (rep_list f) adds t"
    and "sing_reg (t  lt f) (punit.lt (rep_list f)  lt p)"
    and "top_tail t (punit.lt (rep_list p))" by (rule is_sig_red_addsE)
  from this(1-4) 1 2 this(5, 6) show ?thesis unfolding assms(2, 3) by (rule is_sig_red_addsI)
qed

lemma is_sig_red_cong:
  "lt p = lt q  rep_list p = rep_list q 
      is_sig_red sing_reg top_tail F p  is_sig_red sing_reg top_tail F q"
  by (auto intro: is_sig_red_cong')

lemma is_sig_red_top_cong:
  assumes "is_sig_red sing_reg (=) F p" and "rep_list q  0" and "lt p = lt q"
    and "punit.lt (rep_list p) = punit.lt (rep_list q)"
  shows "is_sig_red sing_reg (=) F q"
proof -
  from assms(1) have 1: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
  from assms(1) obtain f where "f  F" and "rep_list f  0" and "rep_list p  0"
    and "punit.lt (rep_list f) adds punit.lt (rep_list p)"
    and "sing_reg (punit.lt (rep_list p)  lt f) (punit.lt (rep_list f)  lt p)"
    by (rule is_sig_red_top_addsE)
  from this(1, 2) assms(2) this(4) 1 this(5) show ?thesis
    unfolding assms(3, 4) by (rule is_sig_red_top_addsI)
qed

lemma sig_irredE_dgrad_max_set:
  assumes "dickson_grading d" and "F  dgrad_max_set d"
  obtains q where "(sig_red sing_reg top_tail F)** p q" and "¬ is_sig_red sing_reg top_tail F q"
proof -
  let ?Q = "{q. (sig_red sing_reg top_tail F)** p q}"
  from assms have "wfP (sig_red sing_reg top_tail F)¯¯" by (rule sig_red_wf_dgrad_max_set)
  moreover have "p  ?Q" by simp
  ultimately obtain q where "q  ?Q" and "x. (sig_red sing_reg top_tail F)¯¯ x q  x  ?Q"
    by (rule wfE_min[to_pred], blast)
  hence 1: "(sig_red sing_reg top_tail F)** p q"
    and 2: "x. sig_red sing_reg top_tail F q x  ¬ (sig_red sing_reg top_tail F)** p x"
    by simp_all
  show ?thesis
  proof
    show "¬ is_sig_red sing_reg top_tail F q"
    proof
      assume "is_sig_red sing_reg top_tail F q"
      then obtain x where 3: "sig_red sing_reg top_tail F q x" unfolding is_sig_red_def ..
      hence "¬ (sig_red sing_reg top_tail F)** p x" by (rule 2)
      moreover from 1 3 have "(sig_red sing_reg top_tail F)** p x" ..
      ultimately show False ..
    qed
  qed fact
qed

lemma is_sig_red_mono:
  "is_sig_red sing_reg top_tail F p  F  F'  is_sig_red sing_reg top_tail F' p"
  by (auto simp: is_sig_red_def dest: sig_red_mono)

lemma is_sig_red_Un:
  "is_sig_red sing_reg top_tail (A  B) p  (is_sig_red sing_reg top_tail A p  is_sig_red sing_reg top_tail B p)"
  by (auto simp: is_sig_red_def sig_red_Un)

lemma is_sig_redD_lt:
  assumes "is_sig_red (≼t) top_tail {f} p"
  shows "lt f t lt p"
proof -
  from assms obtain s where "rep_list f  0" and "s  keys (rep_list p)"
    and 1: "punit.lt (rep_list f) adds s" and 2: "s  lt f t punit.lt (rep_list f)  lt p"
    by (auto elim!: is_sig_red_addsE)
  from 1 obtain t where eq: "s = punit.lt (rep_list f) + t" by (rule addsE)
  hence "punit.lt (rep_list f)  (t  lt f) = s  lt f" by (simp add: splus_assoc)
  also note 2
  finally have "t  lt f t lt p" by (rule ord_term_canc)
  have "0  t" by (fact zero_min)
  hence "0  lt f t t  lt f" by (rule splus_mono_left)
  hence "lt f t t  lt f" by (simp add: term_simps)
  thus ?thesis using t  lt f t lt p by simp
qed

lemma is_sig_red_regularD_lt:
  assumes "is_sig_red (≺t) top_tail {f} p"
  shows "lt f t lt p"
proof -
  from assms obtain s where "rep_list f  0" and "s  keys (rep_list p)"
    and 1: "punit.lt (rep_list f) adds s" and 2: "s  lt f t punit.lt (rep_list f)  lt p"
    by (auto elim!: is_sig_red_addsE)
  from 1 obtain t where eq: "s = punit.lt (rep_list f) + t" by (rule addsE)
  hence "punit.lt (rep_list f)  (t  lt f) = s  lt f" by (simp add: splus_assoc)
  also note 2
  finally have "t  lt f t lt p" by (rule ord_term_strict_canc)
  have "0  t" by (fact zero_min)
  hence "0  lt f t t  lt f" by (rule splus_mono_left)
  hence "lt f t t  lt f" by (simp add: term_simps)
  thus ?thesis using t  lt f t lt p by (rule ord_term_lin.le_less_trans)
qed

lemma sig_irred_regular_self: "¬ is_sig_red (≺t) top_tail {p} p"
  by (auto dest: is_sig_red_regularD_lt)

subsubsection ‹Signature Gr\"obner Bases›

definition sig_red_zero :: "('t 't  bool)  ('t 0 'b) set  ('t 0 'b)  bool"
  where "sig_red_zero sing_reg F r  (s. (sig_red sing_reg (≼) F)** r s  rep_list s = 0)"

definition is_sig_GB_in :: "('a  nat)  ('t 0 'b) set  't  bool"
  where "is_sig_GB_in d G u  (r. lt r = u  r  dgrad_sig_set d  sig_red_zero (≼t) G r)"

definition is_sig_GB_upt :: "('a  nat)  ('t 0 'b) set  't  bool"
  where "is_sig_GB_upt d G u 
            (G  dgrad_sig_set d  (v. v t u  d (pp_of_term v)  dgrad_max d 
                                          component_of_term v < length fs  is_sig_GB_in d G v))"

definition is_min_sig_GB :: "('a  nat)  ('t 0 'b) set  bool"
  where "is_min_sig_GB d G  G  dgrad_sig_set d 
                                (u. d (pp_of_term u)  dgrad_max d  component_of_term u < length fs 
                                      is_sig_GB_in d G u) 
                                (gG. ¬ is_sig_red (≼t) (=) (G - {g}) g)"

definition is_syz_sig :: "('a  nat)  't  bool"
  where "is_syz_sig d u  (sdgrad_sig_set d. s  0  lt s = u  rep_list s = 0)"

lemma sig_red_zeroI:
  assumes "(sig_red sing_reg (≼) F)** r s" and "rep_list s = 0"
  shows "sig_red_zero sing_reg F r"
  unfolding sig_red_zero_def using assms by blast

lemma sig_red_zeroE:
  assumes "sig_red_zero sing_reg F r"
  obtains s where "(sig_red sing_reg (≼) F)** r s" and "rep_list s = 0"
  using assms unfolding sig_red_zero_def by blast

lemma sig_red_zero_monom_mult:
  assumes "sig_red_zero sing_reg F r"
  shows "sig_red_zero sing_reg F (monom_mult c t r)"
proof -
  from assms obtain s where "(sig_red sing_reg (≼) F)** r s" and "rep_list s = 0"
    by (rule sig_red_zeroE)
  from this(1) have "(sig_red sing_reg (≼) F)** (monom_mult c t r) (monom_mult c t s)"
    by (rule sig_red_rtrancl_monom_mult)
  moreover have "rep_list (monom_mult c t s) = 0" by (simp add: rep_list_monom_mult rep_list s = 0)
  ultimately show ?thesis by (rule sig_red_zeroI)
qed

lemma sig_red_zero_sing_regI:
  assumes "sig_red_zero sing_reg G p"
  shows "sig_red_zero (≼t) G p"
proof -
  from assms obtain s where "(sig_red sing_reg (≼) G)** p s" and "rep_list s = 0"
    by (rule sig_red_zeroE)
  from this(1) have "(sig_red (≼t) (≼) G)** p s" by (rule sig_red_rtrancl_sing_regI)
  thus ?thesis using rep_list s = 0 by (rule sig_red_zeroI)
qed

lemma sig_red_zero_nonzero:
  assumes "sig_red_zero sing_reg F r" and "rep_list r  0" and "sing_reg = (≼t)  sing_reg = (≺t)"
  shows "is_sig_red sing_reg (=) F r"
proof -
  from assms(1) obtain s where "(sig_red sing_reg (≼) F)** r s" and "rep_list s = 0"
    by (rule sig_red_zeroE)
  from this(1) assms(2) show ?thesis
  proof (induct rule: converse_rtranclp_induct)
    case base
    thus ?case using rep_list s = 0 ..
  next
    case (step y z)
    from step(1) obtain f t where "f  F" and *: "sig_red_single sing_reg (≼) y z f t"
      unfolding sig_red_def by blast
    from this(2) have 1: "rep_list f  0" and 2: "t + punit.lt (rep_list f)  keys (rep_list y)"
      and 3: "z = y - monom_mult (lookup (rep_list y) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
      and 4: "ord_term_lin.is_le_rel sing_reg" and 5: "sing_reg (t  lt f) (lt y)"
      by (rule sig_red_singleD)+
    show ?case
    proof (cases "t + punit.lt (rep_list f) = punit.lt (rep_list y)")
      case True
      show ?thesis unfolding is_sig_red_def
      proof
        show "sig_red sing_reg (=) F y z" unfolding sig_red_def
        proof (intro bexI exI)
          from 1 2 3 4 ordered_powerprod_lin.is_le_relI(1) 5 True
          show "sig_red_single sing_reg (=) y z f t" by (rule sig_red_singleI)
        qed fact
      qed
    next
      case False
      from 2 have "t + punit.lt (rep_list f)  punit.lt (rep_list y)" by (rule punit.lt_max_keys)
      with False have "t + punit.lt (rep_list f)  punit.lt (rep_list y)" by simp
      with 1 2 3 4 ordered_powerprod_lin.is_le_relI(3) 5 have "sig_red_single sing_reg (≺) y z f t"
        by (rule sig_red_singleI)
      hence "punit.lt (rep_list y)  keys (rep_list z)"
        and lt_z: "punit.lt (rep_list z) = punit.lt (rep_list y)"
        by (rule sig_red_single_tail_lt_in_keys_rep_list, rule sig_red_single_tail_lt_rep_list)
      from this(1) have "rep_list z  0" by auto
      hence "is_sig_red sing_reg (=) F z" by (rule step(3))
      then obtain g where "g  F" and "rep_list g  0"
        and "punit.lt (rep_list g) adds punit.lt (rep_list z)"
        and a: "sing_reg (punit.lt (rep_list z)  lt g) (punit.lt (rep_list g)  lt z)"
        by (rule is_sig_red_top_addsE)
      from this(3) have "punit.lt (rep_list g) adds punit.lt (rep_list y)" by (simp only: lt_z)
      with g  F rep_list g  0 step(4) show ?thesis
      proof (rule is_sig_red_top_addsI)
        from is_sig_red sing_reg (=) F z show "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
      next
        from sig_red_single sing_reg (≺) y z f t have "lt z t lt y" by (rule sig_red_single_lt)
        from assms(3) show "sing_reg (punit.lt (rep_list y)  lt g) (punit.lt (rep_list g)  lt y)"
        proof
          assume "sing_reg = (≼t)"
          from a have "punit.lt (rep_list y)  lt g t punit.lt (rep_list g)  lt z"
            by (simp only: lt_z sing_reg = (≼t))
          also from lt z t lt y have "... t punit.lt (rep_list g)  lt y" by (rule splus_mono)
          finally show ?thesis by (simp only: sing_reg = (≼t))
        next
          assume "sing_reg = (≺t)"
          from a have "punit.lt (rep_list y)  lt g t punit.lt (rep_list g)  lt z"
            by (simp only: lt_z sing_reg = (≺t))
          also from lt z t lt y have "... t punit.lt (rep_list g)  lt y" by (rule splus_mono)
          finally show ?thesis by (simp only: sing_reg = (≺t))
        qed
      qed
    qed
  qed
qed

lemma sig_red_zero_mono: "sig_red_zero sing_reg F p  F  F'  sig_red_zero sing_reg F' p"
  by (auto simp: sig_red_zero_def dest: sig_red_rtrancl_mono)

lemma sig_red_zero_subset:
  assumes "sig_red_zero sing_reg F p" and "sing_reg = (≼t)  sing_reg = (≺t)"
  shows "sig_red_zero sing_reg {fF. sing_reg (lt f) (lt p)} p"
proof -
  from assms(1) obtain s where "(sig_red sing_reg (≼) F)** p s" and "rep_list s = 0"
    by (rule sig_red_zeroE)
  from this(1) assms(2) have "(sig_red sing_reg (≼) {fF. sing_reg (lt f) (lt p)})** p s"
    by (rule sig_red_rtrancl_subset)
  thus ?thesis using rep_list s = 0 by (rule sig_red_zeroI)
qed

lemma sig_red_zero_idealI:
  assumes "sig_red_zero sing_reg F p"
  shows "rep_list p  ideal (rep_list ` F)"
proof -
  from assms obtain s where "(sig_red sing_reg (≼) F)** p s" and "rep_list s = 0" by (rule sig_red_zeroE)
  from this(1) have "(punit.red (rep_list ` F))** (rep_list p) (rep_list s)" by (rule sig_red_red_rtrancl)
  hence "(punit.red (rep_list ` F))** (rep_list p) 0" by (simp only: rep_list s = 0)
  thus ?thesis by (rule punit.red_rtranclp_0_in_pmdl[simplified])
qed

lemma is_sig_GB_inI:
  assumes "r. lt r = u  r  dgrad_sig_set d  sig_red_zero (≼t) G r"
  shows "is_sig_GB_in d G u"
  unfolding is_sig_GB_in_def using assms by blast

lemma is_sig_GB_inD:
  assumes "is_sig_GB_in d G u" and "r  dgrad_sig_set d" and "lt r = u"
  shows "sig_red_zero (≼t) G r"
  using assms unfolding is_sig_GB_in_def by blast

lemma is_sig_GB_inI_triv:
  assumes "¬ d (pp_of_term u)  dgrad_max d  ¬ component_of_term u < length fs"
  shows "is_sig_GB_in d G u"
proof (rule is_sig_GB_inI)
  fix r::"'t 0 'b"
  assume "lt r = u" and "r  dgrad_sig_set d"
  show "sig_red_zero (≼t) G r"
  proof (cases "r = 0")
    case True
    hence "rep_list r = 0" by (simp only: rep_list_zero)
    with rtrancl_refl[to_pred] show ?thesis by (rule sig_red_zeroI)
  next
    case False
    from r  dgrad_sig_set d have "d (lp r)  dgrad_max d" by (rule dgrad_sig_setD_lp)
    moreover from r  dgrad_sig_set d False have "component_of_term (lt r) < length fs"
      by (rule dgrad_sig_setD_lt)
    ultimately show ?thesis using assms by (simp add: lt r = u)
  qed
qed

lemma is_sig_GB_in_mono: "is_sig_GB_in d G u  G  G'  is_sig_GB_in d G' u"
  by (auto simp: is_sig_GB_in_def dest: sig_red_zero_mono)

lemma is_sig_GB_uptI:
  assumes "G  dgrad_sig_set d"
    and "v. v t u  d (pp_of_term v)  dgrad_max d  component_of_term v < length fs 
          is_sig_GB_in d G v"
  shows "is_sig_GB_upt d G u"
  unfolding is_sig_GB_upt_def using assms by blast

lemma is_sig_GB_uptD1:
  assumes "is_sig_GB_upt d G u"
  shows "G  dgrad_sig_set d"
  using assms unfolding is_sig_GB_upt_def by blast

lemma is_sig_GB_uptD2:
  assumes "is_sig_GB_upt d G u" and "v t u"
  shows "is_sig_GB_in d G v"
  using assms is_sig_GB_inI_triv unfolding is_sig_GB_upt_def by blast

lemma is_sig_GB_uptD3:
  assumes "is_sig_GB_upt d G u" and "r  dgrad_sig_set d" and "lt r t u"
  shows "sig_red_zero (≼t) G r"
  by (rule is_sig_GB_inD, rule is_sig_GB_uptD2, fact+, fact refl)

lemma is_sig_GB_upt_le:
  assumes "is_sig_GB_upt d G u" and "v t u"
  shows "is_sig_GB_upt d G v"
proof (rule is_sig_GB_uptI)
  from assms(1) show "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
next
  fix w
  assume "w t v"
  hence "w t u" using assms(2) by (rule ord_term_lin.less_le_trans)
  with assms(1) show "is_sig_GB_in d G w" by (rule is_sig_GB_uptD2)
qed

lemma is_sig_GB_upt_mono:
  "is_sig_GB_upt d G u  G  G'  G'  dgrad_sig_set d  is_sig_GB_upt d G' u"
  by (auto simp: is_sig_GB_upt_def dest!: is_sig_GB_in_mono)

lemma is_sig_GB_upt_is_Groebner_basis:
  assumes "dickson_grading d" and "hom_grading d" and "G  dgrad_sig_set' j d"
    and "u. component_of_term u < j  is_sig_GB_in d G u"
  shows "punit.is_Groebner_basis (rep_list ` G)"
  using assms(1)
proof (rule punit.weak_GB_is_strong_GB_dgrad_p_set[simplified])
  from assms(3) have "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  with assms(1) show "rep_list ` G  punit_dgrad_max_set d" by (rule dgrad_max_3)
next
  fix f::"'a 0 'b"
  assume "f  punit_dgrad_max_set d"
  from assms(3) have G_sub: "G  sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
  assume "f  ideal (rep_list ` G)"
  also from rep_list_subset_ideal_sig_inv_set[OF G_sub] have "...  ideal (set (take j fs))"
    by (rule ideal.span_subset_spanI)
  finally have "f  ideal (set (take j fs))" .
  with assms(2) f  punit_dgrad_max_set d obtain r where "r  dgrad_sig_set d"
    and "r  dgrad_sig_set' j d" and f: "f = rep_list r"
    by (rule in_idealE_rep_list_dgrad_sig_set_take)
  from this(2) have "r  sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
  show "(punit.red (rep_list ` G))** f 0"
  proof (cases "r = 0")
    case True
    thus ?thesis by (simp add: f rep_list_zero)
  next
    case False
    hence "lt r  keys r" by (rule lt_in_keys)
    with r  sig_inv_set' j have "component_of_term (lt r) < j" by (rule sig_inv_setD')
    hence "is_sig_GB_in d G (lt r)" by (rule assms(4))
    hence "sig_red_zero (≼t) G r" using r  dgrad_sig_set d refl by (rule is_sig_GB_inD)
    then obtain s where "(sig_red (≼t) (≼) G)** r s" and s: "rep_list s = 0" by (rule sig_red_zeroE)
    from this(1) have "(punit.red (rep_list ` G))** (rep_list r) (rep_list s)"
      by (rule sig_red_red_rtrancl)
    thus ?thesis by (simp only: f s)
  qed
qed

lemma is_sig_GB_is_Groebner_basis:
  assumes "dickson_grading d" and "hom_grading d" and "G  dgrad_max_set d" and "u. is_sig_GB_in d G u"
  shows "punit.is_Groebner_basis (rep_list ` G)"
  using assms(1)
proof (rule punit.weak_GB_is_strong_GB_dgrad_p_set[simplified])
  from assms(1, 3) show "rep_list ` G  punit_dgrad_max_set d" by (rule dgrad_max_3)
next
  fix f::"'a 0 'b"
  assume "f  punit_dgrad_max_set d"
  assume "f  ideal (rep_list ` G)"
  also from rep_list_subset_ideal have "...  ideal (set fs)" by (rule ideal.span_subset_spanI)
  finally have "f  ideal (set fs)" .
  with assms(2) f  punit_dgrad_max_set d obtain r where "r  dgrad_sig_set d" and f: "f = rep_list r"
    by (rule in_idealE_rep_list_dgrad_sig_set)
  from assms(4) this(1) refl have "sig_red_zero (≼t) G r" by (rule is_sig_GB_inD)
  then obtain s where "(sig_red (≼t) (≼) G)** r s" and s: "rep_list s = 0" by (rule sig_red_zeroE)
  from this(1) have "(punit.red (rep_list ` G))** (rep_list r) (rep_list s)"
    by (rule sig_red_red_rtrancl)
  thus "(punit.red (rep_list ` G))** f 0" by (simp only: f s)
qed

lemma sig_red_zero_is_red:
  assumes "sig_red_zero sing_reg F r" and "rep_list r  0"
  shows "is_sig_red sing_reg (≼) F r"
proof -
  from assms(1) obtain s where *: "(sig_red sing_reg (≼) F)** r s" and "rep_list s = 0"
    by (rule sig_red_zeroE)
  from this(2) assms(2) have "r  s" by auto
  with * show ?thesis by (induct rule: converse_rtranclp_induct, auto simp: is_sig_red_def)
qed

lemma is_sig_red_sing_top_is_red_zero:
  assumes "dickson_grading d" and "is_sig_GB_upt d G u" and "a  dgrad_sig_set d" and "lt a = u"
    and "is_sig_red (=) (=) G a" and "¬ is_sig_red (≺t) (=) G a"
  shows "sig_red_zero (≼t) G a"
proof -
  from assms(5) obtain g where "g  G" and "rep_list g  0" and "rep_list a  0"
    and 1: "punit.lt (rep_list g) adds punit.lt (rep_list a)"
    and 2: "punit.lt (rep_list a)  lt g = punit.lt (rep_list g)  lt a"
    by (rule is_sig_red_top_addsE)
  from this(2, 3) have "g  0" and "a  0" by (auto simp: rep_list_zero)
  hence "lc g  0" and "lc a  0" using lc_not_0 by blast+
  from 1 have 3: "(punit.lt (rep_list a) - punit.lt (rep_list g))  lt g = lt a"
    by (simp add: term_is_le_rel_minus 2)
  define g' where "g' = monom_mult (lc a / lc g) (punit.lt (rep_list a) - punit.lt (rep_list g)) g"
  from g  0 lc a  0 lc g  0 have lt_g': "lt g' = lt a" by (simp add: g'_def lt_monom_mult 3)
  from lc g  0 have lc_g': "lc g' = lc a" by (simp add: g'_def)
  from assms(1) have "g'  dgrad_sig_set d" unfolding g'_def
  proof (rule dgrad_sig_set_closed_monom_mult)
    from assms(1) 1 have "d (punit.lt (rep_list a) - punit.lt (rep_list g))  d (punit.lt (rep_list a))"
      by (rule dickson_grading_minus)
    also from assms(1, 3) have "...  dgrad_max d" by (rule dgrad_sig_setD_rep_list_lt)
    finally show "d (punit.lt (rep_list a) - punit.lt (rep_list g))  dgrad_max d" .
  next
    from assms(2) have "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
    with g  G show "g  dgrad_sig_set d" ..
  qed
  with assms(3) have b_in: "a - g'  dgrad_sig_set d" (is "?b  _")
    by (rule dgrad_sig_set_closed_minus)
  from 1 have 4: "punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g) =
                  punit.lt (rep_list a)"
    by (rule adds_minus)

  show ?thesis
  proof (cases "lc a / lc g = punit.lc (rep_list a) / punit.lc (rep_list g)")
    case True
    have "sig_red_single (=) (=) a ?b g (punit.lt (rep_list a) - punit.lt (rep_list g))"
    proof (rule sig_red_singleI)
      show "punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g)  keys (rep_list a)"
        unfolding 4 using rep_list a  0 by (rule punit.lt_in_keys)
    next
      show "?b =
            a - monom_mult
             (lookup (rep_list a) (punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g)) /
              punit.lc (rep_list g))
             (punit.lt (rep_list a) - punit.lt (rep_list g)) g"
        by (simp add: g'_def 4 punit.lc_def True)
    qed (simp_all add: 3 4 rep_list g  0)
    hence "sig_red (=) (=) G a ?b" unfolding sig_red_def using g  G by blast
    hence "sig_red (≼t) (≼) G a ?b" by (auto dest: sig_red_sing_regI sig_red_top_tailI)
    hence 5: "(sig_red (≼t) (≼) G)** a ?b" ..
    show ?thesis
    proof (cases "?b = 0")
      case True
      hence "rep_list ?b = 0" by (simp only: rep_list_zero)
      with 5 show ?thesis by (rule sig_red_zeroI)
    next
      case False
      hence "lt ?b t lt a" using lt_g' lc_g' by (rule lt_minus_lessI)
      hence "lt ?b t u" by (simp only: assms(4))
      with assms(2) b_in have "sig_red_zero (≼t) G ?b" by (rule is_sig_GB_uptD3)
      then obtain s where "(sig_red (≼t) (≼) G)** ?b s" and "rep_list s = 0" by (rule sig_red_zeroE)
      from 5 this(1) have "(sig_red (≼t) (≼) G)** a s" by (rule rtranclp_trans)
      thus ?thesis using rep_list s = 0 by (rule sig_red_zeroI)
    qed
  next
    case False
    from rep_list g  0 lc g  0 lc a  0 have 5: "punit.lt (rep_list g') = punit.lt (rep_list a)"
      by (simp add: g'_def rep_list_monom_mult punit.lt_monom_mult 4)
    have 6: "punit.lc (rep_list g') = (lc a / lc g) * punit.lc (rep_list g)"
      by (simp add: g'_def rep_list_monom_mult)
    also have 7: "...  punit.lc (rep_list a)"
    proof
      assume "lc a / lc g * punit.lc (rep_list g) = punit.lc (rep_list a)"
      moreover from rep_list g  0 have "punit.lc (rep_list g)  0" by (rule punit.lc_not_0)
      ultimately have "lc a / lc g = punit.lc (rep_list a) / punit.lc (rep_list g)"
        by (simp add: field_simps)
      with False show False ..
    qed
    finally have "punit.lc (rep_list g')  punit.lc (rep_list a)" .
    with 5 have 8: "punit.lt (rep_list ?b) = punit.lt (rep_list a)" unfolding rep_list_minus
      by (rule punit.lt_minus_eqI_3)
    hence "punit.lc (rep_list ?b) = punit.lc (rep_list a) - (lc a / lc g) * punit.lc (rep_list g)"
      unfolding 6[symmetric] by (simp only: punit.lc_def lookup_minus rep_list_minus 5)
    also have "...  0"
    proof
      assume "punit.lc (rep_list a) - lc a / lc g * punit.lc (rep_list g) = 0"
      hence "lc a / lc g * punit.lc (rep_list g) = punit.lc (rep_list a)" by simp
      with 7 show False ..
    qed
    finally have "rep_list ?b  0" by (simp add: punit.lc_eq_zero_iff)
    hence "?b  0" by (auto simp: rep_list_zero)
    hence "lt ?b t lt a" using lt_g' lc_g' by (rule lt_minus_lessI)
    hence "lt ?b t u" by (simp only: assms(4))
    with assms(2) b_in have "sig_red_zero (≼t) G ?b" by (rule is_sig_GB_uptD3)
    moreover note rep_list ?b  0
    moreover have "(≼t) = (≼t)  (≼t) = (≺t)" by simp
    ultimately have "is_sig_red (≼t) (=) G ?b" by (rule sig_red_zero_nonzero)
    then obtain g0 where "g0  G" and "rep_list g0  0"
      and 9: "punit.lt (rep_list g0) adds punit.lt (rep_list ?b)"
      and 10: "punit.lt (rep_list ?b)  lt g0 t punit.lt (rep_list g0)  lt ?b"
      by (rule is_sig_red_top_addsE)
    from 9 have "punit.lt (rep_list g0) adds punit.lt (rep_list a)" by (simp only: 8)
    from 10 have "punit.lt (rep_list a)  lt g0 t punit.lt (rep_list g0)  lt ?b" by (simp only: 8)
    also from lt ?b t lt a have "... t punit.lt (rep_list g0)  lt a" by (rule splus_mono_strict)
    finally have "punit.lt (rep_list a)  lt g0 t punit.lt (rep_list g0)  lt a" .
    have "is_sig_red (≺t) (=) G a"
    proof (rule is_sig_red_top_addsI)
      show "ord_term_lin.is_le_rel (≺t)" by simp
    qed fact+
    with assms(6) show ?thesis ..
  qed
qed

lemma sig_regular_reduced_unique:
  assumes "is_sig_GB_upt d G (lt q)" and "p  dgrad_sig_set d" and "q  dgrad_sig_set d"
    and "lt p = lt q" and "lc p = lc q" and "¬ is_sig_red (≺t) (≼) G p" and "¬ is_sig_red (≺t) (≼) G q"
  shows "rep_list p = rep_list q"
proof (rule ccontr)
  assume "rep_list p  rep_list q"
  hence "rep_list (p - q)  0" by (auto simp: rep_list_minus)
  hence "p - q  0" by (auto simp: rep_list_zero)
  hence "p + (- q)  0" by simp
  moreover from assms(4) have "lt (- q) = lt p" by simp
  moreover from assms(5) have "lc (- q) = - lc p" by simp
  ultimately have "lt (p + (- q)) t lt p" by (rule lt_plus_lessI)
  hence "lt (p - q) t lt q" using assms(4) by simp
  with assms(1) have "is_sig_GB_in d G (lt (p - q))" by (rule is_sig_GB_uptD2)
  moreover from assms(2, 3) have "p - q  dgrad_sig_set d" by (rule dgrad_sig_set_closed_minus)
  ultimately have "sig_red_zero (≼t) G (p - q)" using refl by (rule is_sig_GB_inD)
  hence "is_sig_red (≼t) (≼) G (p - q)" using rep_list (p - q)  0 by (rule sig_red_zero_is_red)
  then obtain g t where "g  G" and t: "t  keys (rep_list (p - q))" and "rep_list g  0"
    and adds: "punit.lt (rep_list g) adds t" and "t  lt g t punit.lt (rep_list g)  lt (p - q)"
    by (rule is_sig_red_addsE)
  note this(5)
  also from lt (p - q) t lt q have "punit.lt (rep_list g)  lt (p - q) t punit.lt (rep_list g)  lt q"
    by (rule splus_mono_strict)
  finally have 1: "t  lt g t punit.lt (rep_list g)  lt q" .
  hence 2: "t  lt g t punit.lt (rep_list g)  lt p" by (simp only: assms(4))
  from t keys_minus have "t  keys (rep_list p)  keys (rep_list q)" unfolding rep_list_minus ..
  thus False
  proof
    assume t_in: "t  keys (rep_list p)"
    hence "t  punit.lt (rep_list p)" by (rule punit.lt_max_keys)
    with g  G t_in rep_list g  0 adds ord_term_lin.is_le_relI(3) ordered_powerprod_lin.is_le_relI(2) 2
    have "is_sig_red (≺t) (≼) G p" by (rule is_sig_red_addsI)
    with assms(6) show False ..
  next
    assume t_in: "t  keys (rep_list q)"
    hence "t  punit.lt (rep_list q)" by (rule punit.lt_max_keys)
    with g  G t_in rep_list g  0 adds ord_term_lin.is_le_relI(3) ordered_powerprod_lin.is_le_relI(2) 1
    have "is_sig_red (≺t) (≼) G q" by (rule is_sig_red_addsI)
    with assms(7) show False ..
  qed
qed

corollary sig_regular_reduced_unique':
  assumes "is_sig_GB_upt d G (lt q)" and "p  dgrad_sig_set d" and "q  dgrad_sig_set d"
    and "lt p = lt q" and "¬ is_sig_red (≺t) (≼) G p" and "¬ is_sig_red (≺t) (≼) G q"
  shows "punit.monom_mult (lc q) 0 (rep_list p) = punit.monom_mult (lc p) 0 (rep_list q)"
proof (cases "p = 0  q = 0")
  case True
  thus ?thesis by (auto simp: rep_list_zero)
next
  case False
  hence "p  0" and "q  0" by simp_all
  hence "lc p  0" and "lc q  0" by (simp_all add: lc_not_0)
  let ?p = "monom_mult (lc q) 0 p"
  let ?q = "monom_mult (lc p) 0 q"
  have "lt ?q = lt q" by (simp add: lt_monom_mult[OF lc p  0 q  0] splus_zero)
  with assms(1) have "is_sig_GB_upt d G (lt ?q)" by simp
  moreover from assms(2) have "?p  dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult_zero)
  moreover from assms(3) have "?q  dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult_zero)
  moreover from lt ?q = lt q have "lt ?p = lt ?q"
    by (simp add: lt_monom_mult[OF lc q  0 p  0] splus_zero assms(4))
  moreover have "lc ?p = lc ?q" by simp
  moreover have "¬ is_sig_red (≺t) (≼) G ?p"
  proof
    assume "is_sig_red (≺t) (≼) G ?p"
    moreover from lc q  0 have "1 / (lc q)  0" by simp
    ultimately have "is_sig_red (≺t) (≼) G (monom_mult (1 / lc q) 0 ?p)" by (rule is_sig_red_monom_mult)
    hence "is_sig_red (≺t) (≼) G p" by (simp add: monom_mult_assoc lc q  0)
    with assms(5) show False ..
  qed
  moreover have "¬ is_sig_red (≺t) (≼) G ?q"
  proof
    assume "is_sig_red (≺t) (≼) G ?q"
    moreover from lc p  0 have "1 / (lc p)  0" by simp
    ultimately have "is_sig_red (≺t) (≼) G (monom_mult (1 / lc p) 0 ?q)" by (rule is_sig_red_monom_mult)
    hence "is_sig_red (≺t) (≼) G q" by (simp add: monom_mult_assoc lc p  0)
    with assms(6) show False ..
  qed
  ultimately have "rep_list ?p = rep_list ?q" by (rule sig_regular_reduced_unique)
  thus ?thesis by (simp only: rep_list_monom_mult)
qed

lemma sig_regular_top_reduced_lt_lc_unique:
  assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p  dgrad_sig_set d" and "q  dgrad_sig_set d"
    and "lt p = lt q" and "(p = 0)  (q = 0)" and "¬ is_sig_red (≺t) (=) G p" and "¬ is_sig_red (≺t) (=) G q"
  shows "punit.lt (rep_list p) = punit.lt (rep_list q)  lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
proof (cases "p = 0")
  case True
  with assms(6) have "q = 0" by simp
  thus ?thesis by (simp add: True)
next
  case False
  with assms(6) have "q  0" by simp
  from False have "lc p  0" by (rule lc_not_0)
  from q  0 have "lc q  0" by (rule lc_not_0)
  from assms(2) have G_sub: "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
  hence "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  with assms(1) obtain p' where p'_red: "(sig_red (≺t) (≺) G)** p p'" and "¬ is_sig_red (≺t) (≺) G p'"
    by (rule sig_irredE_dgrad_max_set)
  from this(1) have lt_p': "lt p' = lt p" and lt_p'': "punit.lt (rep_list p') = punit.lt (rep_list p)"
    and lc_p': "lc p' = lc p" and lc_p'': "punit.lc (rep_list p') = punit.lc (rep_list p)"
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_tail_rtrancl_lt_rep_list,
        rule sig_red_regular_rtrancl_lc, rule sig_red_tail_rtrancl_lc_rep_list)
  have "¬ is_sig_red (≺t) (=) G p'"
  proof
    assume a: "is_sig_red (≺t) (=) G p'"
    hence "rep_list p'  0" using is_sig_red_top_addsE by blast
    hence "rep_list p  0" using (sig_red (≺t) (≺) G)** p p'
      by (auto simp: punit.rtrancl_0 dest!: sig_red_red_rtrancl)
    with a have "is_sig_red (≺t) (=) G p" using lt_p' lt_p'' by (rule is_sig_red_top_cong)
    with assms(7) show False ..
  qed
  with ¬ is_sig_red (≺t) (≺) G p' have 1: "¬ is_sig_red (≺t) (≼) G p'" by (simp add: is_sig_red_top_tail_cases)
  from assms(1) G  dgrad_max_set d obtain q' where q'_red: "(sig_red (≺t) (≺) G)** q q'"
    and "¬ is_sig_red (≺t) (≺) G q'" by (rule sig_irredE_dgrad_max_set)
  from this(1) have lt_q': "lt q' = lt q" and lt_q'': "punit.lt (rep_list q') = punit.lt (rep_list q)"
    and lc_q': "lc q' = lc q" and lc_q'': "punit.lc (rep_list q') = punit.lc (rep_list q)"
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_tail_rtrancl_lt_rep_list,
        rule sig_red_regular_rtrancl_lc, rule sig_red_tail_rtrancl_lc_rep_list)
  have "¬ is_sig_red (≺t) (=) G q'"
  proof
    assume a: "is_sig_red (≺t) (=) G q'"
    hence "rep_list q'  0" using is_sig_red_top_addsE by blast
    hence "rep_list q  0" using (sig_red (≺t) (≺) G)** q q'
      by (auto simp: punit.rtrancl_0 dest!: sig_red_red_rtrancl)
    with a have "is_sig_red (≺t) (=) G q" using lt_q' lt_q'' by (rule is_sig_red_top_cong)
    with assms(8) show False ..
  qed
  with ¬ is_sig_red (≺t) (≺) G q' have 2: "¬ is_sig_red (≺t) (≼) G q'" by (simp add: is_sig_red_top_tail_cases)
  from assms(2) have "is_sig_GB_upt d G (lt q')" by (simp only: lt_q')
  moreover from assms(1) G_sub assms(3) p'_red have "p'  dgrad_sig_set d"
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  moreover from assms(1) G_sub assms(4) q'_red have "q'  dgrad_sig_set d"
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  moreover have "lt p' = lt q'" by (simp only: lt_p' lt_q' assms(5))
  ultimately have eq: "punit.monom_mult (lc q') 0 (rep_list p') = punit.monom_mult (lc p') 0 (rep_list q')"
    using 1 2 by (rule sig_regular_reduced_unique')

  have "lc q * punit.lc (rep_list p) = lc q * punit.lc (rep_list p')" by (simp only: lc_p'')
  also from lc q  0 have "... = punit.lc (punit.monom_mult (lc q') 0 (rep_list p'))"
    by (simp add: lc_q')
  also have "... = punit.lc (punit.monom_mult (lc p') 0 (rep_list q'))" by (simp only: eq)
  also from lc p  0 have "... = lc p * punit.lc (rep_list q')" by (simp add: lc_p')
  also have "... = lc p * punit.lc (rep_list q)" by (simp only: lc_q'')
  finally have *: "lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)" .

  have "punit.lt (rep_list p) = punit.lt (rep_list p')" by (simp only: lt_p'')
  also from lc q  0 have "... = punit.lt (punit.monom_mult (lc q') 0 (rep_list p'))"
    by (simp add: lc_q' punit.lt_monom_mult_zero)
  also have "... = punit.lt (punit.monom_mult (lc p') 0 (rep_list q'))" by (simp only: eq)
  also from lc p  0 have "... = punit.lt (rep_list q')" by (simp add: lc_p' punit.lt_monom_mult_zero)
  also have "... = punit.lt (rep_list q)" by (fact lt_q'')
  finally show ?thesis using * ..
qed

corollary sig_regular_top_reduced_lt_unique:
  assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p  dgrad_sig_set d"
    and "q  dgrad_sig_set d" and "lt p = lt q" and "p  0" and "q  0"
    and "¬ is_sig_red (≺t) (=) G p" and "¬ is_sig_red (≺t) (=) G q"
  shows "punit.lt (rep_list p) = punit.lt (rep_list q)"
proof -
  from assms(6, 7) have "(p = 0)  (q = 0)" by simp
  with assms(1, 2, 3, 4, 5)
  have "punit.lt (rep_list p) = punit.lt (rep_list q)  lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
    using assms(8, 9) by (rule sig_regular_top_reduced_lt_lc_unique)
  thus ?thesis ..
qed

corollary sig_regular_top_reduced_lc_unique:
  assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p  dgrad_sig_set d" and "q  dgrad_sig_set d"
    and "lt p = lt q" and "lc p = lc q" and "¬ is_sig_red (≺t) (=) G p" and "¬ is_sig_red (≺t) (=) G q"
  shows "punit.lc (rep_list p) = punit.lc (rep_list q)"
proof (cases "p = 0")
  case True
  with assms(6) have "q = 0" by (simp add: lc_eq_zero_iff)
  with True show ?thesis by simp
next
  case False
  hence "lc p  0" by (rule lc_not_0)
  hence "lc q  0" by (simp add: assms(6))
  hence "q  0" by (simp add: lc_eq_zero_iff)
  with False have "(p = 0)  (q = 0)" by simp
  with assms(1, 2, 3, 4, 5)
  have "punit.lt (rep_list p) = punit.lt (rep_list q)  lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
    using assms(7, 8) by (rule sig_regular_top_reduced_lt_lc_unique)
  hence "lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)" ..
  also have "... = lc q * punit.lc (rep_list q)" by (simp only: assms(6))
  finally show ?thesis using lc q  0 by simp
qed

text ‹Minimal signature Gr\"obner bases are indeed minimal, at least up to sig-lead-pairs:›

lemma is_min_sig_GB_minimal:
  assumes "is_min_sig_GB d G" and "G'  dgrad_sig_set d"
    and "u. d (pp_of_term u)  dgrad_max d  component_of_term u < length fs  is_sig_GB_in d G' u"
    and "g  G" and "rep_list g  0"
  obtains g' where "g'  G'" and "rep_list g'  0" and "lt g' = lt g"
    and "punit.lt (rep_list g') = punit.lt (rep_list g)"
proof -
  from assms(1) have "G  dgrad_sig_set d"
    and 1: "u. d (pp_of_term u)  dgrad_max d  component_of_term u < length fs  is_sig_GB_in d G u"
    and 2: "g0. g0  G  ¬ is_sig_red (≼t) (=) (G - {g0}) g0"
    by (simp_all add: is_min_sig_GB_def)
  from assms(4) have 3: "¬ is_sig_red (≼t) (=) (G - {g}) g" by (rule 2)

  from assms(5) have "g  0" by (auto simp: rep_list_zero)
  from assms(4) G  dgrad_sig_set d have "g  dgrad_sig_set d" ..
  hence "d (lp g)  dgrad_max d" and "component_of_term (lt g) < length fs"
    by (rule dgrad_sig_setD_lp, rule dgrad_sig_setD_lt[OF _ g  0])
  hence "is_sig_GB_in d G' (lt g)" by (rule assms(3))
  hence "sig_red_zero (≼t) G' g" using g  dgrad_sig_set d refl by (rule is_sig_GB_inD)
  moreover note assms(5)
  moreover have "(≼t) = (≼t)  (≼t) = (≺t)" by simp
  ultimately have "is_sig_red (≼t) (=) G' g" by (rule sig_red_zero_nonzero)
  then obtain g' where "g'  G'" and "rep_list g'  0"
    and adds1: "punit.lt (rep_list g') adds punit.lt (rep_list g)"
    and le1: "punit.lt (rep_list g)  lt g' t punit.lt (rep_list g')  lt g"
    by (rule is_sig_red_top_addsE)

  from rep_list g'  0 have "g'  0" by (auto simp: rep_list_zero)
  from g'  G' assms(2) have "g'  dgrad_sig_set d" ..
  hence "d (lp g')  dgrad_max d" and "component_of_term (lt g') < length fs"
    by (rule dgrad_sig_setD_lp, rule dgrad_sig_setD_lt[OF _ g'  0])
  hence "is_sig_GB_in d G (lt g')" by (rule 1)
  hence "sig_red_zero (≼t) G g'" using g'  dgrad_sig_set d refl by (rule is_sig_GB_inD)
  moreover note rep_list g'  0
  moreover have "(≼t) = (≼t)  (≼t) = (≺t)" by simp
  ultimately have "is_sig_red (≼t) (=) G g'" by (rule sig_red_zero_nonzero)
  then obtain g0 where "g0  G" and "rep_list g0  0"
    and adds2: "punit.lt (rep_list g0) adds punit.lt (rep_list g')"
    and le2: "punit.lt (rep_list g')  lt g0 t punit.lt (rep_list g0)  lt g'"
    by (rule is_sig_red_top_addsE)

  have eq1: "g0 = g"
  proof (rule ccontr)
    assume "g0  g"
    with g0  G have "g0  G - {g}" by simp
    moreover note rep_list g0  0 assms(5)
    moreover from adds2 adds1 have "punit.lt (rep_list g0) adds punit.lt (rep_list g)"
      by (rule adds_trans)
    moreover have "ord_term_lin.is_le_rel (≼t)" by simp
    moreover have "punit.lt (rep_list g)  lt g0 t punit.lt (rep_list g0)  lt g"
    proof (rule ord_term_canc)
      have "punit.lt (rep_list g')  (punit.lt (rep_list g)  lt g0) =
            punit.lt (rep_list g)  (punit.lt (rep_list g')  lt g0)" by (fact splus_left_commute)
      also from le2 have "... t punit.lt (rep_list g)  (punit.lt (rep_list g0)  lt g')"
        by (rule splus_mono)
      also have "... = punit.lt (rep_list g0)  (punit.lt (rep_list g)  lt g')"
        by (fact splus_left_commute)
      also from le1 have "... t punit.lt (rep_list g0)  (punit.lt (rep_list g')  lt g)"
        by (rule splus_mono)
      also have "... = punit.lt (rep_list g')  (punit.lt (rep_list g0)  lt g)"
        by (fact splus_left_commute)
      finally show "punit.lt (rep_list g')  (punit.lt (rep_list g)  lt g0) t
                    punit.lt (rep_list g')  (punit.lt (rep_list g0)  lt g)" .
    qed
    ultimately have "is_sig_red (≼t) (=) (G - {g}) g" by (rule is_sig_red_top_addsI)
    with 3 show False ..
  qed

  from adds2 adds1 have eq2: "punit.lt (rep_list g') = punit.lt (rep_list g)" by (simp add: eq1 adds_antisym)
  with le1 le2 have "punit.lt (rep_list g)  lt g' = punit.lt (rep_list g)  lt g" by (simp add: eq1)
  hence "lt g' = lt g" by (simp only: splus_left_canc)
  with g'  G' rep_list g'  0 show ?thesis using eq2 ..
qed

lemma sig_red_zero_regularI_adds:
  assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)"
    and "p  dgrad_sig_set d" and "q  dgrad_sig_set d" and "p  0" and "sig_red_zero (≺t) G p"
    and "lt p addst lt q"
  shows "sig_red_zero (≺t) G q"
proof (cases "q = 0")
  case True
  hence "rep_list q = 0" by (simp only: rep_list_zero)
  with rtrancl_refl[to_pred] show ?thesis by (rule sig_red_zeroI)
next
  case False
  hence "lc q  0" by (rule lc_not_0)
  moreover from assms(5) have "lc p  0" by (rule lc_not_0)
  ultimately have "lc q / lc p  0" by simp
  from assms(7) have eq1: "(lp q - lp p)  lt p = lt q"
    by (metis add_diff_cancel_right' adds_termE pp_of_term_splus)
  from assms(7) have "lp p adds lp q" by (simp add: adds_term_def)
  with assms(1) have "d (lp q - lp p)  d (lp q)" by (rule dickson_grading_minus)
  also from assms(4) have "...  dgrad_max d" by (rule dgrad_sig_setD_lp)
  finally have "d (lp q - lp p)  dgrad_max d" .
  from assms(2) have G_sub: "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
  hence "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)

  let ?mult = "λr. monom_mult (lc q / lc p) (lp q - lp p) r"
  from assms(6) obtain p' where p_red: "(sig_red (≺t) (≼) G)** p p'" and "rep_list p' = 0"
    by (rule sig_red_zeroE)
  from p_red have "lt p' = lt p" and "lc p' = lc p"
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
  hence "p'  0" using lc p  0 by auto
  with lc q / lc p  0 have "?mult p'  0" by (simp add: monom_mult_eq_zero_iff)
  from lc q / lc p  0 p'  0 have "lt (?mult p') = lt q"
    by (simp add: lt_monom_mult lt p' = lt p eq1)
  from lc p  0 have "lc (?mult p') = lc q" by (simp add: lc p' = lc p)
  from p_red have mult_p_red: "(sig_red (≺t) (≼) G)** (?mult p) (?mult p')"
    by (rule sig_red_rtrancl_monom_mult)
  have "rep_list (?mult p') = 0" by (simp add: rep_list_monom_mult rep_list p' = 0)
  hence mult_p'_irred: "¬ is_sig_red (≺t) (≼) G (?mult p')"
    using is_sig_red_addsE by fastforce
  from assms(1) G_sub assms(3) p_red have "p'  dgrad_sig_set d"
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  with assms(1) d (lp q - lp p)  dgrad_max d have "?mult p'  dgrad_sig_set d"
    by (rule dgrad_sig_set_closed_monom_mult)

  from assms(1) G  dgrad_max_set d obtain q' where q_red: "(sig_red (≺t) (≼) G)** q q'"
    and q'_irred: "¬ is_sig_red (≺t) (≼) G q'" by (rule sig_irredE_dgrad_max_set)
  from q_red have "lt q' = lt q" and "lc q' = lc q"
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
  hence "q'  0" using lc q  0 by auto

  from assms(2) have "is_sig_GB_upt d G (lt (?mult p'))" by (simp only: lt (?mult p') = lt q)
  moreover from assms(1) G_sub assms(4) q_red have "q'  dgrad_sig_set d"
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  moreover note ?mult p'  dgrad_sig_set d
  moreover have "lt q' = lt (?mult p')" by (simp only: lt (?mult p') = lt q lt q' = lt q)
  moreover have "lc q' = lc (?mult p')" by (simp only: lc (?mult p') = lc q lc q' = lc q)
  ultimately have "rep_list q' = rep_list (?mult p')" using q'_irred mult_p'_irred
    by (rule sig_regular_reduced_unique)
  with rep_list (?mult p') = 0 have "rep_list q' = 0" by simp
  with q_red show ?thesis by (rule sig_red_zeroI)
qed

lemma is_syz_sigI:
  assumes "s  0" and "lt s = u" and "s  dgrad_sig_set d" and "rep_list s = 0"
  shows "is_syz_sig d u"
  unfolding is_syz_sig_def using assms by blast

lemma is_syz_sigE:
  assumes "is_syz_sig d u"
  obtains r where "r  0" and "lt r = u" and "r  dgrad_sig_set d" and "rep_list r = 0"
  using assms unfolding is_syz_sig_def by blast

lemma is_syz_sig_adds:
  assumes "dickson_grading d" and "is_syz_sig d u" and "u addst v"
    and "d (pp_of_term v)  dgrad_max d"
  shows "is_syz_sig d v"
proof -
  from assms(2) obtain s where "s  0" and "lt s = u" and "s  dgrad_sig_set d"
    and "rep_list s = 0" by (rule is_syz_sigE)
  from assms(3) obtain t where v: "v = t  u" by (rule adds_termE)
  show ?thesis
  proof (rule is_syz_sigI)
    from s  0 show "monom_mult 1 t s  0" by (simp add: monom_mult_eq_zero_iff)
  next
    from s  0 show "lt (monom_mult 1 t s) = v" by (simp add: lt_monom_mult v lt s = u)
  next
    from assms(4) have "d (t + pp_of_term u)  dgrad_max d" by (simp add: v term_simps)
    with assms(1) have "d t  dgrad_max d" by (simp add: dickson_gradingD1)
    with assms(1) show "monom_mult 1 t s  dgrad_sig_set d" using s  dgrad_sig_set d
      by (rule dgrad_sig_set_closed_monom_mult)
  next
    show "rep_list (monom_mult 1 t s) = 0" by (simp add: rep_list s = 0 rep_list_monom_mult)
  qed
qed

lemma syzygy_crit:
  assumes "dickson_grading d" and "is_sig_GB_upt d G u" and "is_syz_sig d u"
    and "p  dgrad_sig_set d" and "lt p = u"
  shows "sig_red_zero (≺t) G p"
proof -
  from assms(3) obtain s where "s  0" and "lt s = u" and "s  dgrad_sig_set d"
    and "rep_list s = 0" by (rule is_syz_sigE)
  note assms(1)
  moreover from assms(2) have "is_sig_GB_upt d G (lt p)" by (simp only: assms(5))
  moreover note s  dgrad_sig_set d assms(4) s  0
  moreover from rtranclp.rtrancl_refl rep_list s = 0 have "sig_red_zero (≺t) G s"
    by (rule sig_red_zeroI)
  moreover have "lt s addst lt p" by (simp only: assms(5) lt s = u adds_term_refl)
  ultimately show ?thesis by (rule sig_red_zero_regularI_adds)
qed

lemma lemma_21:
  assumes "dickson_grading d" and "is_sig_GB_upt d G (lt p)" and "p  dgrad_sig_set d" and "g  G"
    and "rep_list p  0" and "rep_list g  0" and "lt g addst lt p"
    and "punit.lt (rep_list g) adds punit.lt (rep_list p)"
  shows "is_sig_red (≼t) (=) G p"
proof -
  let ?lp = "punit.lt (rep_list p)"
  define s where "s = ?lp - punit.lt (rep_list g)"
  from assms(8) have s: "?lp = s + punit.lt (rep_list g)" by (simp add: s_def minus_plus)
  from assms(7) obtain t where lt_p: "lt p = t  lt g" by (rule adds_termE)
  show ?thesis
  proof (cases "s  lt g t lt p")
    case True
    hence "?lp  lt g t punit.lt (rep_list g)  lt p"
      by (simp add: s splus_assoc splus_left_commute[of s] splus_mono)
    with assms(4, 6, 5, 8) ord_term_lin.is_le_relI(2) show ?thesis
      by (rule is_sig_red_top_addsI)
  next
    case False
    hence "lt p t s  lt g" by simp
    hence "t  s" by (simp add: lt_p ord_term_strict_canc_left)
    hence "t + punit.lt (rep_list g)  s + punit.lt (rep_list g)" by (rule plus_monotone_strict)
    hence "t + punit.lt (rep_list g)  ?lp" by (simp only: s)
    from assms(5) have "p  0" by (auto simp: rep_list_zero)
    hence "lc p  0" by (rule lc_not_0)
    from assms(6) have "g  0" by (auto simp: rep_list_zero)
    hence "lc g  0" by (rule lc_not_0)
    with lc p  0 have 1: "lc p / lc g  0" by simp

    let ?g = "monom_mult (lc p / lc g) t g"
    from 1 g  0 have "lt ?g = lt p" unfolding lt_p by (rule lt_monom_mult)
    from lc g  0 have "lc ?g = lc p" by simp
    have "punit.lt (rep_list ?g) = t + punit.lt (rep_list g)"
      unfolding rep_list_monom_mult using 1 assms(6) by (rule punit.lt_monom_mult[simplified])
    also have "...  ?lp" by fact
    finally have "punit.lt (rep_list ?g)  ?lp" .
    hence lt_pg: "punit.lt (rep_list (p - ?g)) = ?lp" and "rep_list p  rep_list ?g"
      by (auto simp: rep_list_minus punit.lt_minus_eqI_2)
    from this(2) have "rep_list (p - ?g)  0" and "p - ?g  0"
      by (auto simp: rep_list_minus rep_list_zero)

    from assms(2) have "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
    note assms(1)
    moreover have "d t  dgrad_max d"
    proof (rule le_trans)
      have "lp p = t + lp g" by (simp add: lt_p term_simps)
      with assms(1) show "d t  d (lp p)" by (simp add: dickson_grading_adds_imp_le)
    next
      from assms(3) show "d (lp p)  dgrad_max d" by (rule dgrad_sig_setD_lp)
    qed
    moreover from assms(4) G  dgrad_sig_set d have "g  dgrad_sig_set d" ..
    ultimately have "?g  dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult)

    note assms(2)
    moreover from assms(3) ?g  dgrad_sig_set d have "p - ?g  dgrad_sig_set d"
      by (rule dgrad_sig_set_closed_minus)
    moreover from p - ?g  0 lt ?g = lt p lc ?g = lc p have "lt (p - ?g) t lt p"
      by (rule lt_minus_lessI)
    ultimately have "sig_red_zero (≼t) G (p - ?g)"
      by (rule is_sig_GB_uptD3)
    moreover note rep_list (p - ?g)  0
    moreover have "(≼t) = (≼t)  (≼t) = (≺t)" by simp
    ultimately have "is_sig_red (≼t) (=) G (p - ?g)" by (rule sig_red_zero_nonzero)
    then obtain g1 where "g1  G" and "rep_list g1  0"
      and 2: "punit.lt (rep_list g1) adds punit.lt (rep_list (p - ?g))"
      and 3: "punit.lt (rep_list (p - ?g))  lt g1 t punit.lt (rep_list g1)  lt (p - ?g)"
      by (rule is_sig_red_top_addsE)
    from g1  G rep_list g1  0 assms(5) show ?thesis
    proof (rule is_sig_red_top_addsI)
      from 2 show "punit.lt (rep_list g1) adds punit.lt (rep_list p)" by (simp only: lt_pg)
    next
      have "?lp  lt g1 = punit.lt (rep_list (p - ?g))  lt g1" by (simp only: lt_pg)
      also have "... t punit.lt (rep_list g1)  lt (p - ?g)" by (fact 3)
      also from lt (p - ?g) t lt p have "... t punit.lt (rep_list g1)  lt p"
        by (rule splus_mono_strict)
      finally show "?lp  lt g1 t punit.lt (rep_list g1)  lt p" by (rule ord_term_lin.less_imp_le)
    qed simp
  qed
qed

subsubsection ‹Rewrite Bases›

definition is_rewrite_ord :: "(('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  bool"
  where "is_rewrite_ord rword  (reflp rword  transp rword  (a b. rword a b  rword b a) 
                                  (a b. rword a b  rword b a  fst a = fst b) 
                                  (d G a b. dickson_grading d  is_sig_GB_upt d G (lt b) 
                                          a  G  b  G  a  0  b  0  lt a addst lt b 
                                          ¬ is_sig_red (≺t) (=) G b  rword (spp_of a) (spp_of b)))"

definition is_canon_rewriter :: "(('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  ('t 0 'b) set  't  ('t 0 'b)  bool"
  where "is_canon_rewriter rword A u p 
                  (p  A  p  0  lt p addst u  (aA. a  0  lt a addst u  rword (spp_of a) (spp_of p)))"

definition is_RB_in :: "('a  nat)  (('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  ('t 0 'b) set  't  bool"
  where "is_RB_in d rword G u 
            ((g. is_canon_rewriter rword G u g  ¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)) 
             is_syz_sig d u)"

definition is_RB_upt :: "('a  nat)  (('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  ('t 0 'b) set  't  bool"
  where "is_RB_upt d rword G u 
            (G  dgrad_sig_set d  (v. v t u  d (pp_of_term v)  dgrad_max d 
                                      component_of_term v < length fs  is_RB_in d rword G v))"

lemma is_rewrite_ordI:
  assumes "reflp rword" and "transp rword" and "a b. rword a b  rword b a"
    and "a b. rword a b  rword b a  fst a = fst b"
    and "d G a b. dickson_grading d  is_sig_GB_upt d G (lt b)  a  G  b  G 
                   a  0  b  0  lt a addst lt b  ¬ is_sig_red (≺t) (=) G b  rword (spp_of a) (spp_of b)"
  shows "is_rewrite_ord rword"
  unfolding is_rewrite_ord_def using assms by blast

lemma is_rewrite_ordD1: "is_rewrite_ord rword  rword a a"
  by (simp add: is_rewrite_ord_def reflpD)

lemma is_rewrite_ordD2: "is_rewrite_ord rword  rword a b  rword b c  rword a c"
  by (auto simp: is_rewrite_ord_def dest: transpD)

lemma is_rewrite_ordD3:
  assumes "is_rewrite_ord rword"
    and "rword a b  thesis"
    and "¬ rword a b  rword b a  thesis"
  shows thesis
proof -
  from assms(1) have disj: "rword a b  rword b a"
    by (simp add: is_rewrite_ord_def del: split_paired_All)
  show ?thesis
  proof (cases "rword a b")
    case True
    thus ?thesis by (rule assms(2))
  next
    case False
    moreover from this disj have "rword b a" by simp
    ultimately show ?thesis by (rule assms(3))
  qed
qed

lemma is_rewrite_ordD4:
  assumes "is_rewrite_ord rword" and "rword a b" and "rword b a"
  shows "fst a = fst b"
  using assms unfolding is_rewrite_ord_def by blast

lemma is_rewrite_ordD4':
  assumes "is_rewrite_ord rword" and "rword (spp_of a) (spp_of b)" and "rword (spp_of b) (spp_of a)"
  shows "lt a = lt b"
proof -
  from assms have "fst (spp_of a) = fst (spp_of b)" by (rule is_rewrite_ordD4)
  thus ?thesis by (simp add: spp_of_def)
qed

lemma is_rewrite_ordD5:
  assumes "is_rewrite_ord rword" and "dickson_grading d" and "is_sig_GB_upt d G (lt b)"
    and "a  G" and "b  G" and "a  0" and "b  0" and "lt a addst lt b"
    and "¬ is_sig_red (≺t) (=) G b"
  shows "rword (spp_of a) (spp_of b)"
  using assms unfolding is_rewrite_ord_def by blast

lemma is_canon_rewriterI:
  assumes "p  A" and "p  0" and "lt p addst u"
    and "a. a  A  a  0  lt a addst u  rword (spp_of a) (spp_of p)"
  shows "is_canon_rewriter rword A u p"
  unfolding is_canon_rewriter_def using assms by blast

lemma is_canon_rewriterD1: "is_canon_rewriter rword A u p  p  A"
  by (simp add: is_canon_rewriter_def)

lemma is_canon_rewriterD2: "is_canon_rewriter rword A u p  p  0"
  by (simp add: is_canon_rewriter_def)

lemma is_canon_rewriterD3: "is_canon_rewriter rword A u p  lt p addst u"
  by (simp add: is_canon_rewriter_def)

lemma is_canon_rewriterD4:
  "is_canon_rewriter rword A u p  a  A  a  0  lt a addst u  rword (spp_of a) (spp_of p)"
  by (simp add: is_canon_rewriter_def)

lemmas is_canon_rewriterD = is_canon_rewriterD1 is_canon_rewriterD2 is_canon_rewriterD3 is_canon_rewriterD4

lemma is_rewrite_ord_finite_canon_rewriterE:
  assumes "is_rewrite_ord rword" and "finite A" and "a  A" and "a  0" and "lt a addst u"
  obtains p where "is_canon_rewriter rword A u p"
proof -
  let ?A = "{x. x  A  x  0  lt x addst u}"
  let ?rel = "λx y. strict rword (spp_of y) (spp_of x)"
  have "finite ?A"
  proof (rule finite_subset)
    show "?A  A" by blast
  qed fact
  moreover have "?A  {}"
  proof
    from assms(3, 4, 5) have "a  ?A" by simp
    also assume "?A = {}"
    finally show False by simp
  qed
  moreover have "irreflp ?rel"
  proof -
    from assms(1) have "reflp rword" by (simp add: is_rewrite_ord_def)
    thus ?thesis by (simp add: reflp_def irreflp_def)
  qed
  moreover have "transp ?rel"
  proof -
    from assms(1) have "transp rword" by (simp add: is_rewrite_ord_def)
    thus ?thesis by (auto simp: transp_def simp del: split_paired_All)
  qed
  ultimately obtain p where "p  ?A" and *: "b. ?rel b p  b  ?A" by (rule finite_minimalE, blast)
  from this(1) have "p  A" and "p  0" and "lt p addst u" by simp_all
  show ?thesis
  proof (rule, rule is_canon_rewriterI)
    fix q
    assume "q  A" and "q  0" and "lt q addst u"
    hence "q  ?A" by simp
    with * have "¬ ?rel q p" by blast
    hence disj: "¬ rword (spp_of p) (spp_of q)  rword (spp_of q) (spp_of p)" by simp
    from assms(1) show "rword (spp_of q) (spp_of p)"
    proof (rule is_rewrite_ordD3)
      assume "¬ rword (spp_of q) (spp_of p)" and "rword (spp_of p) (spp_of q)"
      with disj show ?thesis by simp
    qed
  qed fact+
qed

lemma is_rewrite_ord_canon_rewriterD1:
  assumes "is_rewrite_ord rword" and "is_canon_rewriter rword A u p" and "is_canon_rewriter rword A v q"
    and "lt p addst v" and "lt q addst u"
  shows "lt p = lt q"
proof -
  from assms(2) have "p  A" and "p  0"
    and 1: "a. a  A  a  0  lt a addst u  rword (spp_of a) (spp_of p)"
    by (rule is_canon_rewriterD)+
  from assms(3) have "q  A" and "q  0"
    and 2: "a. a  A  a  0  lt a addst v  rword (spp_of a) (spp_of q)"
    by (rule is_canon_rewriterD)+
  note assms(1)
  moreover from p  A p  0 assms(4) have "rword (spp_of p) (spp_of q)" by (rule 2)
  moreover from q  A q  0 assms(5) have "rword (spp_of q) (spp_of p)" by (rule 1)
  ultimately show ?thesis by (rule is_rewrite_ordD4')
qed

corollary is_rewrite_ord_canon_rewriterD2:
  assumes "is_rewrite_ord rword" and "is_canon_rewriter rword A u p" and "is_canon_rewriter rword A u q"
  shows "lt p = lt q"
  using assms
proof (rule is_rewrite_ord_canon_rewriterD1)
  from assms(2) show "lt p addst u" by (rule is_canon_rewriterD)
next
  from assms(3) show "lt q addst u" by (rule is_canon_rewriterD)
qed

lemma is_rewrite_ord_canon_rewriterD3:
  assumes "is_rewrite_ord rword" and "dickson_grading d" and "is_canon_rewriter rword A u p"
    and "a  A" and "a  0" and "lt a addst u" and "is_sig_GB_upt d A (lt a)"
    and "lt p addst lt a" and "¬ is_sig_red (≺t) (=) A a"
  shows "lt p = lt a"
proof -
  note assms(1)
  moreover from assms(1, 2, 7) _ assms(4) _ assms(5, 8, 9) have "rword (spp_of p) (spp_of a)"
  proof (rule is_rewrite_ordD5)
    from assms(3) show "p  A" and "p  0" by (rule is_canon_rewriterD)+
  qed
  moreover from assms(3, 4, 5, 6) have "rword (spp_of a) (spp_of p)" by (rule is_canon_rewriterD4)
  ultimately show ?thesis by (rule is_rewrite_ordD4')
qed

lemma is_RB_inI1:
  assumes "is_canon_rewriter rword G u g" and "¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)"
  shows "is_RB_in d rword G u"
  unfolding is_RB_in_def using assms is_canon_rewriterD1 by blast

lemma is_RB_inI2:
  assumes "is_syz_sig d u"
  shows "is_RB_in d rword G u"
  unfolding is_RB_in_def Let_def using assms by blast

lemma is_RB_inE:
  assumes "is_RB_in d rword G u"
    and "is_syz_sig d u  thesis"
    and "g. ¬ is_syz_sig d u  is_canon_rewriter rword G u g 
            ¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)  thesis"
  shows thesis
  using assms unfolding is_RB_in_def by blast

lemma is_RB_inD:
  assumes "dickson_grading d" and "G  dgrad_sig_set d" and "is_RB_in d rword G u"
    and "¬ is_syz_sig d u" and "d (pp_of_term u)  dgrad_max d"
    and "is_canon_rewriter rword G u g"
  shows "rep_list g  0"
proof
  assume a: "rep_list g = 0"
  from assms(1) have "is_syz_sig d u"
  proof (rule is_syz_sig_adds)
    show "is_syz_sig d (lt g)"
    proof (rule is_syz_sigI)
      from assms(6) show "g  0" by (rule is_canon_rewriterD2)
    next
      from assms(6) have "g  G" by (rule is_canon_rewriterD1)
      thus "g  dgrad_sig_set d" using assms(2) ..
    qed (fact refl, fact a)
  next
    from assms(6) show "lt g addst u" by (rule is_canon_rewriterD3)
  qed fact
  with assms(4) show False ..
qed

lemma is_RB_uptI:
  assumes "G  dgrad_sig_set d"
    and "v. v t u  d (pp_of_term v)  dgrad_max d  component_of_term v < length fs 
            is_RB_in d canon G v"
  shows "is_RB_upt d canon G u"
  unfolding is_RB_upt_def using assms by blast

lemma is_RB_uptD1:
  assumes "is_RB_upt d canon G u"
  shows "G  dgrad_sig_set d"
  using assms unfolding is_RB_upt_def by blast

lemma is_RB_uptD2:
  assumes "is_RB_upt d canon G u" and "v t u" and "d (pp_of_term v)  dgrad_max d"
    and "component_of_term v < length fs"
  shows "is_RB_in d canon G v"
  using assms unfolding is_RB_upt_def by blast

lemma is_RB_in_UnI:
  assumes "is_RB_in d rword G u" and "h. h  H  u t lt h"
  shows "is_RB_in d rword (H  G) u"
  using assms(1)
proof (rule is_RB_inE)
  assume "is_syz_sig d u"
  thus "is_RB_in d rword (H  G) u" by (rule is_RB_inI2)
next
  fix g'
  assume crw: "is_canon_rewriter rword G u g'"
    and irred: "¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term u - lp g') g')"
  from crw have "g'  G" and "g'  0" and "lt g' addst u"
    and max: "a. a  G  a  0  lt a addst u  rword (spp_of a) (spp_of g')"
    by (rule is_canon_rewriterD)+
  show "is_RB_in d rword (H  G) u"
  proof (rule is_RB_inI1)
    show "is_canon_rewriter rword (H  G) u g'"
    proof (rule is_canon_rewriterI)
      from g'  G show "g'  H  G" by simp
    next
      fix a
      assume "a  H  G" and "a  0" and "lt a addst u"
      from this(1) show "rword (spp_of a) (spp_of g')"
      proof
        assume "a  H"
        with lt a addst u have "lt a addst u" by simp
        hence "lt a t u" by (rule ord_adds_term)
        moreover from a  H have "u t lt a" by (rule assms(2))
        ultimately show ?thesis by simp
      next
        assume "a  G"
        thus ?thesis using a  0 lt a addst u by (rule max)
      qed
    qed fact+
  next
    show "¬ is_sig_red (≺t) (=) (H  G) (monom_mult 1 (pp_of_term u - lp g') g')"
      (is "¬ is_sig_red _ _ _ ?g")
    proof
      assume "is_sig_red (≺t) (=) (H  G) ?g"
      with irred have "is_sig_red (≺t) (=) H ?g" by (simp add: is_sig_red_Un del: Un_insert_left)
      then obtain h where "h  H" and "is_sig_red (≺t) (=) {h} ?g" by (rule is_sig_red_singletonI)
      from this(2) have "lt h t lt ?g" by (rule is_sig_red_regularD_lt)
      also from g'  0 lt g' addst u have "... = u"
        by (auto simp: lt_monom_mult adds_term_alt pp_of_term_splus)
      finally have "lt h t u" .
      moreover from h  H have "u t lt h" by (rule assms(2))
      ultimately show False by simp
    qed
  qed
qed

corollary is_RB_in_insertI:
  assumes "is_RB_in d rword G u" and "u t lt g"
  shows "is_RB_in d rword (insert g G) u"
proof -
  from assms(1) have "is_RB_in d rword ({g}  G) u"
  proof (rule is_RB_in_UnI)
    fix h
    assume "h  {g}"
    with assms(2) show "u t lt h" by simp
  qed
  thus ?thesis by simp
qed

corollary is_RB_upt_UnI:
  assumes "is_RB_upt d rword G u" and "H  dgrad_sig_set d" and "h. h  H  u t lt h"
  shows "is_RB_upt d rword (H  G) u"
proof (rule is_RB_uptI)
  from assms(1) have "G  dgrad_sig_set d" by (rule is_RB_uptD1)
  with assms(2) show "H  G  dgrad_sig_set d" by (rule Un_least)
next
  fix v
  assume "v t u" and "d (pp_of_term v)  dgrad_max d" and "component_of_term v < length fs"
  with assms(1) have "is_RB_in d rword G v" by (rule is_RB_uptD2)
  moreover from v t u assms(3) have "h. h  H  v t lt h" by (rule ord_term_lin.less_le_trans)
  ultimately show "is_RB_in d rword (H  G) v" by (rule is_RB_in_UnI)
qed

corollary is_RB_upt_insertI:
  assumes "is_RB_upt d rword G u" and "g  dgrad_sig_set d" and "u t lt g"
  shows "is_RB_upt d rword (insert g G) u"
proof -
  from assms(1) have "is_RB_upt d rword ({g}  G) u"
  proof (rule is_RB_upt_UnI)
    from assms(2) show "{g}  dgrad_sig_set d" by simp
  next
    fix h
    assume "h  {g}"
    with assms(3) show "u t lt h" by simp
  qed
  thus ?thesis by simp
qed

lemma is_RB_upt_is_sig_GB_upt:
  assumes "dickson_grading d" and "is_RB_upt d rword G u"
  shows "is_sig_GB_upt d G u"
proof (rule ccontr)
  let ?Q = "{v. v t u  d (pp_of_term v)  dgrad_max d  component_of_term v < length fs  ¬ is_sig_GB_in d G v}"
  have Q_sub: "pp_of_term ` ?Q  dgrad_set d (dgrad_max d)" by blast
  from assms(2) have G_sub: "G  dgrad_sig_set d" by (rule is_RB_uptD1)
  hence "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  assume "¬ is_sig_GB_upt d G u"
  with G_sub obtain v' where "v'  ?Q" unfolding is_sig_GB_upt_def by blast
  with assms(1) obtain v where "v  ?Q" and min: "y. y t v  y  ?Q" using Q_sub
    by (rule ord_term_minimum_dgrad_set, blast)
  from v  ?Q have "v t u" and "d (pp_of_term v)  dgrad_max d" and "component_of_term v < length fs"
    and "¬ is_sig_GB_in d G v" by simp_all
  from assms(2) this(1, 2, 3) have "is_RB_in d rword G v" by (rule is_RB_uptD2)
  from ¬ is_sig_GB_in d G v obtain r where "lt r = v" and "r  dgrad_sig_set d" and "¬ sig_red_zero (≼t) G r"
    unfolding is_sig_GB_in_def by blast
  from this(3) have "rep_list r  0" by (auto simp: sig_red_zero_def)
  hence "r  0" by (auto simp: rep_list_zero)
  hence "lc r  0" by (rule lc_not_0)

  from G_sub have "is_sig_GB_upt d G v"
  proof (rule is_sig_GB_uptI)
    fix w
    assume dw: "d (pp_of_term w)  dgrad_max d" and cp: "component_of_term w < length fs"
    assume "w t v"
    hence "w  ?Q" by (rule min)
    hence "¬ w t u  is_sig_GB_in d G w" by (simp add: dw cp)
    thus "is_sig_GB_in d G w"
    proof
      assume "¬ w t u"
      moreover from w t v v t u have "w t u" by (rule ord_term_lin.less_trans)
      ultimately show ?thesis ..
    qed
  qed

  from is_RB_in d rword G v have "sig_red_zero (≼t) G r"
  proof (rule is_RB_inE)
    assume "is_syz_sig d v"
    have "sig_red_zero (≺t) G r" by (rule syzygy_crit, fact+)
    thus ?thesis by (rule sig_red_zero_sing_regI)
  next
    fix g
    assume a: "¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
    assume "is_canon_rewriter rword G v g"
    hence "g  G" and "g  0" and "lt g addst v" by (rule is_canon_rewriterD)+
    assume "¬ is_syz_sig d v"
    from g  G G_sub have "g  dgrad_sig_set d" ..
    from g  0 have "lc g  0" by (rule lc_not_0)
    with lc r  0 have "lc r / lc g  0" by simp
    from lt g addst v have "lt g addst lt r" by (simp only: lt r = v)
    hence eq1: "(lp r - lp g)  lt g = lt r" by (metis add_implies_diff adds_termE pp_of_term_splus)

    let ?h = "monom_mult (lc r / lc g) (lp r - lp g) g"
    from lc g  0 lc r  0 g  0 have "?h  0" by (simp add: monom_mult_eq_zero_iff)
    have h_irred: "¬ is_sig_red (≺t) (=) G ?h"
    proof
      assume "is_sig_red (≺t) (=) G ?h"
      moreover from lc g  0 lc r  0 have "lc g / lc r  0" by simp
      ultimately have "is_sig_red (≺t) (=) G (monom_mult (lc g / lc r) 0 ?h)" by (rule is_sig_red_monom_mult)
      with lc g  0 lc r  0 have "is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
        by (simp add: monom_mult_assoc lt r = v)
      with a show False ..
    qed
    from lc r / lc g  0 g  0 have "lt ?h = lt r" by (simp add: lt_monom_mult eq1)
    hence "lt ?h = v" by (simp only: lt r = v)
    from lc g  0 have "lc ?h = lc r" by simp
    from assms(1) _ g  dgrad_sig_set d have "?h  dgrad_sig_set d"
    proof (rule dgrad_sig_set_closed_monom_mult)
      from lt g addst lt r have "lp g adds lp r" by (simp add: adds_term_def)
      with assms(1) have "d (lp r - lp g)  d (lp r)" by (rule dickson_grading_minus)
      also from r  dgrad_sig_set d have "...  dgrad_max d" by (rule dgrad_sig_setD_lp)
      finally show "d (lp r - lp g)  dgrad_max d" .
    qed
    have "rep_list ?h  0"
    proof
      assume "rep_list ?h = 0"
      with ?h  0 lt ?h = v ?h  dgrad_sig_set d have "is_syz_sig d v" by (rule is_syz_sigI)
      with ¬ is_syz_sig d v show False ..
    qed
    hence "rep_list g  0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
    hence "punit.lc (rep_list g)  0" by (rule punit.lc_not_0)
    from assms(1) G  dgrad_max_set d obtain s where s_red: "(sig_red (≺t) (≼) G)** r s"
      and s_irred: "¬ is_sig_red (≺t) (≼) G s" by (rule sig_irredE_dgrad_max_set)
    from s_red have s_red': "(sig_red (≼t) (≼) G)** r s" by (rule sig_red_rtrancl_sing_regI)
    have "rep_list s  0"
    proof
      assume "rep_list s = 0"
      with s_red' have "sig_red_zero (≼t) G r" by (rule sig_red_zeroI)
      with ¬ sig_red_zero (≼t) G r show False ..
    qed
    from assms(1) G_sub r  dgrad_sig_set d s_red have "s  dgrad_sig_set d"
      by (rule dgrad_sig_set_closed_sig_red_rtrancl)
    from s_red have "lt s = lt r" and "lc s = lc r"
      by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
    hence "lt ?h = lt s" and "lc ?h = lc s" and "s  0"
      using lc r  0 by (auto simp: lt ?h = lt r lc ?h = lc r simp del: lc_monom_mult)
    from s_irred have "¬ is_sig_red (≺t) (=) G s" by (simp add: is_sig_red_top_tail_cases)
    from is_sig_GB_upt d G v have "is_sig_GB_upt d G (lt s)" by (simp only: lt s = lt r lt r = v)
    have "punit.lt (rep_list ?h) = punit.lt (rep_list s)"
      by (rule sig_regular_top_reduced_lt_unique, fact+)
    hence eq2: "lp r - lp g + punit.lt (rep_list g) = punit.lt (rep_list s)"
      using lc r / lc g  0 rep_list g  0 by (simp add: rep_list_monom_mult punit.lt_monom_mult)
    have "punit.lc (rep_list ?h) = punit.lc (rep_list s)"
      by (rule sig_regular_top_reduced_lc_unique, fact+)
    hence eq3: "lc r / lc g = punit.lc (rep_list s) / punit.lc (rep_list g)"
      using punit.lc (rep_list g)  0 by (simp add: rep_list_monom_mult field_simps)
    have "sig_red_single (=) (=) s (s - ?h) g (lp r - lp g)"
      by (rule sig_red_singleI, auto simp: eq1 eq2 eq3 punit.lc_def[symmetric] lt s = lt r
            rep_list g  0 rep_list s  0 intro!: punit.lt_in_keys)
    with g  G have "sig_red (=) (=) G s (s - ?h)" unfolding sig_red_def by blast
    hence "sig_red (≼t) (≼) G s (s - ?h)" by (auto dest: sig_red_sing_regI sig_red_top_tailI)
    with s_red' have r_red: "(sig_red (≼t) (≼) G)** r (s - ?h)" ..
    show ?thesis
    proof (cases "s - ?h = 0")
      case True
      hence "rep_list (s - ?h) = 0" by (simp only: rep_list_zero)
      with r_red show ?thesis by (rule sig_red_zeroI)
    next
      case False
      note is_sig_GB_upt d G (lt s)
      moreover from s  dgrad_sig_set d ?h  dgrad_sig_set d have "s - ?h  dgrad_sig_set d"
        by (rule dgrad_sig_set_closed_minus)
      moreover from False lt ?h = lt s lc ?h = lc s have "lt (s - ?h) t lt s" by (rule lt_minus_lessI)
      ultimately have "sig_red_zero (≼t) G (s - ?h)" by (rule is_sig_GB_uptD3)
      then obtain s' where "(sig_red (≼t) (≼) G)** (s - ?h) s'" and "rep_list s' = 0"
        by (rule sig_red_zeroE)
      from r_red this(1) have "(sig_red (≼t) (≼) G)** r s'" by simp
      thus ?thesis using rep_list s' = 0 by (rule sig_red_zeroI)
    qed
  qed
  with ¬ sig_red_zero (≼t) G r show False ..
qed

corollary is_RB_upt_is_syz_sigD:
  assumes "dickson_grading d" and "is_RB_upt d rword G u"
    and "is_syz_sig d u" and "p  dgrad_sig_set d" and "lt p = u"
  shows "sig_red_zero (≺t) G p"
proof -
  note assms(1)
  moreover from assms(1, 2) have "is_sig_GB_upt d G u" by (rule is_RB_upt_is_sig_GB_upt)
  ultimately show ?thesis using assms(3, 4, 5) by (rule syzygy_crit)
qed

subsubsection ‹S-Pairs›

definition spair :: "('t 0 'b)  ('t 0 'b)  ('t 0 'b)"
  where "spair p q = (let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
                        (monom_mult (1 / punit.lc (rep_list p)) (l - t1) p) -
                        (monom_mult (1 / punit.lc (rep_list q)) (l - t2) q))"

definition is_regular_spair :: "('t 0 'b)  ('t 0 'b)  bool"
  where "is_regular_spair p q 
                    (rep_list p  0  rep_list q  0 
                      (let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
                        (l - t1)  lt p  (l - t2)  lt q))"

lemma rep_list_spair: "rep_list (spair p q) = punit.spoly (rep_list p) (rep_list q)"
  by (simp add: spair_def punit.spoly_def Let_def rep_list_minus rep_list_monom_mult punit.lc_def)

lemma spair_comm: "spair p q = - spair q p"
  by (simp add: spair_def Let_def lcs_comm)

lemma dgrad_sig_set_closed_spair:
  assumes "dickson_grading d" and "p  dgrad_sig_set d" and "q  dgrad_sig_set d"
  shows "spair p q  dgrad_sig_set d"
proof -
  define t1 where "t1 = punit.lt (rep_list p)"
  define t2 where "t2 = punit.lt (rep_list q)"
  let ?l = "lcs t1 t2"
  have "d t1  dgrad_max d"
  proof (cases "rep_list p = 0")
    case True
    show ?thesis by (simp add: t1_def True dgrad_max_0)
  next
    case False
    from assms(2) have "p  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
    with assms(1) have "rep_list p  punit_dgrad_max_set d" by (rule dgrad_max_2)
    thus ?thesis unfolding t1_def using False by (rule punit.dgrad_p_setD_lp[simplified])
  qed
  moreover have "d t2  dgrad_max d"
  proof (cases "rep_list q = 0")
    case True
    show ?thesis by (simp add: t2_def True dgrad_max_0)
  next
    case False
    from assms(3) have "q  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
    with assms(1) have "rep_list q  punit_dgrad_max_set d" by (rule dgrad_max_2)
    thus ?thesis unfolding t2_def using False by (rule punit.dgrad_p_setD_lp[simplified])
  qed
  ultimately have "ord_class.max (d t1) (d t2)  dgrad_max d" by simp
  moreover from assms(1) have "d ?l  ord_class.max (d t1) (d t2)" by (rule dickson_grading_lcs)
  ultimately have *: "d ?l  dgrad_max d" by auto
  thm dickson_grading_minus
  show ?thesis
  proof (simp add: spair_def Let_def t1_def[symmetric] t2_def[symmetric],
      intro dgrad_sig_set_closed_minus dgrad_sig_set_closed_monom_mult[OF assms(1)])
    from assms(1) adds_lcs have "d (?l - t1)  d ?l" by (rule dickson_grading_minus)
    thus "d (?l - t1)  dgrad_max d" using * by (rule le_trans)
  next
    from assms(1) adds_lcs_2 have "d (?l - t2)  d ?l" by (rule dickson_grading_minus)
    thus "d (?l - t2)  dgrad_max d" using * by (rule le_trans)
  qed fact+
qed

lemma lt_spair:
  assumes "rep_list p  0" and "punit.lt (rep_list p)  lt q t punit.lt (rep_list q)  lt p"
  shows "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p"
proof -
  define l where "l = lcs (punit.lt (rep_list p)) (punit.lt (rep_list q))"
  have 1: "punit.lt (rep_list p) adds l" and 2: "punit.lt (rep_list q) adds l"
    unfolding l_def by (rule adds_lcs, rule adds_lcs_2)
  have eq1: "spair p q = (monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) p) +
                         (monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) q)"
    by (simp add: spair_def Let_def l_def monom_mult_uminus_left)
  from assms(1) have "punit.lc (rep_list p)  0" by (rule punit.lc_not_0)
  hence "1 / punit.lc (rep_list p)  0" by simp
  moreover from assms(1) have "p  0" by (auto simp: rep_list_zero)
  ultimately have eq2: "lt (monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) p) =
                        (l - punit.lt (rep_list p))  lt p"
    by (rule lt_monom_mult)
  have "lt (monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) q) t
                        (l - punit.lt (rep_list q))  lt q"
    by (fact lt_monom_mult_le)
  also from assms(2) have "... t (l - punit.lt (rep_list p))  lt p"
    by (simp add: term_is_le_rel_minus_minus[OF _ 2 1])
  finally show ?thesis unfolding eq2[symmetric] eq1 l_def[symmetric] by (rule lt_plus_eqI_2)
qed

lemma lt_spair':
  assumes "rep_list p  0" and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)" and "b  lt q t a  lt p"
  shows "lt (spair p q) = (a - gcs a b)  lt p"
proof -
  from assms(3) have "punit.lt (rep_list p)  (b  lt q) t punit.lt (rep_list p)  (a  lt p)"
    by (fact splus_mono_strict)
  hence "(b + punit.lt (rep_list p))  lt q t (b + punit.lt (rep_list q))  lt p"
    by (simp only: splus_assoc[symmetric] add.commute assms(2))
  hence "punit.lt (rep_list p)  lt q t punit.lt (rep_list q)  lt p"
    by (simp only: splus_assoc ord_term_strict_canc)
  with assms(1)
  have "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p"
    by (rule lt_spair)
  with assms(2) show ?thesis by (simp add: lcs_minus_1)
qed

lemma lt_rep_list_spair:
  assumes "rep_list p  0" and "rep_list q  0" and "rep_list (spair p q)  0"
    and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)"
  shows "punit.lt (rep_list (spair p q))  (a - gcs a b) + punit.lt (rep_list p)"
proof -
  from assms(1) have 1: "punit.lc (rep_list p)  0" by (rule punit.lc_not_0)
  from assms(2) have 2: "punit.lc (rep_list q)  0" by (rule punit.lc_not_0)
  define l where "l = lcs (punit.lt (rep_list p)) (punit.lt (rep_list q))"
  have eq: "rep_list (spair p q) = (punit.monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) (rep_list p)) +
                               (punit.monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) (rep_list q))"
    (is "_ = ?a + ?b")
    by (simp add: spair_def Let_def rep_list_minus rep_list_monom_mult punit.monom_mult_uminus_left l_def)
  have "?a + ?b  0" unfolding eq[symmetric] by (fact assms(3))
  moreover from 1 2 assms(1, 2) have "punit.lt ?b = punit.lt ?a"
    by (simp add: lp_monom_mult l_def minus_plus adds_lcs adds_lcs_2)
  moreover have "punit.lc ?b = - punit.lc ?a" by (simp add: 1 2)
  ultimately have "punit.lt (rep_list (spair p q))  punit.lt ?a" unfolding eq by (rule punit.lt_plus_lessI)
  also from 1 assms(1) have "... = (l - punit.lt (rep_list p)) + punit.lt (rep_list p)" by (simp add: lp_monom_mult)
  also have "... = l" by (simp add: l_def minus_plus adds_lcs)
  also have "... = (a + punit.lt (rep_list p)) - gcs a b" unfolding l_def using assms(4) by (rule lcs_alt_1)
  also have "... = (a - gcs a b) + punit.lt (rep_list p)" by (simp add: minus_plus gcs_adds)
  finally show ?thesis .
qed

lemma is_regular_spair_sym: "is_regular_spair p q  is_regular_spair q p"
  by (auto simp: is_regular_spair_def Let_def lcs_comm)

lemma is_regular_spairI:
  assumes "rep_list p  0" and "rep_list q  0"
  and "punit.lt (rep_list q)  lt p  punit.lt (rep_list p)  lt q"
  shows "is_regular_spair p q"
proof -
  have *: "(lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p 
           (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q))  lt q"
    (is "?l  ?r")
  proof
    assume "?l = ?r"
    hence "punit.lt (rep_list q)  lt p = punit.lt (rep_list p)  lt q"
      by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
    with assms(3) show False ..
  qed
  with assms(1, 2) show ?thesis by (simp add: is_regular_spair_def)
qed

lemma is_regular_spairI':
  assumes "rep_list p  0" and "rep_list q  0"
  and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)" and "a  lt p  b  lt q"
  shows "is_regular_spair p q"
proof -
  have "punit.lt (rep_list q)  lt p  punit.lt (rep_list p)  lt q"
  proof
    assume "punit.lt (rep_list q)  lt p = punit.lt (rep_list p)  lt q"
    hence "a  (punit.lt (rep_list q)  lt p) = a  (punit.lt (rep_list p)  lt q)" by (simp only:)
    hence "(a + punit.lt (rep_list q))  lt p = (b + punit.lt (rep_list q))  lt q"
      by (simp add: splus_assoc[symmetric] assms(3))
    hence "punit.lt (rep_list q)  (a  lt p) = punit.lt (rep_list q)  (b  lt q)"
      by (simp only: add.commute[of _ "punit.lt (rep_list q)"] splus_assoc)
    hence "a  lt p = b  lt q" by (simp only: splus_left_canc)
    with assms(4) show False ..
  qed
  with assms(1, 2) show ?thesis by (rule is_regular_spairI)
qed

lemma is_regular_spairD1: "is_regular_spair p q  rep_list p  0"
  by (simp add: is_regular_spair_def)

lemma is_regular_spairD2: "is_regular_spair p q  rep_list q  0"
  by (simp add: is_regular_spair_def)

lemma is_regular_spairD3:
  fixes p q
  defines "t1  punit.lt (rep_list p)"
  defines "t2  punit.lt (rep_list q)"
  assumes "is_regular_spair p q"
  shows "t2  lt p  t1  lt q" (is ?thesis1)
    and "lt (monom_mult (1 / punit.lc (rep_list p)) (lcs t1 t2 - t1) p) 
         lt (monom_mult (1 / punit.lc (rep_list q)) (lcs t1 t2 - t2) q)"  (is "?l  ?r")
proof -
  from assms(3) have "rep_list p  0" by (rule is_regular_spairD1)
  hence "punit.lc (rep_list p)  0" and "p  0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
  from assms(3) have "rep_list q  0" by (rule is_regular_spairD2)
  hence "punit.lc (rep_list q)  0" and "q  0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)

  have "?l = (lcs t1 t2 - t1)  lt p"
    using punit.lc (rep_list p)  0 p  0 by (simp add: lt_monom_mult)
  also from assms(3) have *: "...  (lcs t1 t2 - t2)  lt q"
    by (simp add: is_regular_spair_def t1_def t2_def Let_def)
  also have "(lcs t1 t2 - t2)  lt q = ?r"
    using punit.lc (rep_list q)  0 q  0 by (simp add: lt_monom_mult)
  finally show "?l  ?r" .

  show ?thesis1
  proof
    assume "t2  lt p = t1  lt q"
    hence "(lcs t1 t2 - t1)  lt p = (lcs t1 t2 - t2)  lt q"
      by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
    with * show False ..
  qed
qed

lemma is_regular_spair_nonzero: "is_regular_spair p q  spair p q  0"
  by (auto simp: spair_def Let_def dest: is_regular_spairD3)

lemma is_regular_spair_lt:
  assumes "is_regular_spair p q"
  shows "lt (spair p q) = ord_term_lin.max
              ((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p)
              ((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q))  lt q)"
proof -
  let ?t1 = "punit.lt (rep_list p)"
  let ?t2 = "punit.lt (rep_list q)"
  let ?l = "lcs ?t1 ?t2"
  show ?thesis
  proof (rule ord_term_lin.linorder_cases)
    assume a: "?t2  lt p t ?t1  lt q"
    hence "(?l - ?t1)  lt p t (?l - ?t2)  lt q"
      by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
    hence le: "(?l - ?t1)  lt p t (?l - ?t2)  lt q" by (rule ord_term_lin.less_imp_le)
    from assms have "rep_list q  0" by (rule is_regular_spairD2)
    have "lt (spair p q) = lt (spair q p)" by (simp add: spair_comm[of p])
    also from rep_list q  0 a have "... = (lcs ?t2 ?t1 - ?t2)  lt q" by (rule lt_spair)
    also have "... = (?l - ?t2)  lt q" by (simp only: lcs_comm)
    finally show ?thesis using le by (simp add: ord_term_lin.max_def)
  next
    assume a: "?t1  lt q t ?t2  lt p"
    hence "(?l - ?t2)  lt q t (?l - ?t1)  lt p"
      by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
    hence le: "¬ ((?l - ?t1)  lt p t (?l - ?t2)  lt q)" by simp
    from assms have "rep_list p  0" by (rule is_regular_spairD1)
    hence "lt (spair p q) = (lcs ?t1 ?t2 - ?t1)  lt p" using a by (rule lt_spair)
    thus ?thesis using le by (simp add: ord_term_lin.max_def)
  next
    from assms have "?t2  lt p  ?t1  lt q" by (rule is_regular_spairD3)
    moreover assume "?t2  lt p = ?t1  lt q"
    ultimately show ?thesis ..
  qed
qed

lemma is_regular_spair_lt_ge_1:
  assumes "is_regular_spair p q"
  shows "lt p t lt (spair p q)"
proof -
  have "lt p = 0  lt p" by (simp only: term_simps)
  also from zero_min have "... t (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p"
    by (rule splus_mono_left)
  also have "... t ord_term_lin.max
              ((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p)
              ((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q))  lt q)"
    by (rule ord_term_lin.max.cobounded1)
  also from assms have "... = lt (spair p q)" by (simp only: is_regular_spair_lt)
  finally show ?thesis .
qed

corollary is_regular_spair_lt_ge_2:
  assumes "is_regular_spair p q"
  shows "lt q t lt (spair p q)"
proof -
  from assms have "is_regular_spair q p" by (rule is_regular_spair_sym)
  hence "lt q t lt (spair q p)" by (rule is_regular_spair_lt_ge_1)
  also have "... = lt (spair p q)" by (simp add: spair_comm[of q])
  finally show ?thesis .
qed

lemma is_regular_spair_component_lt_cases:
  assumes "is_regular_spair p q"
  shows "component_of_term (lt (spair p q)) = component_of_term (lt p) 
          component_of_term (lt (spair p q)) = component_of_term (lt q)"
proof (rule ord_term_lin.linorder_cases)
  from assms have "rep_list q  0" by (rule is_regular_spairD2)
  moreover assume "punit.lt (rep_list q)  lt p t punit.lt (rep_list p)  lt q"
  ultimately have "lt (spair q p) = (lcs (punit.lt (rep_list q)) (punit.lt (rep_list p)) - punit.lt (rep_list q))  lt q"
    by (rule lt_spair)
  thus ?thesis by (simp add: spair_comm[of p] term_simps)
next
  from assms have "rep_list p  0" by (rule is_regular_spairD1)
  moreover assume "punit.lt (rep_list p)  lt q t punit.lt (rep_list q)  lt p"
  ultimately have "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p))  lt p"
    by (rule lt_spair)
  thus ?thesis by (simp add: term_simps)
next
  from assms have "punit.lt (rep_list q)  lt p  punit.lt (rep_list p)  lt q"
    by (rule is_regular_spairD3)
  moreover assume "punit.lt (rep_list q)  lt p = punit.lt (rep_list p)  lt q"
  ultimately show ?thesis ..
qed

lemma lemma_9:
  assumes "dickson_grading d" and "is_rewrite_ord rword" and "is_RB_upt d rword G u"
    and "inj_on lt G" and "¬ is_syz_sig d u" and "is_canon_rewriter rword G u g1" and "h  G"
    and "is_sig_red (≺t) (=) {h} (monom_mult 1 (pp_of_term u - lp g1) g1)"
    and "d (pp_of_term u)  dgrad_max d"
  shows "lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h)) - punit.lt (rep_list g1) =
            pp_of_term u - lp g1" (is ?thesis1)
    and "lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h)) - punit.lt (rep_list h) =
            ((pp_of_term u - lp g1) + punit.lt (rep_list g1)) - punit.lt (rep_list h)" (is ?thesis2)
    and "is_regular_spair g1 h" (is ?thesis3)
    and "lt (spair g1 h) = u" (is ?thesis4)
proof -
  from assms(8) have "rep_list (monom_mult 1 (pp_of_term u - lp g1) g1)  0"
    using is_sig_red_top_addsE by fastforce
  hence "rep_list g1  0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
  hence "g1  0" by (auto simp: rep_list_zero)
  from assms(6) have "g1  G" and "lt g1 addst u" by (rule is_canon_rewriterD)+
  from assms(3) have "G  dgrad_sig_set d" by (rule is_RB_uptD1)
  with g1  G have "g1  dgrad_sig_set d" ..
  hence "component_of_term (lt g1) < length fs" using g1  0 by (rule dgrad_sig_setD_lt)
  with lt g1 addst u have "component_of_term u < length fs" by (simp add: adds_term_def)

  from lt g1 addst u obtain a where u: "u = a  lt g1" by (rule adds_termE)
  hence a: "a = pp_of_term u - lp g1" by (simp add: term_simps)
  from assms(8) have "is_sig_red (≺t) (=) {h} (monom_mult 1 a g1)" by (simp only: a)
  hence "rep_list h  0" and "rep_list (monom_mult 1 a g1)  0" and
      2: "punit.lt (rep_list h) adds punit.lt (rep_list (monom_mult 1 a g1))" and
      3: "punit.lt (rep_list (monom_mult 1 a g1))  lt h t punit.lt (rep_list h)  lt (monom_mult 1 a g1)"
    by (auto elim: is_sig_red_top_addsE)
  from this(2) have "rep_list g1  0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
  hence "g1  0" by (auto simp: rep_list_zero)
  from rep_list h  0 have "h  0" by (auto simp: rep_list_zero)
  from 2 rep_list g1  0 have "punit.lt (rep_list h) adds a + punit.lt (rep_list g1)"
    by (simp add: rep_list_monom_mult punit.lt_monom_mult)
  then obtain b where eq1: "a + punit.lt (rep_list g1) = b + punit.lt (rep_list h)"
    by (auto elim: addsE simp: add.commute)
  hence b: "b = ((pp_of_term u - lp g1) + punit.lt (rep_list g1)) - punit.lt (rep_list h)"
    by (simp add: a)

  define g where "g = gcs a b"
  have "g = 0"
  proof (rule ccontr)
    assume "g  0"
    have "g adds a" unfolding g_def by (fact gcs_adds)
    also have "... addsp u" unfolding u by (fact adds_pp_triv)
    finally obtain v where u2: "u = g  v" by (rule adds_ppE)
    hence v: "v = u  g" by (simp add: term_simps)
    from u2 have "v addst u" by (rule adds_termI)
    hence "v t u" by (rule ord_adds_term)
    moreover have "v  u"
    proof
      assume "v = u"
      hence "g  v = 0  v" by (simp add: u2 term_simps)
      hence "g = 0" by (simp only: splus_right_canc)
      with g  0 show False ..
    qed
    ultimately have "v t u" by simp
    note assms(3) v t u
    moreover have "d (pp_of_term v)  dgrad_max d"
    proof (rule le_trans)
      from assms(1) show "d (pp_of_term v)  d (pp_of_term u)"
        by (simp add: u2 term_simps dickson_gradingD1)
    qed fact
    moreover from component_of_term u < length fs have "component_of_term v < length fs"
      by (simp only: v term_simps)
    ultimately have "is_RB_in d rword G v" by (rule is_RB_uptD2)
    thus False
    proof (rule is_RB_inE)
      assume "is_syz_sig d v"
      with assms(1) have "is_syz_sig d u" using v addst u assms(9) by (rule is_syz_sig_adds)
      with assms(5) show False ..
    next
      fix g2
      assume *: "¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g2) g2)"
      assume "is_canon_rewriter rword G v g2"
      hence "g2  G" and "g2  0" and "lt g2 addst v" by (rule is_canon_rewriterD)+
      assume "¬ is_syz_sig d v"
      note assms(2) is_canon_rewriter rword G v g2 assms(6)
      moreover from lt g2 addst v v addst u have "lt g2 addst u" by (rule adds_term_trans)
      moreover from g adds a have "lt g1 addst v" by (simp add: v u minus_splus[symmetric] adds_termI)
      ultimately have "lt g2 = lt g1" by (rule is_rewrite_ord_canon_rewriterD1)
      with assms(4) have "g2 = g1" using g2  G g1  G by (rule inj_onD)
      have "pp_of_term v - lp g1 = a - g" by (simp add: u v term_simps diff_diff_add)
  
      have "is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g2) g2)"
        unfolding g2 = g1 pp_of_term v - lp g1 = a - g using assms(7) rep_list h  0
      proof (rule is_sig_red_top_addsI)
        from rep_list g1  0 show "rep_list (monom_mult 1 (a - g) g1)  0"
          by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
      next
        have eq3: "(a - g) + punit.lt (rep_list g1) = lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h))"
          by (simp add: g_def lcs_minus_1[OF eq1, symmetric] adds_minus adds_lcs)
        from rep_list g1  0
        show "punit.lt (rep_list h) adds punit.lt (rep_list (monom_mult 1 (a - g) g1))"
          by (simp add: rep_list_monom_mult punit.lt_monom_mult eq3 adds_lcs_2)
      next
        from 3 rep_list g1  0 g1  0
        show "punit.lt (rep_list (monom_mult 1 (a - g) g1))  lt h t
              punit.lt (rep_list h)  lt (monom_mult 1 (a - g) g1)"
          by (auto simp: rep_list_monom_mult punit.lt_monom_mult lt_monom_mult splus_assoc splus_left_commute
                dest!: ord_term_strict_canc intro: splus_mono_strict)
      next
        show "ord_term_lin.is_le_rel (≺t)" by (fact ord_term_lin.is_le_relI)
      qed
      with * show False ..
    qed
  qed
  thus ?thesis1 and ?thesis2 by (simp_all add: a b lcs_minus_1[OF eq1] lcs_minus_2[OF eq1] g_def)
  hence eq3: "spair g1 h = monom_mult (1 / punit.lc (rep_list g1)) a g1 -
                            monom_mult (1 / punit.lc (rep_list h)) b h"
    by (simp add: spair_def Let_def a b)

  from 3 rep_list g1  0 g1  0 have "b  lt h t a  lt g1"
    by (auto simp: rep_list_monom_mult punit.lt_monom_mult lt_monom_mult eq1 splus_assoc
        splus_left_commute[of b] dest!: ord_term_strict_canc)
  hence "a  lt g1  b  lt h" by simp
  with rep_list g1  0 rep_list h  0 eq1 show ?thesis3
    by (rule is_regular_spairI')

  have "lt (monom_mult (1 / punit.lc (rep_list h)) b h) = b  lt h"
  proof (rule lt_monom_mult)
    from rep_list h  0 show "1 / punit.lc (rep_list h)  0" by (simp add: punit.lc_eq_zero_iff)
  qed fact
  also have "... t a  lt g1" by fact
  also have "... = lt (monom_mult (1 / punit.lc (rep_list g1)) a g1)"
  proof (rule HOL.sym, rule lt_monom_mult)
    from rep_list g1  0 show "1 / punit.lc (rep_list g1)  0" by (simp add: punit.lc_eq_zero_iff)
  qed fact
  finally have "lt (spair g1 h) = lt (monom_mult (1 / punit.lc (rep_list g1)) a g1)"
    unfolding eq3 by (rule lt_minus_eqI_2)
  also have "... = a  lt g1" by (rule HOL.sym, fact)
  finally show ?thesis4 by (simp only: u)
qed

lemma is_RB_upt_finite:
  assumes "dickson_grading d" and "is_rewrite_ord rword" and "G  dgrad_sig_set d" and "inj_on lt G"
    and "finite G"
    and "g1 g2. g1  G  g2  G  is_regular_spair g1 g2  lt (spair g1 g2) t u 
              is_RB_in d rword G (lt (spair g1 g2))"
    and "i. i < length fs  term_of_pair (0, i) t u  is_RB_in d rword G (term_of_pair (0, i))"
  shows "is_RB_upt d rword G u"
proof (rule ccontr)
  let ?Q = "{v. v t u  d (pp_of_term v)  dgrad_max d  component_of_term v < length fs  ¬ is_RB_in d rword G v}"
  have Q_sub: "pp_of_term ` ?Q  dgrad_set d (dgrad_max d)" by blast
  from assms(3) have "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  assume "¬ is_RB_upt d rword G u"
  with assms(3) obtain v' where "v'  ?Q" unfolding is_RB_upt_def by blast
  with assms(1) obtain v where "v  ?Q" and min: "y. y t v  y  ?Q" using Q_sub
    by (rule ord_term_minimum_dgrad_set, blast)
  from v  ?Q have "v t u" and "d (pp_of_term v)  dgrad_max d" and "component_of_term v < length fs"
    and "¬ is_RB_in d rword G v" by simp_all
  from this(4)
  have impl: "g. g  G  is_canon_rewriter rword G v g 
                    is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
    and "¬ is_syz_sig d v" by (simp_all add: is_RB_in_def Let_def)

  from assms(3) have "is_RB_upt d rword G v"
  proof (rule is_RB_uptI)
    fix w
    assume dw: "d (pp_of_term w)  dgrad_max d" and cp: "component_of_term w < length fs"
    assume "w t v"
    hence "w  ?Q" by (rule min)
    hence "¬ w t u  is_RB_in d rword G w" by (simp add: dw cp)
    thus "is_RB_in d rword G w"
    proof
      assume "¬ w t u"
      moreover from w t v v t u have "w t u" by (rule ord_term_lin.less_trans)
      ultimately show ?thesis ..
    qed
  qed

  show False
  proof (cases "gG. g  0  lt g addst v")
    case False
    hence x: "g. g  G  lt g addst v  g = 0" by blast
    let ?w = "term_of_pair (0, component_of_term v)"
    have "?w addst v" by (simp add: adds_term_def term_simps)
    hence "?w t v" by (rule ord_adds_term)
    also have "... t u" by fact
    finally have "?w t u" .
    with component_of_term v < length fs have "is_RB_in d rword G ?w" by (rule assms(7))
    thus ?thesis
    proof (rule is_RB_inE)
      assume "is_syz_sig d ?w"
      with assms(1) have "is_syz_sig d v" using ?w addst v d (pp_of_term v)  dgrad_max d
        by (rule is_syz_sig_adds)
      with ¬ is_syz_sig d v show ?thesis ..
    next
      fix g1
      assume "is_canon_rewriter rword G ?w g1"
      hence "g1  0" and "g1  G" and "lt g1 addst ?w" by (rule is_canon_rewriterD)+
      from this(3) have "lt g1 addst v" using ?w addst v by (rule adds_term_trans)
      with g1  G have "g1 = 0" by (rule x)
      with g1  0 show ?thesis ..
    qed
  next
    case True
    then obtain g' where "g'  G" and "g'  0" and "lt g' addst v" by blast
    with assms(2, 5) obtain g1 where crw: "is_canon_rewriter rword G v g1"
      by (rule is_rewrite_ord_finite_canon_rewriterE)
    hence "g1  G" by (rule is_canon_rewriterD1)
    hence "is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term v - lp g1) g1)" using crw by (rule impl)
    then obtain h where "h  G" and "is_sig_red (≺t) (=) {h} (monom_mult 1 (pp_of_term v - lp g1) g1)"
      by (rule is_sig_red_singletonI)
    with assms(1, 2) is_RB_upt d rword G v assms(4) ¬ is_syz_sig d v crw
    have "is_regular_spair g1 h" and eq: "lt (spair g1 h) = v"
      using d (pp_of_term v)  dgrad_max d by (rule lemma_9)+
    from v t u have "lt (spair g1 h) t u" by (simp only: eq)
    with g1  G h  G is_regular_spair g1 h have "is_RB_in d rword G (lt (spair g1 h))"
      by (rule assms(6))
    hence "is_RB_in d rword G v" by (simp only: eq)
    with ¬ is_RB_in d rword G v show ?thesis ..
  qed
qed

text ‹Note that the following lemma actually holds for @{emph ‹all›} regularly reducible power-products
  in @{term "rep_list p"}, not just for the leading power-product.›

lemma lemma_11:
  assumes "dickson_grading d" and "is_rewrite_ord rword" and "is_RB_upt d rword G (lt p)"
    and "p  dgrad_sig_set d" and "is_sig_red (≺t) (=) G p"
  obtains u g where "u t lt p" and "d (pp_of_term u)  dgrad_max d" and "component_of_term u < length fs"
    and "¬ is_syz_sig d u" and "is_canon_rewriter rword G u g"
    and "u = (punit.lt (rep_list p) - punit.lt (rep_list g))  lt g" and "is_sig_red (≺t) (=) {g} p"
proof -
  from assms(3) have G_sub: "G  dgrad_sig_set d" by (rule is_RB_uptD1)
  from assms(5) have "rep_list p  0" using is_sig_red_addsE by fastforce
  hence "p  0" by (auto simp: rep_list_zero)
  let ?lc = "punit.lc (rep_list p)"
  let ?lp = "punit.lt (rep_list p)"
  from rep_list p  0 have "?lc  0" by (rule punit.lc_not_0)
  from assms(4) have "p  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  from assms(4) have "d (lp p)  dgrad_max d" by (rule dgrad_sig_setD_lp)
  from assms(4) p  0 have "component_of_term (lt p) < length fs" by (rule dgrad_sig_setD_lt)
  from assms(1) p  dgrad_max_set d have "rep_list p  punit_dgrad_max_set d" by (rule dgrad_max_2)
  hence "d ?lp  dgrad_max d" using rep_list p  0 by (rule punit.dgrad_p_setD_lp[simplified])

  from assms(5) obtain g0 where "g0  G" and "is_sig_red (≺t) (=) {g0} p"
    by (rule is_sig_red_singletonI)
  from g0  G G_sub have "g0  dgrad_sig_set d" ..
  let ?g0 = "monom_mult (?lc / punit.lc (rep_list g0)) (?lp - punit.lt (rep_list g0)) g0"

  define M where "M = {monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g |
                          g. g  dgrad_sig_set d  is_sig_red (≺t) (=) {g} p}"
  from g0  dgrad_sig_set d is_sig_red (≺t) (=) {g0} p have "?g0  M" by (auto simp: M_def)
  have "0  rep_list ` M"
  proof
    assume "0  rep_list ` M"
    then obtain g where 1: "is_sig_red (≺t) (=) {g} p"
      and 2: "rep_list (monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g) = 0"
      unfolding M_def by fastforce
    from 1 have "rep_list g  0" using is_sig_red_addsE by fastforce
    moreover from this have "punit.lc (rep_list g)  0" by (rule punit.lc_not_0)
    ultimately have "rep_list (monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g)  0"
      using ?lc  0 by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
    thus False using 2 ..
  qed
  with rep_list_zero have "0  M" by auto
  have "M  dgrad_sig_set d"
  proof
    fix m
    assume "m  M"
    then obtain g where "g  dgrad_sig_set d" and 1: "is_sig_red (≺t) (=) {g} p"
      and m: "m = monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g"
      unfolding M_def by fastforce
    from 1 have "punit.lt (rep_list g) adds ?lp" using is_sig_red_top_addsE by fastforce
    note assms(1)
    thm dickson_grading_minus
    moreover have "d (?lp - punit.lt (rep_list g))  dgrad_max d"
      by (rule le_trans, rule dickson_grading_minus, fact+)
    ultimately show "m  dgrad_sig_set d" unfolding m using g  dgrad_sig_set d
      by (rule dgrad_sig_set_closed_monom_mult)
  qed
  hence "M  sig_inv_set" by (simp add: dgrad_sig_set'_def)

  let ?M = "lt ` M"
  note assms(1)
  moreover from ?g0  M have "lt ?g0  ?M" by (rule imageI)
  moreover from M  dgrad_sig_set d have "pp_of_term ` ?M  dgrad_set d (dgrad_max d)"
    by (auto intro!: dgrad_sig_setD_lp)
  ultimately obtain u where "u  ?M" and min: "v. v t u  v  ?M"
    by (rule ord_term_minimum_dgrad_set, blast)
  from u  ?M obtain m where "m  M" and u': "u = lt m" ..
  from this(1) obtain g1 where "g1  dgrad_sig_set d" and 1: "is_sig_red (≺t) (=) {g1} p"
    and m: "m = monom_mult (?lc / punit.lc (rep_list g1)) (?lp - punit.lt (rep_list g1)) g1"
    unfolding M_def by fastforce
  from 1 have adds: "punit.lt (rep_list g1) adds ?lp" and "?lp  lt g1 t punit.lt (rep_list g1)  lt p"
    and "rep_list g1  0" using is_sig_red_top_addsE by fastforce+
  from this(3) have lc_g1: "punit.lc (rep_list g1)  0" by (rule punit.lc_not_0)
  from m  M 0  rep_list ` M have "rep_list m  0" by fastforce
  from m  M 0  M have "m  0" by blast
  hence "lc m  0" by (rule lc_not_0)
  from lc_g1 have eq0: "punit.lc (rep_list m) = ?lc" by (simp add: m rep_list_monom_mult)
  from ?lc  0 rep_list g1  0 adds have eq1: "punit.lt (rep_list m) = ?lp"
    by (simp add: m rep_list_monom_mult punit.lt_monom_mult punit.lc_eq_zero_iff adds_minus)
  from m  M M  dgrad_sig_set d have "m  dgrad_sig_set d" ..

  from rep_list g1  0 have "punit.lc (rep_list g1)  0" and "g1  0"
    by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
  with ?lc  0 have u: "u = (?lp - punit.lt (rep_list g1))  lt g1"
    by (simp add: u' m lt_monom_mult lc_eq_zero_iff)
  hence "punit.lt (rep_list g1)  u = punit.lt (rep_list g1)  ((?lp - punit.lt (rep_list g1))  lt g1)"
    by simp
  also from adds have "... = ?lp  lt g1"
    by (simp only: splus_assoc[symmetric], metis add.commute adds_minus)
  also have "... t punit.lt (rep_list g1)  lt p" by fact
  finally have "u t lt p" by (rule ord_term_strict_canc)

  from u  ?M have "pp_of_term u  pp_of_term ` ?M" by (rule imageI)
  also have "...  dgrad_set d (dgrad_max d)" by fact
  finally have "d (pp_of_term u)  dgrad_max d" by (rule dgrad_setD)

  from u  ?M have "component_of_term u  component_of_term ` ?M" by (rule imageI)
  also from M  sig_inv_set 0  M sig_inv_setD_lt have "...  {0..<length fs}" by fastforce
  finally have "component_of_term u < length fs" by simp

  have "¬ is_syz_sig d u"
  proof
    assume "is_syz_sig d u"
    then obtain s where "s  0" and "lt s = u" and "s  dgrad_sig_set d" and "rep_list s = 0"
      by (rule is_syz_sigE)
    let ?s = "monom_mult (lc m / lc s) 0 s"
    have "rep_list ?s = 0" by (simp add: rep_list_monom_mult rep_list s = 0)
    from s  0 have "lc s  0" by (rule lc_not_0)
    hence "lc m / lc s  0" using lc m  0 by simp
    have "m - ?s  0"
    proof
      assume "m - ?s = 0"
      hence "m = ?s" by simp
      with rep_list ?s = 0 have "rep_list m = 0" by simp
      with rep_list m  0 show False ..
    qed
    moreover from lc m / lc s  0 have "lt ?s = lt m" by (simp add: lt_monom_mult_zero lt s = u u')
    moreover from lc s  0 have "lc ?s = lc m" by simp
    ultimately have "lt (m - ?s) t u" unfolding u' by (rule lt_minus_lessI)
    hence "lt (m - ?s)  ?M" by (rule min)
    hence "m - ?s  M" by blast
    moreover have "m - ?s  M"
    proof -
      have "?s = monom_mult (?lc / lc s) 0 (monom_mult (lc g1 / punit.lc (rep_list g1)) 0 s)"
        by (simp add: m monom_mult_assoc mult.commute)
      define m' where "m' = m - ?s"
      have eq: "rep_list m' = rep_list m" by (simp add: m'_def rep_list_minus rep_list ?s = 0)
      from ?lc  0 have "m' = monom_mult (?lc / punit.lc (rep_list m')) (?lp - punit.lt (rep_list m')) m'"
        by (simp add: eq eq0 eq1)
      also have "...  M" unfolding M_def
      proof (rule, intro exI conjI)
        from s  dgrad_sig_set d have "?s  dgrad_sig_set d"
          by (rule dgrad_sig_set_closed_monom_mult_zero)
        with m  dgrad_sig_set d show "m'  dgrad_sig_set d" unfolding m'_def
          by (rule dgrad_sig_set_closed_minus)
      next
        show "is_sig_red (≺t) (=) {m'} p"
        proof (rule is_sig_red_top_addsI)
          show "m'  {m'}" by simp
        next
          from rep_list m  0 show "rep_list m'  0" by (simp add: eq)
        next
          show "punit.lt (rep_list m') adds punit.lt (rep_list p)" by (simp add: eq eq1)
        next
          have "punit.lt (rep_list p)  lt m' t punit.lt (rep_list p)  u"
            by (rule splus_mono_strict, simp only: m'_def lt (m - ?s) t u)
          also have "... t punit.lt (rep_list m')  lt p"
            unfolding eq eq1 using u t lt p by (rule splus_mono_strict)
          finally show "punit.lt (rep_list p)  lt m' t punit.lt (rep_list m')  lt p" .
        next
          show "ord_term_lin.is_le_rel (≺t)" by simp
        qed fact
      qed (fact refl)
      finally show ?thesis by (simp only: m'_def)
    qed
    ultimately show False ..
  qed

  have "is_RB_in d rword G u" by (rule is_RB_uptD2, fact+)
  thus ?thesis
  proof (rule is_RB_inE)
    assume "is_syz_sig d u"
    with ¬ is_syz_sig d u show ?thesis ..
  next
    fix g
    assume "is_canon_rewriter rword G u g"
    hence "g  G" and "g  0" and adds': "lt g addst u" by (rule is_canon_rewriterD)+
    assume irred: "¬ is_sig_red (≺t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)"

    define b where "b = monom_mult 1 (pp_of_term u - lp g) g"
    note assms(1)
    moreover have "is_sig_GB_upt d G (lt m)" unfolding u'[symmetric]
      by (rule is_sig_GB_upt_le, rule is_RB_upt_is_sig_GB_upt, fact+, rule ord_term_lin.less_imp_le, fact)
    moreover from assms(1) have "b  dgrad_sig_set d" unfolding b_def
    proof (rule dgrad_sig_set_closed_monom_mult)
      from adds' have "lp g adds pp_of_term u" by (simp add: adds_term_def)
      with assms(1) have "d (pp_of_term u - lp g)  d (pp_of_term u)" by (rule dickson_grading_minus)
      thus "d (pp_of_term u - lp g)  dgrad_max d" using d (pp_of_term u)  dgrad_max d
        by (rule le_trans)
    next
      from g  G G_sub show "g  dgrad_sig_set d" ..
    qed
    moreover note m  dgrad_sig_set d
    moreover from g  0 have "lt b = lt m"
      by (simp add: b_def u'[symmetric] lt_monom_mult,
          metis adds' add_diff_cancel_right' adds_termE pp_of_term_splus)
    moreover from g  0 have "b  0" by (simp add: b_def monom_mult_eq_zero_iff)
    moreover note m  0
    moreover from irred have "¬ is_sig_red (≺t) (=) G b" by (simp add: b_def)
    moreover have "¬ is_sig_red (≺t) (=) G m"
    proof
      assume "is_sig_red (≺t) (=) G m"
      then obtain g2 where 1: "g2  G" and 2: "rep_list g2  0"
        and 3: "punit.lt (rep_list g2) adds punit.lt (rep_list m)"
        and 4: "punit.lt (rep_list m)  lt g2 t punit.lt (rep_list g2)  lt m"
        by (rule is_sig_red_top_addsE)
      from 2 have "g2  0" and "punit.lc (rep_list g2)  0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
      with 3 4 have "lt (monom_mult (?lc / punit.lc (rep_list g2)) (?lp - punit.lt (rep_list g2)) g2) t u"
        (is "lt ?g2 t u")
        using ?lc  0 by (simp add: term_is_le_rel_minus u' eq1 lt_monom_mult)
      hence "lt ?g2  ?M" by (rule min)
      hence "?g2  M" by blast
      hence "g2  dgrad_sig_set d  ¬ is_sig_red (≺t) (=) {g2} p" by (simp add: M_def)
      thus False
      proof
        assume "g2  dgrad_sig_set d"
        moreover from g2  G G_sub have "g2  dgrad_sig_set d" ..
        ultimately show ?thesis ..
      next
        assume "¬ is_sig_red (≺t) (=) {g2} p"
        moreover have "is_sig_red (≺t) (=) {g2} p"
        proof (rule is_sig_red_top_addsI)
          show "g2  {g2}" by simp
        next
          from 3 show "punit.lt (rep_list g2) adds punit.lt (rep_list p)" by (simp only: eq1)
        next
          from 4 have "?lp  lt g2 t punit.lt (rep_list g2)  u" by (simp only: eq1 u')
          also from u t lt p have "... t punit.lt (rep_list g2)  lt p" by (rule splus_mono_strict)
          finally show "?lp  lt g2 t punit.lt (rep_list g2)  lt p" .
        next
          show "ord_term_lin.is_le_rel (≺t)" by simp
        qed fact+
        ultimately show ?thesis ..
      qed
    qed
    ultimately have eq2: "punit.lt (rep_list b) = punit.lt (rep_list m)"
      by (rule sig_regular_top_reduced_lt_unique)
    have "rep_list g  0" by (rule is_RB_inD, fact+)
    moreover from adds' have "lp g adds pp_of_term u" and "component_of_term (lt g) = component_of_term u"
      by (simp_all add: adds_term_def)
    ultimately have "u = (?lp - punit.lt (rep_list g))  lt g"
      by (simp add: eq1[symmetric] eq2[symmetric] b_def rep_list_monom_mult punit.lt_monom_mult
          splus_def adds_minus term_simps)
    have "is_sig_red (≺t) (=) {b} p"
    proof (rule is_sig_red_top_addsI)
      show "b  {b}" by simp
    next
      from rep_list g  0 show "rep_list b  0"
        by (simp add: b_def rep_list_monom_mult punit.monom_mult_eq_zero_iff)
    next
      show "punit.lt (rep_list b) adds punit.lt (rep_list p)" by (simp add: eq1 eq2)
    next
      show "punit.lt (rep_list p)  lt b t punit.lt (rep_list b)  lt p"
        by (simp add: eq1 eq2 lt b = lt m u'[symmetric] u t lt p splus_mono_strict)
    next
      show "ord_term_lin.is_le_rel (≺t)" by simp
    qed fact
    hence "is_sig_red (≺t) (=) {g} p" unfolding b_def by (rule is_sig_red_singleton_monom_multD)
    show ?thesis by (rule, fact+)
  qed
qed

subsubsection ‹Termination›

definition term_pp_rel :: "('t  't  bool)  ('t × 'a)  ('t × 'a)  bool"
  where "term_pp_rel r a b  r (snd b  fst a) (snd a  fst b)"

definition canon_term_pp_pair :: "('t × 'a)  bool"
  where "canon_term_pp_pair a  (gcs (pp_of_term (fst a)) (snd a) = 0)"

definition cancel_term_pp_pair :: "('t × 'a)  ('t × 'a)"
  where "cancel_term_pp_pair a = (fst a  (gcs (pp_of_term (fst a)) (snd a)), snd a - (gcs (pp_of_term (fst a)) (snd a)))"

lemma term_pp_rel_refl: "reflp r  term_pp_rel r a a"
  by (simp add: term_pp_rel_def reflp_def)

lemma term_pp_rel_irrefl: "irreflp r  ¬ term_pp_rel r a a"
  by (simp add: term_pp_rel_def irreflp_def)

lemma term_pp_rel_sym: "symp r  term_pp_rel r a b  term_pp_rel r b a"
  by (auto simp: term_pp_rel_def symp_def)

lemma term_pp_rel_trans:
  assumes "ord_term_lin.is_le_rel r" and "term_pp_rel r a b" and "term_pp_rel r b c"
  shows "term_pp_rel r a c"
proof -
  from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
  from assms(2) have 1: "r (snd b  fst a) (snd a  fst b)" by (simp only: term_pp_rel_def)
  from assms(3) have 2: "r (snd c  fst b) (snd b  fst c)" by (simp only: term_pp_rel_def)
  have "snd b  (snd c  fst a) = snd c  (snd b  fst a)" by (rule splus_left_commute)
  also from assms(1) 1 have "r ... (snd a  (snd c  fst b))"
    by (simp add: splus_left_commute[of "snd a"] term_is_le_rel_canc_left)
  also from assms(1) 2 have "r ... (snd b  (snd a  fst c))"
    by (simp add: splus_left_commute[of "snd b"] term_is_le_rel_canc_left)
  finally(transpD[OF transp r]) show ?thesis using assms(1)
    by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed

lemma term_pp_rel_trans_eq_left:
  assumes "ord_term_lin.is_le_rel r" and "term_pp_rel (=) a b" and "term_pp_rel r b c"
  shows "term_pp_rel r a c"
proof -
  from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
  from assms(2) have 1: "snd b  fst a = snd a  fst b" by (simp only: term_pp_rel_def)
  from assms(3) have 2: "r (snd c  fst b) (snd b  fst c)" by (simp only: term_pp_rel_def)
  have "snd b  (snd c  fst a) = snd c  (snd b  fst a)" by (rule splus_left_commute)
  also from assms(1) 1 have "... = (snd a  (snd c  fst b))"
    by (simp add: splus_left_commute[of "snd a"])
  finally have eq: "snd b  (snd c  fst a) = snd a  (snd c  fst b)" .
  from assms(1) 2 have "r (snd b  (snd c  fst a)) (snd b  (snd a  fst c))"
    unfolding eq by (simp add: splus_left_commute[of "snd b"] term_is_le_rel_canc_left)
  thus ?thesis using assms(1) by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed

lemma term_pp_rel_trans_eq_right:
  assumes "ord_term_lin.is_le_rel r" and "term_pp_rel r a b" and "term_pp_rel (=) b c"
  shows "term_pp_rel r a c"
proof -
  from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
  from assms(2) have 1: "r (snd b  fst a) (snd a  fst b)" by (simp only: term_pp_rel_def)
  from assms(3) have 2: "snd c  fst b = snd b  fst c" by (simp only: term_pp_rel_def)
  have "snd b  (snd a  fst c) = snd a  (snd b  fst c)" by (rule splus_left_commute)
  also from assms(1) 2 have "... = (snd a  (snd c  fst b))"
    by (simp add: splus_left_commute[of "snd a"])
  finally have eq: "snd b  (snd a  fst c) = snd a  (snd c  fst b)" .
  from assms(1) 1 have "r (snd b  (snd c  fst a)) (snd b  (snd a  fst c))"
    unfolding eq by (simp add: splus_left_commute[of _ "snd c"] term_is_le_rel_canc_left)
  thus ?thesis using assms(1) by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed

lemma canon_term_pp_cancel: "canon_term_pp_pair (cancel_term_pp_pair a)"
  by (simp add: cancel_term_pp_pair_def canon_term_pp_pair_def gcs_minus_gcs term_simps)

lemma term_pp_rel_cancel:
  assumes "reflp r"
  shows "term_pp_rel r a (cancel_term_pp_pair a)"
proof -
  obtain u s where a: "a = (u, s)" by (rule prod.exhaust)
  show ?thesis
  proof (simp add: a cancel_term_pp_pair_def)
    let ?g = "gcs (pp_of_term u) s"
    have "?g adds s" by (fact gcs_adds_2)
    hence "(s - ?g)  (u  0) = s  u  (?g + 0)" using zero_adds_pp
      by (rule minus_splus_sminus)
    also have "... = s  (u  ?g)"
      by (metis add.left_neutral add.right_neutral adds_pp_def diff_zero gcs_adds_2 gcs_comm
          minus_splus_sminus zero_adds)
    finally have "r ((s - ?g)  u) (s  (u  ?g))" using assms by (simp add: term_simps reflp_def)
    thus "term_pp_rel r (u, s) (u  ?g, s - ?g)" by (simp add: a term_pp_rel_def)
  qed
qed

lemma canon_term_pp_rel_id:
  assumes "term_pp_rel (=) a b" and "canon_term_pp_pair a" and "canon_term_pp_pair b"
  shows "a = b"
proof -
  obtain u s where a: "a = (u, s)" by (rule prod.exhaust)
  obtain v t where b: "b = (v, t)" by (rule prod.exhaust)
  from assms(1) have "t  u = s  v" by (simp add: term_pp_rel_def a b)
  hence 1: "t + pp_of_term u = s + pp_of_term v" by (metis pp_of_term_splus)
  from assms(2) have 2: "gcs (pp_of_term u) s = 0" by (simp add: canon_term_pp_pair_def a)
  from assms(3) have 3: "gcs (pp_of_term v) t = 0" by (simp add: canon_term_pp_pair_def b)
  have "t = t + gcs (pp_of_term u) s" by (simp add: 2)
  also have "... = gcs (t + pp_of_term u) (t + s)" by (simp only: gcs_plus_left)
  also have "... = gcs (s + pp_of_term v) (s + t)" by (simp only: 1 add.commute)
  also have "... = s + gcs (pp_of_term v) t" by (simp only: gcs_plus_left)
  also have "... = s" by (simp add: 3)
  finally have "t = s" .
  moreover from t  u = s  v have "u = v" by (simp only: t = s splus_left_canc)
  ultimately show ?thesis by (simp add: a b)
qed

lemma min_set_finite:
  fixes seq :: "nat  ('t 0 'b::field)"
  assumes "dickson_grading d" and "range seq  dgrad_sig_set d" and "0  rep_list ` range seq"
    and "i j. i < j  lt (seq i) t lt (seq j)"
  shows "finite {i. ¬ (j<i. lt (seq j) addst lt (seq i) 
                             punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i)))}"
proof -
  have "inj (λi. lt (seq i))"
  proof
    fix i j
    assume eq: "lt (seq i) = lt (seq j)"
    show "i = j"
    proof (rule linorder_cases)
      assume "i < j"
      hence "lt (seq i) t lt (seq j)" by (rule assms(4))
      thus ?thesis by (simp add: eq)
    next
      assume "j < i"
      hence "lt (seq j) t lt (seq i)" by (rule assms(4))
      thus ?thesis by (simp add: eq)
    qed
  qed
  hence "inj seq" unfolding comp_def[symmetric] by (rule inj_on_imageI2)

  let ?P1 = "λp q. lt p addst lt q"
  let ?P2 = "λp q. punit.lt (rep_list p) adds punit.lt (rep_list q)"
  let ?P = "λp q. ?P1 p q  ?P2 p q"
  have "reflp ?P" by (simp add: reflp_def adds_term_refl)
  have "almost_full_on ?P1 (range seq)"
  proof (rule almost_full_on_map)
    let ?B = "{t. pp_of_term t  dgrad_set d (dgrad_max d)  component_of_term t  {0..<length fs}}"
    from assms(1) finite_atLeastLessThan show "almost_full_on (addst) ?B" by (rule Dickson_term)
    show "lt ` range seq  ?B"
    proof
      fix v
      assume "v  lt ` range seq"
      then obtain p where "p  range seq" and v: "v = lt p" ..
      from this(1) assms(3) have "rep_list p  0" by auto
      hence "p  0" by (auto simp: rep_list_zero)
      from p  range seq assms(2) have "p  dgrad_sig_set d" ..
      hence "d (lp p)  dgrad_max d" by (rule dgrad_sig_setD_lp)
      hence "lp p  dgrad_set d (dgrad_max d)" by (simp add: dgrad_set_def)
      moreover from p  dgrad_sig_set d p  0 have "component_of_term (lt p) < length fs"
        by (rule dgrad_sig_setD_lt)
      ultimately show "v  ?B" by (simp add: v)
    qed
  qed
  moreover have "almost_full_on ?P2 (range seq)"
  proof (rule almost_full_on_map)
    let ?B = "dgrad_set d (dgrad_max d)"
    from assms(1) show "almost_full_on (adds) ?B" by (rule dickson_gradingD_dgrad_set)
    show "(λp. punit.lt (rep_list p)) ` range seq  ?B"
    proof
      fix t
      assume "t  (λp. punit.lt (rep_list p)) ` range seq"
      then obtain p where "p  range seq" and t: "t = punit.lt (rep_list p)" ..
      from this(1) assms(3) have "rep_list p  0" by auto
      from p  range seq assms(2) have "p  dgrad_sig_set d" ..
      hence "p  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
      with assms(1) have "rep_list p  punit_dgrad_max_set d" by (rule dgrad_max_2)
      from this rep_list p  0 have "d (punit.lt (rep_list p))  dgrad_max d"
        by (rule punit.dgrad_p_setD_lp[simplified])
      thus "t  ?B" by (simp add: t dgrad_set_def)
    qed
  qed
  ultimately have "almost_full_on ?P (range seq)" by (rule almost_full_on_same)
  with reflp ?P obtain T where "finite T" and "T  range seq" and *: "p. p  range seq  (qT. ?P q p)"
    by (rule almost_full_on_finite_subsetE, blast)
  from T  range seq obtain I where T: "T = seq ` I" by (meson subset_image_iff)
  have "{i. ¬ (j<i. ?P (seq j) (seq i))}  I"
  proof
    fix i
    assume "i  {i. ¬ (j<i. ?P (seq j) (seq i))}"
    hence x: "¬ (j<i. ?P (seq j) (seq i))" by simp
    obtain j where "j  I" and "?P (seq j) (seq i)"
    proof -
      have "seq i  range seq" by simp
      hence "qT. ?P q (seq i)" by (rule *)
      then obtain q where "q  T" and "?P q (seq i)" ..
      from this(1) obtain j where "j  I" and "q = seq j" unfolding T ..
      from this(1) ?P q (seq i) show ?thesis unfolding q = seq j ..
    qed
    from this(2) x have "i  j" by auto
    moreover have "¬ i < j"
    proof
      assume "i < j"
      hence "lt (seq i) t lt (seq j)" by (rule assms(4))
      hence "¬ ?P1 (seq j) (seq i)" using ord_adds_term ord_term_lin.leD by blast
      with ?P (seq j) (seq i) show False by simp
    qed
    ultimately show "i  I" using j  I by simp
  qed
  moreover from inj seq finite T have "finite I" by (simp add: finite_image_iff inj_on_subset T)
  ultimately show ?thesis by (rule finite_subset)
qed

lemma rb_termination:
  fixes seq :: "nat  ('t 0 'b::field)"
  assumes "dickson_grading d" and "range seq  dgrad_sig_set d" and "0  rep_list ` range seq"
    and "i j. i < j  lt (seq i) t lt (seq j)"
    and "i. ¬ is_sig_red (≺t) (≼) (seq ` {0..<i}) (seq i)"
    and "i. (j<length fs. lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j))) 
                punit.lt (rep_list (seq i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
              (j k. is_regular_spair (seq j) (seq k)  rep_list (spair (seq j) (seq k))  0 
                    lt (seq i) = lt (spair (seq j) (seq k)) 
                    punit.lt (rep_list (seq i))  punit.lt (rep_list (spair (seq j) (seq k))))"
    and "i. is_sig_GB_upt d (seq ` {0..<i}) (lt (seq i))"
  shows thesis
proof -
  from assms(3) have "0  range seq" using rep_list_zero by auto
  have "ord_term_lin.is_le_rel (=)" and "ord_term_lin.is_le_rel (≺t)" by (rule ord_term_lin.is_le_relI)+
  have "reflp (=)" and "symp (=)" by (simp_all add: symp_def)
  have "irreflp (≺t)" by (simp add: irreflp_def)
  have "inj (λi. lt (seq i))"
  proof
    fix i j
    assume eq: "lt (seq i) = lt (seq j)"
    show "i = j"
    proof (rule linorder_cases)
      assume "i < j"
      hence "lt (seq i) t lt (seq j)" by (rule assms(4))
      thus ?thesis by (simp add: eq)
    next
      assume "j < i"
      hence "lt (seq j) t lt (seq i)" by (rule assms(4))
      thus ?thesis by (simp add: eq)
    qed
  qed
  hence "inj seq" unfolding comp_def[symmetric] by (rule inj_on_imageI2)

  define R where "R = (λx. {i. term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) x})"
  let ?A = "{x. canon_term_pp_pair x  R x  {}}"

  have "finite ?A"
  proof -
    define min_set where "min_set = {i. ¬ (j<i. lt (seq j) addst lt (seq i) 
                                      punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i)))}"
    have "?A  (λi. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set"
    proof
      fix u t
      assume "(u, t)  ?A"
      hence "canon_term_pp_pair (u, t)" and "R (u, t)  {}" by simp_all
      from this(2) obtain i where x: "term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) (u, t)"
        by (auto simp: R_def)
      let ?equiv = "(λi j. term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) (lt (seq j), punit.lt (rep_list (seq j))))"
      obtain j where "j  min_set" and "?equiv j i"
      proof (cases "i  min_set")
        case True
        moreover have "?equiv i i" by (simp add: term_pp_rel_refl)
        ultimately show ?thesis ..
      next
        case False
        let ?Q = "{seq j | j. j < i  is_sig_red (=) (=) {seq j} (seq i)}"
        have "?Q  range seq" by blast
        also have "...  dgrad_sig_set d" by (fact assms(2))
        finally have "?Q  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
        moreover from ?Q  range seq 0  range seq have "0  ?Q" by blast
        ultimately have Q_sub: "pp_of_term ` lt ` ?Q  dgrad_set d (dgrad_max d)"
          unfolding image_image by (smt CollectI dgrad_p_setD_lp dgrad_set_def image_subset_iff subsetCE)
        have *: "gseq ` {0..<k}. is_sig_red (=) (=) {g} (seq k)" if "k  min_set" for k
          proof -
          from that obtain j where "j < k" and a: "lt (seq j) addst lt (seq k)"
            and b: "punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq k))" by (auto simp: min_set_def)
          note assms(1, 7)
          moreover from assms(2) have "seq k  dgrad_sig_set d" by fastforce
          moreover from j < k have "seq j  seq ` {0..<k}" by simp
          moreover from assms(3) have "rep_list (seq k)  0" and "rep_list (seq j)  0" by fastforce+
          ultimately have "is_sig_red (≼t) (=) (seq ` {0..<k}) (seq k)" using a b by (rule lemma_21)
          moreover from assms(5)[of k] have "¬ is_sig_red (≺t) (=) (seq ` {0..<k}) (seq k)"
            by (simp add: is_sig_red_top_tail_cases)
          ultimately have "is_sig_red (=) (=) (seq ` {0..<k}) (seq k)"
            by (simp add: is_sig_red_sing_reg_cases)
          then obtain g0 where "g0  seq ` {0..<k}" and "is_sig_red (=) (=) {g0} (seq k)"
            by (rule is_sig_red_singletonI)
          thus ?thesis ..
        qed

        from this[OF False] obtain g0 where "g0  seq ` {0..<i}" and "is_sig_red (=) (=) {g0} (seq i)" ..
        hence "g0  ?Q" by fastforce
        hence "lt g0  lt ` ?Q" by (rule imageI)
        with assms(1) obtain v where "v  lt ` ?Q" and min: "v'. v' t v  v'  lt ` ?Q"
          using Q_sub by (rule ord_term_minimum_dgrad_set, blast)
        from this(1) obtain j where "j < i" and "is_sig_red (=) (=) {seq j} (seq i)"
          and v: "v = lt (seq j)" by fastforce
        hence 1: "punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i))"
          and 2: "punit.lt (rep_list (seq i))  lt (seq j) = punit.lt (rep_list (seq j))  lt (seq i)"
          by (auto elim: is_sig_red_top_addsE)
        show ?thesis
        proof
          show "?equiv j i" by (simp add: term_pp_rel_def 2)
        next
          show "j  min_set"
          proof (rule ccontr)
            assume "j  min_set"
            from *[OF this] obtain g1 where "g1  seq ` {0..<j}" and red: "is_sig_red (=) (=) {g1} (seq j)" ..
            from this(1) obtain j0 where "j0 < j" and "g1 = seq j0" by fastforce+

            from red have 3: "punit.lt (rep_list (seq j0)) adds punit.lt (rep_list (seq j))"
              and 4: "punit.lt (rep_list (seq j))  lt (seq j0) = punit.lt (rep_list (seq j0))  lt (seq j)"
              by (auto simp: g1 = seq j0 elim: is_sig_red_top_addsE)

            from j0 < j j < i have "j0 < i" by simp
            from j0 < j have "lt (seq j0) t v" unfolding v by (rule assms(4))
            hence "lt (seq j0)  lt `?Q" by (rule min)
            with j0 < i have "¬ is_sig_red (=) (=) {seq j0} (seq i)" by blast
            moreover have "is_sig_red (=) (=) {seq j0} (seq i)"
            proof (rule is_sig_red_top_addsI)
              from assms(3) show "rep_list (seq j0)  0" by fastforce
            next
              from assms(3) show "rep_list (seq i)  0" by fastforce
            next
              from 3 1 show "punit.lt (rep_list (seq j0)) adds punit.lt (rep_list (seq i))"
                by (rule adds_trans)
            next
              from 4 have "?equiv j0 j" by (simp add: term_pp_rel_def)
              also from 2 have "?equiv j i" by (simp add: term_pp_rel_def)
              finally(term_pp_rel_trans[OF ord_term_lin.is_le_rel (=)])
              show "punit.lt (rep_list (seq i))  lt (seq j0) = punit.lt (rep_list (seq j0))  lt (seq i)"
                by (simp add: term_pp_rel_def)
            next
              show "ord_term_lin.is_le_rel (=)" by simp
            qed simp_all
            ultimately show False ..
          qed
        qed
      qed
      have "term_pp_rel (=) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))) (lt (seq j), punit.lt (rep_list (seq j)))"
        by (rule term_pp_rel_sym, fact symp (=), rule term_pp_rel_cancel, fact reflp (=))
      also note ?equiv j i
      also(term_pp_rel_trans[OF ord_term_lin.is_le_rel (=)]) note x
      finally(term_pp_rel_trans[OF ord_term_lin.is_le_rel (=)])
      have "term_pp_rel (=) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))) (u, t)" .
      with symp (=) have "term_pp_rel (=) (u, t) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j))))"
        by (rule term_pp_rel_sym)
      hence "(u, t) = cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))"
        using canon_term_pp_pair (u, t) canon_term_pp_cancel by (rule canon_term_pp_rel_id)
      with j  min_set show "(u, t)  (λi. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set"
        by fastforce
    qed
    moreover have "finite ((λi. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set)"
    proof (rule finite_imageI)
      show "finite min_set" unfolding min_set_def using assms(1-4) by (rule min_set_finite)
    qed
    ultimately show ?thesis by (rule finite_subset)
  qed

  have "range seq  seq ` ( (R ` ?A))"
  proof (rule image_mono, rule)
    fix i
    show "i  ( (R ` ?A))"
    proof
      show "i  R (cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i))))"
        by (simp add: R_def term_pp_rel_cancel)
      thus "cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))  ?A"
        using canon_term_pp_cancel by blast
    qed
  qed
  moreover from inj seq have "infinite (range seq)" by (rule range_inj_infinite)
  ultimately have "infinite (seq ` ( (R ` ?A)))" by (rule infinite_super)
  moreover have "finite (seq ` ( (R ` ?A)))"
  proof (rule finite_imageI, rule finite_UN_I)
    fix x
    assume "x  ?A"
    let ?rel = "term_pp_rel (≺t)"
    have "irreflp ?rel" by (rule irreflpI, rule term_pp_rel_irrefl, fact)
    moreover have "transp ?rel" by (rule transpI, drule term_pp_rel_trans[OF ord_term_lin.is_le_rel (≺t)])
    ultimately have "wfp_on ?rel ?A" using finite ?A by (rule wfp_on_finite)
    thus "finite (R x)" using x  ?A
    proof (induct rule: wfp_on_induct)
      case (less x)
      from less(1) have "canon_term_pp_pair x" by simp
      define R' where " R' =  (R ` ({x. canon_term_pp_pair x  R x  {}}  {z. term_pp_rel (≺t) z x}))"
      define red_set where "red_set = (λp::'t 0 'b. {k. lt (seq k) = lt p 
                                              punit.lt (rep_list (seq k))  punit.lt (rep_list p)})"
      have finite_red_set: "finite (red_set p)" for p
      proof (cases "red_set p = {}")
        case True
        thus ?thesis by simp
      next
        case False
        then obtain k where lt_k: "lt (seq k) = lt p" by (auto simp: red_set_def)
        have "red_set p  {k}"
        proof
          fix k'
          assume "k'  red_set p"
          hence "lt (seq k') = lt p" by (simp add: red_set_def)
          hence "lt (seq k') = lt (seq k)" by (simp only: lt_k)
          with inj (λi. lt (seq i)) have "k' = k" by (rule injD)
          thus "k'  {k}" by simp
        qed
        thus ?thesis using infinite_super by auto
      qed

      have "R x  (iR'. jR'. red_set (spair (seq i) (seq j))) 
                   (j{0..<length fs}. red_set (monomial 1 (term_of_pair (0, j))))"
        (is "_  ?B  ?C")
      proof
        fix i
        assume "i  R x"
        hence i_x: "term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) x"
          by (simp add: R_def term_pp_rel_def)
        from assms(6)[of i] show "i  ?B  ?C"
        proof (elim disjE exE conjE)
          fix j
          assume "j < length fs"
          hence "j  {0..<length fs}" by simp
          assume "lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j)))"
            and "punit.lt (rep_list (seq i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))"
          hence "i  red_set (monomial 1 (term_of_pair (0, j)))" by (simp add: red_set_def)
          with j  {0..<length fs} have "i  ?C" ..
          thus ?thesis ..
        next
          fix j k
          let ?li = "punit.lt (rep_list (seq i))"
          let ?lj = "punit.lt (rep_list (seq j))"
          let ?lk = "punit.lt (rep_list (seq k))"
          assume lt_i: "lt (seq i) = lt (spair (seq j) (seq k))"
            and lt_i': "?li  punit.lt (rep_list (spair (seq j) (seq k)))"
            and spair_0: "rep_list (spair (seq j) (seq k))  0"
          hence "i  red_set (spair (seq j) (seq k))" by (simp add: red_set_def)
          from assms(3) have i_0: "rep_list (seq i)  0" and j_0: "rep_list (seq j)  0"
            and k_0: "rep_list (seq k)  0" by fastforce+

          have R'I: "a  R'" if "term_pp_rel (≺t) (lt (seq a), punit.lt (rep_list (seq a))) x" for a
          proof -
            let ?x = "cancel_term_pp_pair (lt (seq a), punit.lt (rep_list (seq a)))"
            show ?thesis unfolding R'_def
            proof (rule UN_I, simp, intro conjI)
              show "a  R ?x" by (simp add: R_def term_pp_rel_cancel)
              thus "R ?x  {}" by blast
            next
              note ord_term_lin.is_le_rel (≺t)
              moreover have "term_pp_rel (=) ?x (lt (seq a), punit.lt (rep_list (seq a)))"
                by (rule term_pp_rel_sym, fact, rule term_pp_rel_cancel, fact)
              ultimately show "term_pp_rel (≺t) ?x x" using that by (rule term_pp_rel_trans_eq_left)
            qed (fact canon_term_pp_cancel)
          qed

          assume "is_regular_spair (seq j) (seq k)"
          hence "?lk  lt (seq j)  ?lj  lt (seq k)" by (rule is_regular_spairD3)
          hence "term_pp_rel (≺t) (lt (seq j), ?lj) x  term_pp_rel (≺t) (lt (seq k), ?lk) x"
          proof (rule ord_term_lin.neqE)
            assume c: "?lk  lt (seq j) t ?lj  lt (seq k)"
            hence j_k: "term_pp_rel (≺t) (lt (seq j), ?lj) (lt (seq k), ?lk)"
              by (simp add: term_pp_rel_def)
            note ord_term_lin.is_le_rel (≺t)
            moreover have "term_pp_rel (≺t) (lt (seq k), ?lk) (lt (seq i), ?li)"
            proof (simp add: term_pp_rel_def)
              from lt_i' have "?li  lt (seq k) t
                                punit.lt (rep_list (spair (seq j) (seq k)))  lt (seq k)"
                by (rule splus_mono_left)
              also have "... t (?lk - gcs ?lk ?lj + ?lj)  lt (seq k)"
                by (rule splus_mono_strict_left, rule lt_rep_list_spair, fact+, simp only: add.commute)
              also have "... = ((?lk + ?lj) - gcs ?lj ?lk)  lt (seq k)"
                by (simp add: minus_plus gcs_adds_2 gcs_comm)
              also have "... = ?lk  ((?lj - gcs ?lj ?lk)  lt (seq k))"
                by (simp add: minus_plus' gcs_adds splus_assoc[symmetric])
              also have "... = ?lk  lt (seq i)"
                by (simp add: lt_spair'[OF k_0 _ c] add.commute spair_comm[of "seq j"] lt_i)
              finally show "?li  lt (seq k) t ?lk  lt (seq i)" .
            qed
            ultimately have "term_pp_rel (≺t) (lt (seq k), ?lk) x" using i_x
              by (rule term_pp_rel_trans_eq_right)
            moreover from ord_term_lin.is_le_rel (≺t) j_k this
            have "term_pp_rel (≺t) (lt (seq j), ?lj) x" by (rule term_pp_rel_trans)
            ultimately show ?thesis by simp
          next
            assume c: "?lj  lt (seq k) t ?lk  lt (seq j)"
            hence j_k: "term_pp_rel (≺t) (lt (seq k), ?lk) (lt (seq j), ?lj)"
              by (simp add: term_pp_rel_def)
            note ord_term_lin.is_le_rel (≺t)
            moreover have "term_pp_rel (≺t) (lt (seq j), ?lj) (lt (seq i), ?li)"
            proof (simp add: term_pp_rel_def)
              from lt_i' have "?li  lt (seq j) t
                                punit.lt (rep_list (spair (seq j) (seq k)))  lt (seq j)"
                by (rule splus_mono_left)
              thm lt_rep_list_spair
              also have "... t (?lk - gcs ?lk ?lj + ?lj)  lt (seq j)"
                by (rule splus_mono_strict_left, rule lt_rep_list_spair, fact+, simp only: add.commute)
              also have "... = ((?lk + ?lj) - gcs ?lk ?lj)  lt (seq j)"
                by (simp add: minus_plus gcs_adds_2 gcs_comm)
              also have "... = ?lj  ((?lk - gcs ?lk ?lj)  lt (seq j))"
                by (simp add: minus_plus' gcs_adds splus_assoc[symmetric] add.commute)
              also have "... = ?lj  lt (seq i)" by (simp add: lt_spair'[OF j_0 _ c] lt_i add.commute)
              finally show "?li  lt (seq j) t ?lj  lt (seq i)" .
            qed
            ultimately have "term_pp_rel (≺t) (lt (seq j), ?lj) x" using i_x
              by (rule term_pp_rel_trans_eq_right)
            moreover from ord_term_lin.is_le_rel (≺t) j_k this
            have "term_pp_rel (≺t) (lt (seq k), ?lk) x" by (rule term_pp_rel_trans)
            ultimately show ?thesis by simp
          qed
          with i  red_set (spair (seq j) (seq k)) have "i  ?B" using R'I by blast
          thus ?thesis ..
        qed
      qed
      moreover have "finite (?B  ?C)"
      proof (rule finite_UnI)
        have "finite R'" unfolding R'_def
        proof (rule finite_UN_I)
          from finite ?A show "finite (?A  {z. term_pp_rel (≺t) z x})" by simp
        next
          fix y
          assume "y  ?A  {z. term_pp_rel (≺t) z x}"
          hence "y  ?A" and "term_pp_rel (≺t) y x" by simp_all
          thus "finite (R y)" by (rule less(2))
        qed
        show "finite ?B" by (intro finite_UN_I finite R' finite_red_set)
      next
        show "finite ?C" by (intro finite_UN_I finite_atLeastLessThan finite_red_set)
      qed
      ultimately show ?case by (rule finite_subset)
    qed
  qed fact
  ultimately show ?thesis ..
qed

subsubsection ‹Concrete Rewrite Orders›

definition is_strict_rewrite_ord :: "(('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  bool"
  where "is_strict_rewrite_ord rel  is_rewrite_ord (λx y. ¬ rel y x)"

lemma is_strict_rewrite_ordI: "is_rewrite_ord (λx y. ¬ rel y x)  is_strict_rewrite_ord rel"
  unfolding is_strict_rewrite_ord_def by blast

lemma is_strict_rewrite_ordD: "is_strict_rewrite_ord rel  is_rewrite_ord (λx y. ¬ rel y x)"
  unfolding is_strict_rewrite_ord_def by blast

lemma is_strict_rewrite_ord_antisym:
  assumes "is_strict_rewrite_ord rel" and "¬ rel x y" and "¬ rel y x"
  shows "fst x = fst y"
  by (rule is_rewrite_ordD4, rule is_strict_rewrite_ordD, fact+)

lemma is_strict_rewrite_ord_asym:
  assumes "is_strict_rewrite_ord rel" and "rel x y"
  shows "¬ rel y x"
proof -
  from assms(1) have "is_rewrite_ord (λx y. ¬ rel y x)" by (rule is_strict_rewrite_ordD)
  thus ?thesis
  proof (rule is_rewrite_ordD3)
    assume "¬ ¬ rel y x"
    assume "¬ rel x y"
    thus ?thesis using rel x y ..
  qed
qed

lemma is_strict_rewrite_ord_irrefl: "is_strict_rewrite_ord rel  ¬ rel x x"
  using is_strict_rewrite_ord_asym by blast

definition rw_rat :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "rw_rat p q  (let u = punit.lt (snd q)  fst p; v = punit.lt (snd p)  fst q in
                          u t v  (u = v  fst p t fst q))"

definition rw_rat_strict :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "rw_rat_strict p q  (let u = punit.lt (snd q)  fst p; v = punit.lt (snd p)  fst q in
                          u t v  (u = v  fst p t fst q))"

definition rw_add :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "rw_add p q  (fst p t fst q)"

definition rw_add_strict :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "rw_add_strict p q  (fst p t fst q)"

lemma rw_rat_alt: "rw_rat = (λp q. ¬ rw_rat_strict q p)"
  by (intro ext, auto simp: rw_rat_def rw_rat_strict_def Let_def)

lemma rw_rat_is_rewrite_ord: "is_rewrite_ord rw_rat"
proof (rule is_rewrite_ordI)
  show "reflp rw_rat" by (simp add: reflp_def rw_rat_def)
next
  have 1: "ord_term_lin.is_le_rel (≺t)" and 2: "ord_term_lin.is_le_rel (=)"
    by (rule ord_term_lin.is_le_relI)+
  have "rw_rat p q  (term_pp_rel (≺t) (fst p, punit.lt (snd p)) (fst q, punit.lt (snd q)) 
                        (term_pp_rel (=) (fst p, punit.lt (snd p)) (fst q, punit.lt (snd q)) 
                          fst p t fst q))"
    for p q
  by (simp add: rw_rat_def term_pp_rel_def Let_def)
  thus "transp rw_rat"
    by (auto simp: transp_def dest: term_pp_rel_trans[OF 1] term_pp_rel_trans_eq_left[OF 1]
        term_pp_rel_trans_eq_right[OF 1] term_pp_rel_trans[OF 2])
next
  fix p q
  show "rw_rat p q  rw_rat q p" by (auto simp: rw_rat_def Let_def)
next
  fix p q
  assume "rw_rat p q" and "rw_rat q p"
  thus "fst p = fst q" by (auto simp: rw_rat_def Let_def)
next
  fix d G p q
  assume d: "dickson_grading d" and gb: "is_sig_GB_upt d G (lt q)" and "p  G" and "q  G"
    and "p  0" and "q  0" and "lt p addst lt q" and "¬ is_sig_red (≺t) (=) G q"
  let ?u = "punit.lt (rep_list q)  lt p"
  let ?v = "punit.lt (rep_list p)  lt q"
  from lt p addst lt q obtain t where lt_q: "lt q = t  lt p" by (rule adds_termE)
  from gb have "G  dgrad_sig_set d" by (rule is_sig_GB_uptD1)
  hence "G  dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
  with d obtain p' where red: "(sig_red (≺t) (=) G)** (monom_mult 1 t p) p'"
    and "¬ is_sig_red (≺t) (=) G p'" by (rule sig_irredE_dgrad_max_set)
  from red have "lt p' = lt (monom_mult 1 t p)" and "lc p' = lc (monom_mult 1 t p)"
    and 2: "punit.lt (rep_list p')  punit.lt (rep_list (monom_mult 1 t p))"
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc, rule sig_red_rtrancl_lt_rep_list)
  with p  0 have "lt p' = lt q" and "lc p' = lc p" by (simp_all add: lt_q lt_monom_mult)
  from 2 punit.lt_monom_mult_le[simplified] have 3: "punit.lt (rep_list p')  t + punit.lt (rep_list p)"
    unfolding rep_list_monom_mult by (rule ordered_powerprod_lin.order_trans)
  have "punit.lt (rep_list p') = punit.lt (rep_list q)"
  proof (rule sig_regular_top_reduced_lt_unique)
    show "p'  dgrad_sig_set d"
    proof (rule dgrad_sig_set_closed_sig_red_rtrancl)
      note d
      moreover have "d t  dgrad_max d"
      proof (rule le_trans)
        have "t adds lp q" by (simp add: lt_q term_simps)
        with d show "d t  d (lp q)" by (rule dickson_grading_adds_imp_le)
      next
        from q  G G  dgrad_max_set d have "q  dgrad_max_set d" ..
        thus "d (lp q)  dgrad_max d" using q  0 by (rule dgrad_p_setD_lp)
      qed
      moreover from p  G G  dgrad_sig_set d have "p  dgrad_sig_set d" ..
      ultimately show "monom_mult 1 t p  dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult)
    qed fact+
  next
    from q  G G  dgrad_sig_set d show "q  dgrad_sig_set d" ..
  next
    from p  0 lc p' = lc p show "p'  0" by (auto simp: lc_eq_zero_iff)
  qed fact+
  with 3 have "punit.lt (rep_list q)  t + punit.lt (rep_list p)" by simp
  hence "?u t (t + punit.lt (rep_list p))  lt p" by (rule splus_mono_left)
  also have "... = ?v" by (simp add: lt_q splus_assoc splus_left_commute)
  finally have "?u t ?v" by (simp only: rel_def)
  moreover from lt p addst lt q have "lt p t lt q" by (rule ord_adds_term)
  ultimately show "rw_rat (spp_of p) (spp_of q)" by (auto simp: rw_rat_def Let_def spp_of_def)
qed

lemma rw_rat_strict_is_strict_rewrite_ord: "is_strict_rewrite_ord rw_rat_strict"
proof (rule is_strict_rewrite_ordI)
  show "is_rewrite_ord (λx y. ¬ rw_rat_strict y x)"
    unfolding rw_rat_alt[symmetric] by (fact rw_rat_is_rewrite_ord)
qed

lemma rw_add_alt: "rw_add = (λp q. ¬ rw_add_strict q p)"
  by (intro ext, auto simp: rw_add_def rw_add_strict_def)

lemma rw_add_is_rewrite_ord: "is_rewrite_ord rw_add"
proof (rule is_rewrite_ordI)
  show "reflp rw_add" by (simp add: reflp_def rw_add_def)
next
  show "transp rw_add" by (auto simp: transp_def rw_add_def)
next
  fix p q
  show "rw_add p q  rw_add q p" by (simp only: rw_add_def ord_term_lin.linear)
next
  fix p q
  assume "rw_add p q" and "rw_add q p"
  thus "fst p = fst q" unfolding rw_add_def
    by simp
next
  fix p q :: "'t 0 'b"
  assume "lt p addst lt q"
  thus "rw_add (spp_of p) (spp_of q)" unfolding rw_add_def spp_of_def fst_conv by (rule ord_adds_term)
qed

lemma rw_add_strict_is_strict_rewrite_ord: "is_strict_rewrite_ord rw_add_strict"
proof (rule is_strict_rewrite_ordI)
  show "is_rewrite_ord (λx y. ¬ rw_add_strict y x)"
    unfolding rw_add_alt[symmetric] by (fact rw_add_is_rewrite_ord)
qed

subsubsection ‹Preparations for Sig-Poly-Pairs›

context
  fixes dgrad :: "'a  nat"
begin

definition spp_rel :: "('t × ('a 0 'b))  ('t 0 'b)  bool"
  where "spp_rel sp r  (r  0  r  dgrad_sig_set dgrad  lt r = fst sp  rep_list r = snd sp)"

definition spp_inv :: "('t × ('a 0 'b))  bool"
  where "spp_inv sp  Ex (spp_rel sp)"

definition vec_of :: "('t × ('a 0 'b))  ('t 0 'b)"
  where "vec_of sp = (if spp_inv sp then Eps (spp_rel sp) else 0)"

lemma spp_inv_spp_of:
  assumes "r  0" and "r  dgrad_sig_set dgrad"
  shows "spp_inv (spp_of r)"
  unfolding spp_inv_def spp_rel_def
proof (intro exI conjI)
  show "lt r = fst (spp_of r)" by (simp add: spp_of_def)
next
  show "rep_list r = snd (spp_of r)" by (simp add: spp_of_def)
qed fact+

context
  fixes sp :: "'t × ('a 0 'b)"
  assumes spi: "spp_inv sp"
begin

lemma sig_poly_rel_vec_of: "spp_rel sp (vec_of sp)"
proof -
  from spi have eq: "vec_of sp = Eps (spp_rel sp)" by (simp add: vec_of_def)
  from spi show ?thesis unfolding eq spp_inv_def by (rule someI_ex)
qed

lemma vec_of_nonzero: "vec_of sp  0"
  using sig_poly_rel_vec_of by (simp add: spp_rel_def)

lemma lt_vec_of: "lt (vec_of sp) = fst sp"
  using sig_poly_rel_vec_of by (simp add: spp_rel_def)

lemma rep_list_vec_of: "rep_list (vec_of sp) = snd sp"
  using sig_poly_rel_vec_of by (simp add: spp_rel_def)

lemma spp_of_vec_of: "spp_of (vec_of sp) = sp"
  by (simp add: spp_of_def lt_vec_of rep_list_vec_of)

end

lemma map_spp_of_vec_of:
  assumes "list_all spp_inv sps"
  shows "map (spp_of  vec_of) sps = sps"
proof (rule map_idI)
  fix sp
  assume "sp  set sps"
  with assms have "spp_inv sp" by (simp add: list_all_def)
  hence "spp_of (vec_of sp) = sp" by (rule spp_of_vec_of)
  thus "(spp_of  vec_of) sp = sp" by simp
qed

lemma vec_of_dgrad_sig_set: "vec_of sp  dgrad_sig_set dgrad"
proof (cases "spp_inv sp")
  case True
  hence "spp_rel sp (vec_of sp)" by (rule sig_poly_rel_vec_of)
  thus ?thesis by (simp add: spp_rel_def)
next
  case False
  moreover have "0  dgrad_sig_set dgrad" unfolding dgrad_sig_set'_def
  proof
    show "0  dgrad_max_set dgrad" by (rule dgrad_p_setI) simp
  next
    show "0  sig_inv_set" by (rule sig_inv_setI) (simp add: term_simps)
  qed
  ultimately show ?thesis by (simp add: vec_of_def)
qed

lemma spp_invD_fst:
  assumes "spp_inv sp"
  shows "dgrad (pp_of_term (fst sp))  dgrad_max dgrad" and "component_of_term (fst sp) < length fs"
proof -
  from vec_of_dgrad_sig_set have "dgrad (lp (vec_of sp))  dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
  with assms show "dgrad (pp_of_term (fst sp))  dgrad_max dgrad" by (simp add: lt_vec_of)
  from vec_of_dgrad_sig_set vec_of_nonzero[OF assms] have "component_of_term (lt (vec_of sp)) < length fs"
    by (rule dgrad_sig_setD_lt)
  with assms show "component_of_term (fst sp) < length fs" by (simp add: lt_vec_of)
qed

lemma spp_invD_snd:
  assumes "dickson_grading dgrad" and "spp_inv sp"
  shows "snd sp  punit_dgrad_max_set dgrad"
proof -
  from vec_of_dgrad_sig_set[of sp] have "vec_of sp  dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
  with assms(1) have "rep_list (vec_of sp)  punit_dgrad_max_set dgrad" by (rule dgrad_max_2)
  with assms(2) show ?thesis by (simp add: rep_list_vec_of)
qed

lemma vec_of_inj:
  assumes "spp_inv sp" and "vec_of sp = vec_of sp'"
  shows "sp = sp'"
proof -
  from assms(1) have "vec_of sp  0" by (rule vec_of_nonzero)
  hence "vec_of sp'  0" by (simp add: assms(2))
  hence "spp_inv sp'" by (simp add: vec_of_def split: if_split_asm)
  from assms(1) have "sp = spp_of (vec_of sp)" by (simp only: spp_of_vec_of)
  also have "... = spp_of (vec_of sp')" by (simp only: assms(2))
  also from spp_inv sp' have "... = sp'" by (rule spp_of_vec_of)
  finally show ?thesis .
qed

lemma spp_inv_alt: "spp_inv sp  (vec_of sp  0)"
proof -
  have "spp_inv sp" if "vec_of sp  0"
  proof (rule ccontr)
    assume "¬ spp_inv sp"
    hence "vec_of sp = 0" by (simp add: vec_of_def)
    with that show False ..
  qed
  thus ?thesis by (auto dest: vec_of_nonzero)
qed

lemma spp_of_vec_of_spp_of:
  assumes "p  dgrad_sig_set dgrad"
  shows "spp_of (vec_of (spp_of p)) = spp_of p"
proof (cases "p = 0")
  case True
  show ?thesis
  proof (cases "spp_inv (spp_of p)")
    case True
    thus ?thesis by (rule spp_of_vec_of)
  next
    case False
    hence "vec_of (spp_of p) = 0" by (simp add: spp_inv_alt)
    thus ?thesis by (simp only: True)
  qed
next
  case False
  have "spp_inv (spp_of p)" unfolding spp_inv_def
  proof
    from False assms show "spp_rel (spp_of p) p" by (simp add: spp_rel_def spp_of_def)
  qed
  thus ?thesis by (rule spp_of_vec_of)
qed

subsubsection ‹Total Reduction›

primrec find_sig_reducer :: "('t × ('a 0 'b)) list  't  'a  nat  nat option" where
  "find_sig_reducer [] _ _ _ = None"|
  "find_sig_reducer (b # bs) u t i =
          (if snd b  0  punit.lt (snd b) adds t  (t - punit.lt (snd b))  fst b t u then Some i
           else find_sig_reducer bs u t (Suc i))"

lemma find_sig_reducer_SomeD_aux:
  assumes "find_sig_reducer bs u t i = Some j"
  shows "i  j" and "j - i < length bs"
proof -
  from assms have "i  j  j - i < length bs"
  proof (induct bs arbitrary: i)
    case Nil
    thus ?case by simp
  next
    case (Cons b bs)
    from Cons(2) show ?case
    proof (simp split: if_split_asm)
      assume "find_sig_reducer bs u t (Suc i) = Some j"
      hence "Suc i  j  j - Suc i < length bs" by (rule Cons(1))
      thus "i  j  j - i < Suc (length bs)" by auto
    qed
  qed
  thus "i  j" and "j - i < length bs" by simp_all
qed

lemma find_sig_reducer_SomeD':
  assumes "find_sig_reducer bs u t i = Some j" and "b = bs ! (j - i)"
  shows "b  set bs" and "snd b  0" and "punit.lt (snd b) adds t" and "(t - punit.lt (snd b))  fst b t u"
proof -
  from assms(1) have "j - i < length bs" by (rule find_sig_reducer_SomeD_aux)
  thus "b  set bs" unfolding assms(2) by (rule nth_mem)
next
  from assms have "snd b  0  punit.lt (snd b) adds t  (t - punit.lt (snd b))  fst b t u"
  proof (induct bs arbitrary: i)
    case Nil
    from Nil(1) show ?case by simp
  next
    case (Cons a bs)
    from Cons(2) show ?case
    proof (simp split: if_split_asm)
      assume "i = j"
      with Cons(3) have "b = a" by simp
      moreover assume "snd a  0" and "punit.lt (snd a) adds t" and "(t - punit.lt (snd a))  fst a t u"
      ultimately show ?case by simp
    next
      assume *: "find_sig_reducer bs u t (Suc i) = Some j"
      hence "Suc i  j" by (rule find_sig_reducer_SomeD_aux)
      note Cons(3)
      also from Suc i  j have "(a # bs) ! (j - i) = bs ! (j - Suc i)" by simp
      finally have "b = bs ! (j - Suc i)" .
      with * show ?case by (rule Cons(1))
    qed
  qed
  thus "snd b  0" and "punit.lt (snd b) adds t" and "(t - punit.lt (snd b))  fst b t u" by simp_all
qed

corollary find_sig_reducer_SomeD:
  assumes "find_sig_reducer (map spp_of bs) u t 0 = Some i"
  shows "i < length bs" and "rep_list (bs ! i)  0" and "punit.lt (rep_list (bs ! i)) adds t"
    and "(t - punit.lt (rep_list (bs ! i)))  lt (bs ! i) t u"
proof -
  from assms have "i - 0 < length (map spp_of bs)" by (rule find_sig_reducer_SomeD_aux)
  thus "i < length bs" by simp
  hence "spp_of (bs ! i) = (map spp_of bs) ! (i - 0)" by simp
  with assms have "snd (spp_of (bs ! i))  0" and "punit.lt (snd (spp_of (bs ! i))) adds t"
    and "(t - punit.lt (snd (spp_of (bs ! i))))  fst (spp_of (bs ! i)) t u"
    by (rule find_sig_reducer_SomeD')+
  thus "rep_list (bs ! i)  0" and "punit.lt (rep_list (bs ! i)) adds t"
    and "(t - punit.lt (rep_list (bs ! i)))  lt (bs ! i) t u" by (simp_all add: fst_spp_of snd_spp_of)
qed

lemma find_sig_reducer_NoneE:
  assumes "find_sig_reducer bs u t i = None" and "b  set bs"
  assumes "snd b = 0  thesis" and "snd b  0  ¬ punit.lt (snd b) adds t  thesis"
    and "snd b  0  punit.lt (snd b) adds t  ¬ (t - punit.lt (snd b))  fst b t u  thesis"
  shows thesis
  using assms
proof (induct bs arbitrary: thesis i)
  case Nil
  from Nil(2) show ?case by simp
next
  case (Cons a bs)
  from Cons(2) have 1: "snd a = 0  ¬ punit.lt (snd a) adds t  ¬ (t - punit.lt (snd a))  fst a t u"
    and eq: "find_sig_reducer bs u t (Suc i) = None" by (simp_all split: if_splits)
  from Cons(3) have "b = a  b  set bs" by simp
  thus ?case
  proof
    assume "b = a"
    show ?thesis
    proof (cases "snd a = 0")
      case True
      show ?thesis by (rule Cons(4), simp add: b = a True)
    next
      case False
      with 1 have 2: "¬ punit.lt (snd a) adds t  ¬ (t - punit.lt (snd a))  fst a t u" by simp
      show ?thesis
      proof (cases "punit.lt (snd a) adds t")
        case True
        with 2 have 3: "¬ (t - punit.lt (snd a))  fst a t u" by simp
        show ?thesis by (rule Cons(6), simp_all add: b = a snd a  0 True 3)
      next
        case False
        show ?thesis by (rule Cons(5), simp_all add: b = a snd a  0 False)
      qed
    qed
  next
    assume "b  set bs"
    with eq show ?thesis
    proof (rule Cons(1))
      assume "snd b = 0"
      thus ?thesis by (rule Cons(4))
    next
      assume "snd b  0" and "¬ punit.lt (snd b) adds t"
      thus ?thesis by (rule Cons(5))
    next
      assume "snd b  0" and "punit.lt (snd b) adds t" and "¬ (t - punit.lt (snd b))  fst b t u"
      thus ?thesis by (rule Cons(6))
    qed
  qed
qed

lemma find_sig_reducer_SomeD_red_single:
  assumes "t  keys (rep_list p)" and "find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i"
  shows "sig_red_single (≺t) (≼) p (p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
            (t - punit.lt (rep_list (bs ! i))) (bs ! i)) (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
proof -
  from assms(2) have "punit.lt (rep_list (bs ! i)) adds t" and 1: "rep_list (bs ! i)  0"
    and 2: "(t - punit.lt (rep_list (bs ! i)))  lt (bs ! i) t lt p"
    by (rule find_sig_reducer_SomeD)+
  from this(1) have eq: "t - punit.lt (rep_list (bs ! i)) + punit.lt (rep_list (bs ! i)) = t"
    by (rule adds_minus)
  from assms(1) have 3: "t  punit.lt (rep_list p)" by (rule punit.lt_max_keys)
  show ?thesis by (rule sig_red_singleI, simp_all add: eq 1 2 3 assms(1))
qed

corollary find_sig_reducer_SomeD_red:
  assumes "t  keys (rep_list p)" and "find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i"
  shows "sig_red (≺t) (≼) (set bs) p (p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
            (t - punit.lt (rep_list (bs ! i))) (bs ! i))"
  unfolding sig_red_def
proof (intro bexI exI, rule find_sig_reducer_SomeD_red_single)
  from assms(2) have "i - 0 < length (map spp_of bs)" by (rule find_sig_reducer_SomeD_aux)
  hence "i < length bs" by simp
  thus "bs ! i  set bs" by (rule nth_mem)
qed fact+

context
  fixes bs :: "('t 0 'b) list"
begin

definition sig_trd_term :: "('a  nat)  (('a × ('t 0 'b)) × ('a × ('t 0 'b))) set"
  where "sig_trd_term d = {(x, y). punit.dgrad_p_set_le d {rep_list (snd x)}
                                      (insert (rep_list (snd y)) (rep_list ` set bs)) 
                                 fst x  keys (rep_list (snd x))  fst y  keys (rep_list (snd y)) 
                                 fst x  fst y}"

lemma sig_trd_term_wf:
  assumes "dickson_grading d"
  shows "wf (sig_trd_term d)"
proof (rule wfI_min)
  fix x :: "'a × ('t 0 'b)" and Q
  assume "x  Q"
  show "zQ. y. (y, z)  sig_trd_term d  y  Q"
  proof (cases "fst x  keys (rep_list (snd x))")
    case True
    define X where "X = rep_list ` set bs"
    let ?A = "insert (rep_list (snd x)) X"
    have "finite X" unfolding X_def by simp
    hence "finite ?A" by (simp only: finite_insert)
    then obtain m where A: "?A  punit.dgrad_p_set d m" by (rule punit.dgrad_p_set_exhaust)
    hence x: "rep_list (snd x)  punit.dgrad_p_set d m" and X: "X  punit.dgrad_p_set d m"
      by simp_all
    let ?Q = "{q  Q. rep_list (snd q)  punit.dgrad_p_set d m  fst q  keys (rep_list (snd q))}"
    from x  Q x True have "x  ?Q" by simp
    have "Q x. x  Q  Q  {q. d q  m}  (zQ. y. y  z  y  Q)"
      by (rule wfp_on_imp_minimal, rule wfp_on_ord_strict, fact assms)
    hence 1: "fst x  fst ` ?Q  fst ` ?Q  {q. d q  m}  (zfst ` ?Q. y. y  z  y  fst ` ?Q)"
      by meson
  
    have "fst x  fst ` ?Q" by (rule, fact refl, fact)
    moreover have "fst ` ?Q  {q. d q  m}"
    proof -
      {
        fix q
        assume a: "rep_list (snd q)  punit.dgrad_p_set d m" and b: "fst q  keys (rep_list (snd q))"
        from a have "keys (rep_list (snd q))  dgrad_set d m" by (simp add: punit.dgrad_p_set_def)
        with b have "fst q  dgrad_set d m" ..
        hence "d (fst q)  m" by (simp add: dgrad_set_def)
      }
      thus ?thesis by auto
    qed
    ultimately have "zfst ` ?Q. y. y  z  y  fst ` ?Q" by (rule 1)
    then obtain z0 where "z0  fst ` ?Q" and 2: "y. y  z0  y  fst ` ?Q" by blast
    from this(1) obtain z where "z  ?Q" and z0: "z0 = fst z" ..
    hence "z  Q" and z: "rep_list (snd z)  punit.dgrad_p_set d m" by simp_all
    from this(1) show "zQ. y. (y, z)  sig_trd_term d  y  Q"
    proof
      show "y. (y, z)  sig_trd_term d  y  Q"
      proof (intro allI impI)
        fix y
        assume "(y, z)  sig_trd_term d"
        hence 3: "punit.dgrad_p_set_le d {rep_list (snd y)} (insert (rep_list (snd z)) X)"
          and 4: "fst y  keys (rep_list (snd y))" and "fst y  z0"
          by (simp_all add: sig_trd_term_def X_def z0)
        from this(3) have "fst y  fst ` ?Q" by (rule 2)
        hence "y  Q  rep_list (snd y)  punit.dgrad_p_set d m  fst y  keys (rep_list (snd y))"
          by auto
        thus "y  Q"
        proof (elim disjE)
          assume 5: "rep_list (snd y)  punit.dgrad_p_set d m"
          from z X have "insert (rep_list (snd z)) X  punit.dgrad_p_set d m" by simp
          with 3 have "{rep_list (snd y)}  punit.dgrad_p_set d m" by (rule punit.dgrad_p_set_le_dgrad_p_set)
          hence "rep_list (snd y)  punit.dgrad_p_set d m" by simp
          with 5 show ?thesis ..
        next
          assume "fst y  keys (rep_list (snd y))"
          thus ?thesis using 4 ..
        qed
      qed
    qed
  next
    case False
    from x  Q show ?thesis
    proof
      show "y. (y, x)  sig_trd_term d  y  Q"
      proof (intro allI impI)
        fix y
        assume "(y, x)  sig_trd_term d"
        hence "fst x  keys (rep_list (snd x))" by (simp add: sig_trd_term_def)
        with False show "y  Q" ..
      qed
    qed
  qed
qed

function (domintros) sig_trd_aux :: "('a × ('t 0 'b))  ('t 0 'b)" where
  "sig_trd_aux (t, p) =
    (let p' =
      (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
          None    p
        | Some i  p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                          (t - punit.lt (rep_list (bs ! i))) (bs ! i));
      p'' = punit.lower (rep_list p') t in
    if p'' = 0 then p' else sig_trd_aux (punit.lt p'', p'))"
  by auto

lemma sig_trd_aux_domI:
  assumes "fst args0  keys (rep_list (snd args0))"
  shows "sig_trd_aux_dom args0"
proof -
  from ex_hgrad obtain d::"'a  nat" where "dickson_grading d  hom_grading d" ..
  hence dg: "dickson_grading d" ..
  hence "wf (sig_trd_term d)" by (rule sig_trd_term_wf)
  thus ?thesis using assms
  proof (induct args0)
    case (less args)
    obtain t p where args: "args = (t, p)" using prod.exhaust by blast
    with less(1) have 1: "s q. ((s, q), (t, p))  sig_trd_term d  s  keys (rep_list q)  sig_trd_aux_dom (s, q)"
      using prod.exhaust by auto
    from less(2) have "t  keys (rep_list p)" by (simp add: args)
    show ?case unfolding args
    proof (rule sig_trd_aux.domintros)
      define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
                                None  p
                              | Some i  p -
                                  monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                   (t - punit.lt (rep_list (bs ! i))) (bs ! i))"
      define p'' where "p'' = punit.lower (rep_list p') t"
      assume "p''  0"
      from p''  0 have "punit.lt p''  keys p''" by (rule punit.lt_in_keys)
      also have "...  keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
      finally have "punit.lt p''  keys (rep_list p')" .
      with _ show "sig_trd_aux_dom (punit.lt p'', p')"
      proof (rule 1)
        have "punit.dgrad_p_set_le d {rep_list p'} (insert (rep_list p) (rep_list ` set bs))"
        proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
          case None
          hence "p' = p" by (simp add: p'_def)
          hence "{rep_list p'}  insert (rep_list p) (rep_list ` set bs)" by simp
          thus ?thesis by (rule punit.dgrad_p_set_le_subset)
        next
          case (Some i)
          hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                    (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
          have "sig_red (≺t) (≼) (set bs) p p'" unfolding p' using t  keys (rep_list p) Some
            by (rule find_sig_reducer_SomeD_red)
          hence "punit.red (rep_list ` set bs) (rep_list p) (rep_list p')" by (rule sig_red_red)
          with dg show ?thesis by (rule punit.dgrad_p_set_le_red)
        qed
        moreover note punit.lt p''  keys (rep_list p') t  keys (rep_list p)
        moreover from p''  0 have "punit.lt p''  t" unfolding p''_def by (rule punit.lt_lower_less)
        ultimately show "((punit.lt p'', p'), t, p)  sig_trd_term d" by (simp add: sig_trd_term_def)
      qed
    qed
  qed
qed

definition sig_trd :: "('t 0 'b)  ('t 0 'b)"
  where "sig_trd p = (if rep_list p = 0 then p else sig_trd_aux (punit.lt (rep_list p), p))"

lemma sig_trd_aux_red_rtrancl:
  assumes "fst args0  keys (rep_list (snd args0))"
  shows "(sig_red (≺t) (≼) (set bs))** (snd args0) (sig_trd_aux args0)"
proof -
  from assms have "sig_trd_aux_dom args0" by (rule sig_trd_aux_domI)
  thus ?thesis using assms
  proof (induct args0 rule: sig_trd_aux.pinduct)
    case (1 t p)
    define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
                              None  p
                            | Some i  p -
                                monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                 (t - punit.lt (rep_list (bs ! i))) (bs ! i))"
    define p'' where "p'' = punit.lower (rep_list p') t"
    from 1(3) have "t  keys (rep_list p)" by simp
    have *: "(sig_red (≺t) (≼) (set bs))** p p'"
    proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
      case None
      hence "p' = p" by (simp add: p'_def)
      thus ?thesis by simp
    next
      case (Some i)
      hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                  (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
      have "sig_red (≺t) (≼) (set bs) p p'" unfolding p' using t  keys (rep_list p) Some
        by (rule find_sig_reducer_SomeD_red)
      thus ?thesis ..
    qed
    show ?case
    proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric] *, intro impI)
      assume "p''  0"
      from * have "(sig_red (≺t) (≼) (set bs))** p (snd (punit.lt p'', p'))" by (simp only: snd_conv)
      moreover have "(sig_red (≺t) (≼) (set bs))** (snd (punit.lt p'', p')) (sig_trd_aux (punit.lt p'', p'))"
        using p'_def p''_def p''  0
      proof (rule 1(2))
        from p''  0 have "punit.lt p''  keys p''" by (rule punit.lt_in_keys)
        also have "...  keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
        finally show "fst (punit.lt p'', p')  keys (rep_list (snd (punit.lt p'', p')))" by simp
      qed
      ultimately show "(sig_red (≺t) (≼) (set bs))** p (sig_trd_aux (punit.lt p'', p'))"
        by (rule rtranclp_trans)
    qed
  qed
qed

corollary sig_trd_red_rtrancl: "(sig_red (≺t) (≼) (set bs))** p (sig_trd p)"
  unfolding sig_trd_def
proof (split if_split, intro conjI impI rtranclp.rtrancl_refl)
  let ?args = "(punit.lt (rep_list p), p)"
  assume "rep_list p  0"
  hence "punit.lt (rep_list p)  keys (rep_list p)" by (rule punit.lt_in_keys)
  hence "fst (punit.lt (rep_list p), p)  keys (rep_list (snd (punit.lt (rep_list p), p)))"
    by (simp only: fst_conv snd_conv)
  hence "(sig_red (≺t) (≼) (set bs))** (snd ?args) (sig_trd_aux ?args)" by (rule sig_trd_aux_red_rtrancl)
  thus "(sig_red (≺t) (≼) (set bs))** p (sig_trd_aux (punit.lt (rep_list p), p))" by (simp only: snd_conv)
qed

lemma sig_trd_aux_irred:
  assumes "fst args0  keys (rep_list (snd args0))"
    and "b s. b  set bs  rep_list b  0  fst args0  s + punit.lt (rep_list b) 
              s  lt b t lt (snd (args0))  lookup (rep_list (snd args0)) (s + punit.lt (rep_list b)) = 0"
  shows "¬ is_sig_red (≺t) (≼) (set bs) (sig_trd_aux args0)"
proof -
  from assms(1) have "sig_trd_aux_dom args0" by (rule sig_trd_aux_domI)
  thus ?thesis using assms
  proof (induct args0 rule: sig_trd_aux.pinduct)
    case (1 t p)
    define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
                              None  p
                            | Some i  p -
                                monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                 (t - punit.lt (rep_list (bs ! i))) (bs ! i))"
    define p'' where "p'' = punit.lower (rep_list p') t"
    from 1(3) have "t  keys (rep_list p)" by simp
    from 1(4) have a: "b  set bs  rep_list b  0  t  s + punit.lt (rep_list b) 
                       s  lt b t lt p  lookup (rep_list p) (s + punit.lt (rep_list b)) = 0"
      for b s by (simp only: fst_conv snd_conv)
    have "lt p' = lt p  (s. t  s  lookup (rep_list p') s = lookup (rep_list p) s)"
    proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
      case None
      thus ?thesis by (simp add: p'_def)
    next
      case (Some i)
      hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                  (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
      have "sig_red_single (≺t) (≼) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
        unfolding p' using t  keys (rep_list p) Some by (rule find_sig_reducer_SomeD_red_single)
      hence r: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
        and "lt p' = lt p" by (rule sig_red_single_red_single, rule sig_red_single_regular_lt)
      have "s. t  s  lookup (rep_list p') s = lookup (rep_list p) s"
      proof (intro allI impI)
        fix s
        assume "t  s"
        from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
        hence eq0: "(t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i)) = t" (is "?t = t")
          by (rule adds_minus)
        from t  s have "lookup (rep_list p') s = lookup (punit.higher (rep_list p') ?t) s"
          by (simp add: eq0 punit.lookup_higher_when)
        also from r have "... = lookup (punit.higher (rep_list p) ?t) s"
          by (simp add: punit.red_single_higher[simplified])
        also from t  s have "... = lookup (rep_list p) s" by (simp add: eq0 punit.lookup_higher_when)
        finally show "lookup (rep_list p') s = lookup (rep_list p) s" .
      qed
      with lt p' = lt p show ?thesis ..
    qed
    hence lt_p': "lt p' = lt p" and b: "s. t  s  lookup (rep_list p') s = lookup (rep_list p) s"
      by blast+
    have c: "lookup (rep_list p') (s + punit.lt (rep_list b)) = 0"
      if "b  set bs" and "rep_list b  0" and "t  s + punit.lt (rep_list b)" and "s  lt b t lt p'" for b s
    proof (cases "t  s + punit.lt (rep_list b)")
      case True
      hence "lookup (rep_list p') (s + punit.lt (rep_list b)) =
             lookup (rep_list p) (s + punit.lt (rep_list b))" by (rule b)
      also from that(1, 2) True that(4) have "... = 0" unfolding lt_p' by (rule a)
      finally show ?thesis .
    next
      case False
      with that(3) have t: "t = s + punit.lt (rep_list b)" by simp
      show ?thesis
      proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
        case None
        from that(1) have "spp_of b  set (map spp_of bs)" by fastforce
        with None show ?thesis
        proof (rule find_sig_reducer_NoneE)
          assume "snd (spp_of b) = 0"
          with that(2) show ?thesis by (simp add: snd_spp_of)
        next
          assume "¬ punit.lt (snd (spp_of b)) adds t"
          thus ?thesis by (simp add: snd_spp_of t)
        next
          assume "¬ (t - punit.lt (snd (spp_of b)))  fst (spp_of b) t lt p"
          with that(4) show ?thesis by (simp add: fst_spp_of snd_spp_of t lt_p')
        qed
      next
        case (Some i)
        hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                  (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
        have "sig_red_single (≺t) (≼) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
          unfolding p' using t  keys (rep_list p) Some by (rule find_sig_reducer_SomeD_red_single)
        hence r: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
          by (rule sig_red_single_red_single)
        from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
        hence eq0: "(t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i)) = t" (is "?t = t")
          by (rule adds_minus)
        from r have "lookup (rep_list p') ((t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i))) = 0"
          by (rule punit.red_single_lookup[simplified])
        thus ?thesis by (simp only: eq0 t[symmetric])
      qed
    qed
    show ?case
    proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric], intro conjI impI)
      assume "p'' = 0"
      show "¬ is_sig_red (≺t) (≼) (set bs) p'"
      proof
        assume "is_sig_red (≺t) (≼) (set bs) p'"
        then obtain b s where "b  set bs" and "s  keys (rep_list p')" and "rep_list b  0"
          and adds: "punit.lt (rep_list b) adds s" and "s  lt b t punit.lt (rep_list b)  lt p'"
          by (rule is_sig_red_addsE)
        let ?s = "s - punit.lt (rep_list b)"
        from adds have eq0: "?s + punit.lt (rep_list b) = s" by (simp add: adds_minus)
        show False
        proof (cases "t  s")
          case True
          note b  set bs rep_list b  0
          moreover from True have "t  ?s + punit.lt (rep_list b)" by (simp only: eq0)
          moreover from adds s  lt b t punit.lt (rep_list b)  lt p' have "?s  lt b t lt p'"
            by (simp add: term_is_le_rel_minus)
          ultimately have "lookup (rep_list p') (?s + punit.lt (rep_list b)) = 0" by (rule c)
          hence "s  keys (rep_list p')" by (simp add: eq0 in_keys_iff)
          thus ?thesis using s  keys (rep_list p') ..
        next
          case False
          hence "s  t" by simp
          hence "lookup (rep_list p') s = lookup (punit.lower (rep_list p') t) s"
            by (simp add: punit.lookup_lower_when)
          also from p'' = 0 have "... = 0" by (simp add: p''_def)
          finally have "s  keys (rep_list p')" by (simp add: in_keys_iff)
          thus ?thesis using s  keys (rep_list p') ..
        qed
      qed
    next
      assume "p''  0"
      with p'_def p''_def show "¬ is_sig_red (≺t) (≼) (set bs) (sig_trd_aux (punit.lt p'', p'))"
      proof (rule 1(2))
        from p''  0 have "punit.lt p''  keys p''" by (rule punit.lt_in_keys)
        also have "...  keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
        finally show "fst (punit.lt p'', p')  keys (rep_list (snd (punit.lt p'', p')))" by simp
      next
        fix b s
        assume "b  set bs" and "rep_list b  0"
        assume "fst (punit.lt p'', p')  s + punit.lt (rep_list b)"
          and "s  lt b t lt (snd (punit.lt p'', p'))"
        hence "punit.lt p''  s + punit.lt (rep_list b)" and "s  lt b t lt p'" by simp_all
        have "lookup (rep_list p') (s + punit.lt (rep_list b)) = 0"
        proof (cases "t  s + punit.lt (rep_list b)")
          case True
          with b  set bs rep_list b  0 show ?thesis using s  lt b t lt p' by (rule c)
        next
          case False
          hence "s + punit.lt (rep_list b)  t" by simp
          hence "lookup (rep_list p') (s + punit.lt (rep_list b)) =
                  lookup (punit.lower (rep_list p') t) (s + punit.lt (rep_list b))"
            by (simp add: punit.lookup_lower_when)
          also have "... = 0"
          proof (rule ccontr)
            assume "lookup (punit.lower (rep_list p') t) (s + punit.lt (rep_list b))  0"
            hence "s + punit.lt (rep_list b)  punit.lt (punit.lower (rep_list p') t)"
              by (rule punit.lt_max)
            also have "... = punit.lt p''" by (simp only: p''_def)
            finally show False using punit.lt p''  s + punit.lt (rep_list b) by simp
          qed
          finally show ?thesis .
        qed
        thus "lookup (rep_list (snd (punit.lt p'', p'))) (s + punit.lt (rep_list b)) = 0"
          by (simp only: snd_conv)
      qed
    qed
  qed
qed

corollary sig_trd_irred: "¬ is_sig_red (≺t) (≼) (set bs) (sig_trd p)"
  unfolding sig_trd_def
proof (split if_split, intro conjI impI)
  assume "rep_list p = 0"
  show "¬ is_sig_red (≺t) (≼) (set bs) p"
  proof
    assume "is_sig_red (≺t) (≼) (set bs) p"
    then obtain t where "t  keys (rep_list p)" by (rule is_sig_red_addsE)
    thus False by (simp add: rep_list p = 0)
  qed
next
  assume "rep_list p  0"
  show "¬ is_sig_red (≺t) (≼) (set bs) (sig_trd_aux (punit.lt (rep_list p), p))"
  proof (rule sig_trd_aux_irred)
    from rep_list p  0 have "punit.lt (rep_list p)  keys (rep_list p)" by (rule punit.lt_in_keys)
    thus "fst (punit.lt (rep_list p), p)  keys (rep_list (snd (punit.lt (rep_list p), p)))" by simp
  next
    fix b s
    assume "fst (punit.lt (rep_list p), p)  s + punit.lt (rep_list b)"
    thus "lookup (rep_list (snd (punit.lt (rep_list p), p))) (s + punit.lt (rep_list b)) = 0"
      using punit.lt_max by force
  qed
qed

end

context
  fixes bs :: "('t × ('a 0 'b)) list"
begin

context
  fixes v :: 't
begin

fun sig_trd_spp_body :: "(('a 0 'b) × ('a 0 'b))  (('a 0 'b) × ('a 0 'b))" where
  "sig_trd_spp_body (p, r) =
    (case find_sig_reducer bs v (punit.lt p) 0 of
        None    (punit.tail p, r + monomial (punit.lc p) (punit.lt p))
      | Some i  let b = snd (bs ! i) in
          (punit.tail p - punit.monom_mult (punit.lc p / punit.lc b) (punit.lt p - punit.lt b) (punit.tail b), r))"

definition sig_trd_spp_aux :: "(('a 0 'b) × ('a 0 'b))  ('a 0 'b)"
  where sig_trd_spp_aux_def [code del]: "sig_trd_spp_aux = tailrec.fun (λx. fst x = 0) snd sig_trd_spp_body"

lemma sig_trd_spp_aux_simps [code]:
  "sig_trd_spp_aux (p, r) = (if p = 0 then r else sig_trd_spp_aux (sig_trd_spp_body (p, r)))"
  by (simp add: sig_trd_spp_aux_def tailrec.simps)

end

fun sig_trd_spp :: "('t × ('a 0 'b))  ('t × ('a 0 'b))" where
  "sig_trd_spp (v, p) = (v, sig_trd_spp_aux v (p, 0))"

text ‹We define function @{const sig_trd_spp}, operating on sig-poly-pairs, already here, to have
  its definition in the right context. Lemmas are proved about it below in Section Sig-Poly-Pairs›.›

end

subsubsection ‹Koszul Syzygies›

text ‹A @{emph ‹Koszul syzygy›} of the list @{term fs} of scalar polynomials is a syzygy of the form
  @{term "(fs ! i)  (monomial 1 (term_of_pair (0, j))) - (fs ! j)  (monomial 1 (term_of_pair (0, i)))"},
  for @{prop "i < j"} and @{prop "j < length fs"}.›

primrec Koszul_syz_sigs_aux :: "('a 0 'b) list  nat  't list" where
  "Koszul_syz_sigs_aux [] i = []" |
  "Koszul_syz_sigs_aux (b # bs) i =
    map_idx (λb' j. ord_term_lin.max (term_of_pair (punit.lt b, j)) (term_of_pair (punit.lt b', i))) bs (Suc i) @
    Koszul_syz_sigs_aux bs (Suc i)"

definition Koszul_syz_sigs :: "('a 0 'b) list  't list"
  where "Koszul_syz_sigs bs = filter_min (addst) (Koszul_syz_sigs_aux bs 0)"

fun new_syz_sigs :: "'t list  ('t 0 'b) list  (('t 0 'b) × ('t 0 'b)) + nat  't list"
  where
    "new_syz_sigs ss bs (Inl (a, b)) = ss" |
    "new_syz_sigs ss bs (Inr j) =
      (if is_pot_ord then
        filter_min_append (addst) ss (filter_min (addst) (map (λb. term_of_pair (punit.lt (rep_list b), j)) bs))
      else ss)"

fun new_syz_sigs_spp :: "'t list  ('t × ('a 0 'b)) list  (('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat  't list"
  where
    "new_syz_sigs_spp ss bs (Inl (a, b)) = ss" |
    "new_syz_sigs_spp ss bs (Inr j) =
      (if is_pot_ord then
        filter_min_append (addst) ss (filter_min (addst) (map (λb. term_of_pair (punit.lt (snd b), j)) bs))
      else ss)"

lemma Koszul_syz_sigs_auxI:
  assumes "i < j" and "j < length bs"
  shows "ord_term_lin.max (term_of_pair (punit.lt (bs ! i), k + j)) (term_of_pair (punit.lt (bs ! j), k + i)) 
          set (Koszul_syz_sigs_aux bs k)"
  using assms
proof (induct bs arbitrary: i j k)
  case Nil
  from Nil(2) show ?case by simp
next
  case (Cons b bs)
  from Cons(2) obtain j0 where j: "j = Suc j0" by (meson lessE)
  from Cons(3) have "j0 < length bs" by (simp add: j)
  let ?A = "(λj. ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k))) (term_of_pair (punit.lt (bs ! j), k))) `
                {0..<length bs}"
  let ?B = "set (Koszul_syz_sigs_aux bs (Suc k))"
  show ?case
  proof (cases i)
    case 0
    from j0 < length bs have "j0  {0..<length bs}" by simp
    hence "ord_term_lin.max (term_of_pair (punit.lt b, Suc (j0 + k)))
                (term_of_pair (punit.lt (bs ! j0), k))  ?A" by (rule imageI)
    thus ?thesis by (simp add: i = 0 j set_map_idx ac_simps)
  next
    case (Suc i0)
    from Cons(2) have "i0 < j0" by (simp add: i = Suc i0 j)
    hence "ord_term_lin.max (term_of_pair (punit.lt (bs ! i0), Suc k + j0))
              (term_of_pair (punit.lt (bs ! j0), Suc k + i0))  ?B"
      using j0 < length bs by (rule Cons(1))
    thus ?thesis by (simp add: i = Suc i0 j set_map_idx ac_simps)
  qed
qed

lemma Koszul_syz_sigs_auxE:
  assumes "v  set (Koszul_syz_sigs_aux bs k)"
  obtains i j where "i < j" and "j < length bs"
    and "v = ord_term_lin.max (term_of_pair (punit.lt (bs ! i), k + j)) (term_of_pair (punit.lt (bs ! j), k + i))"
  using assms
proof (induct bs arbitrary: k thesis)
  case Nil
  from Nil(2) show ?case by simp
next
  case (Cons b bs)
  have "v  (λj. ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k))) (term_of_pair (punit.lt (bs ! j), k))) `
                {0..<length bs}  set (Koszul_syz_sigs_aux bs (Suc k))" (is "v  ?A  ?B")
    using Cons(3) by (simp add: set_map_idx)
  thus ?case
  proof
    assume "v  ?A"
    then obtain j where "j  {0..<length bs}"
      and v: "v = ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k)))
                                    (term_of_pair (punit.lt (bs ! j), k))" ..
    from this(1) have "j < length bs" by simp
    show ?thesis
    proof (rule Cons(2))
      show "0 < Suc j" by simp
    next
      from j < length bs show "Suc j < length (b # bs)" by simp
    next
      show "v = ord_term_lin.max (term_of_pair (punit.lt ((b # bs) ! 0), k + Suc j))
                                  (term_of_pair (punit.lt ((b # bs) ! Suc j), k + 0))"
        by (simp add: v ac_simps)
    qed
  next
    assume "v  ?B"
    obtain i j where "i < j" and "j < length bs"
      and v: "v = ord_term_lin.max (term_of_pair (punit.lt (bs ! i), Suc k + j))
                                    (term_of_pair (punit.lt (bs ! j), Suc k + i))"
      by (rule Cons(1), assumption, rule v  ?B)
    show ?thesis
    proof (rule Cons(2))
      from i < j show "Suc i < Suc j" by simp
    next
      from j < length bs show "Suc j < length (b # bs)" by simp
    next
      show "v = ord_term_lin.max (term_of_pair (punit.lt ((b # bs) ! Suc i), k + Suc j))
                                  (term_of_pair (punit.lt ((b # bs) ! Suc j), k + Suc i))"
        by (simp add: v)
    qed
  qed
qed

lemma lt_Koszul_syz_comp:
  assumes "0  set fs" and "i < length fs"
  shows "lt ((fs ! i)  monomial 1 (term_of_pair (0, j))) = term_of_pair (punit.lt (fs ! i), j)"
proof -
  from assms(2) have "fs ! i  set fs" by (rule nth_mem)
  with assms(1) have "fs ! i  0" by auto
  thus ?thesis by (simp add: lt_mult_scalar_monomial_right splus_def term_simps)
qed

lemma Koszul_syz_nonzero_lt:
  assumes "rep_list a  0" and "rep_list b  0" and "component_of_term (lt a) < component_of_term (lt b)"
  shows "rep_list a  b - rep_list b  a  0" (is "?p - ?q  0")
    and "lt (rep_list a  b - rep_list b  a) =
          ord_term_lin.max (punit.lt (rep_list a)  lt b) (punit.lt (rep_list b)  lt a)" (is "_ = ?r")
proof -
  from assms(2) have "b  0" by (auto simp: rep_list_zero)
  with assms(1) have lt_p: "lt ?p = punit.lt (rep_list a)  lt b" by (rule lt_mult_scalar)
  from assms(1) have "a  0" by (auto simp: rep_list_zero)
  with assms(2) have lt_q: "lt ?q = punit.lt (rep_list b)  lt a" by (rule lt_mult_scalar)
  from assms(3) have "component_of_term (lt ?p)  component_of_term (lt ?q)"
    by (simp add: lt_p lt_q component_of_term_splus)
  hence "lt ?p  lt ?q" by auto
  hence "lt (?p - ?q) = ord_term_lin.max (lt ?p) (lt ?q)" by (rule lt_minus_distinct_eq_max)
  also have "... = ?r" by (simp only: lt_p lt_q)
  finally show "lt (?p - ?q) = ?r" .
  
  from lt ?p  lt ?q show "?p - ?q  0" by auto
qed

lemma Koszul_syz_is_syz: "rep_list (rep_list a  b - rep_list b  a) = 0"
  by (simp add: rep_list_minus rep_list_mult_scalar)

lemma dgrad_sig_set_closed_Koszul_syz:
  assumes "dickson_grading dgrad" and "a  dgrad_sig_set dgrad" and "b  dgrad_sig_set dgrad"
  shows "rep_list a  b - rep_list b  a  dgrad_sig_set dgrad"
proof -
  from assms(2, 3) have 1: "a  dgrad_max_set dgrad" and 2: "b  dgrad_max_set dgrad"
    by (simp_all add: dgrad_sig_set'_def)
  show ?thesis
    by (intro dgrad_sig_set_closed_minus dgrad_sig_set_closed_mult_scalar dgrad_max_2 assms 1 2)
qed

corollary Koszul_syz_is_syz_sig:
  assumes "dickson_grading dgrad" and "a  dgrad_sig_set dgrad" and "b  dgrad_sig_set dgrad"
    and "rep_list a  0" and "rep_list b  0" and "component_of_term (lt a) < component_of_term (lt b)"
  shows "is_syz_sig dgrad (ord_term_lin.max (punit.lt (rep_list a)  lt b) (punit.lt (rep_list b)  lt a))"
proof (rule is_syz_sigI)
  from assms(4-6) show "rep_list a  b - rep_list b  a  0"
    and "lt (rep_list a  b - rep_list b  a) =
          ord_term_lin.max (punit.lt (rep_list a)  lt b) (punit.lt (rep_list b)  lt a)"
    by (rule Koszul_syz_nonzero_lt)+
next
  from assms(1-3) show "rep_list a  b - rep_list b  a  dgrad_sig_set dgrad"
    by (rule dgrad_sig_set_closed_Koszul_syz)
qed (fact Koszul_syz_is_syz)

corollary lt_Koszul_syz_in_Koszul_syz_sigs_aux:
  assumes "distinct fs" and "0  set fs" and "i < j" and "j < length fs"
  shows "lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i))) 
          set (Koszul_syz_sigs_aux fs 0)" (is "?l  ?K")
proof -
  let ?a = "monomial (1::'b) (term_of_pair (0, i))"
  let ?b = "monomial (1::'b) (term_of_pair (0, j))"
  from assms(3, 4) have "i < length fs" by simp
  with assms(1) have a: "rep_list ?a = fs ! i" by (simp add: rep_list_monomial term_simps)
  from assms(1, 4) have b: "rep_list ?b = fs ! j" by (simp add: rep_list_monomial term_simps)
  have "?l = lt (rep_list ?a  ?b - rep_list ?b  ?a)" by (simp only: a b)
  also have "... = ord_term_lin.max (punit.lt (rep_list ?a)  lt ?b) (punit.lt (rep_list ?b)  lt ?a)"
  proof (rule Koszul_syz_nonzero_lt)
    from i < length fs have "fs ! i  set fs" by (rule nth_mem)
    with assms(2) show "rep_list ?a  0" by (auto simp: a)
  next
    from assms(4) have "fs ! j  set fs" by (rule nth_mem)
    with assms(2) show "rep_list ?b  0" by (auto simp: b)
  next
    from assms(3) show "component_of_term (lt ?a) < component_of_term (lt ?b)"
      by (simp add: lt_monomial component_of_term_of_pair)
  qed
  also have "... = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), 0 + j)) (term_of_pair (punit.lt (fs ! j), 0 + i))"
    by (simp add: a b lt_monomial splus_def term_simps)
  also from assms(3, 4) have "...  ?K" by (rule Koszul_syz_sigs_auxI)
  thm Koszul_syz_sigs_auxI[OF assms(3, 4)]
  finally show ?thesis .
qed

corollary lt_Koszul_syz_in_Koszul_syz_sigs:
  assumes "¬ is_pot_ord" and "distinct fs" and "0  set fs" and "i < j" and "j < length fs"
  obtains v where "v  set (Koszul_syz_sigs fs)"
    and "v addst lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i)))"
proof -
  have "transp (addst)" by (rule transpI, drule adds_term_trans)
  moreover have "lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i))) 
                  set (Koszul_syz_sigs_aux fs 0)" (is "?l  set ?ks")
    using assms(2-5) by (rule lt_Koszul_syz_in_Koszul_syz_sigs_aux)
  ultimately show ?thesis
  proof (rule filter_min_cases)
    assume "?l  set (filter_min (addst) ?ks)"
    hence "?l  set (Koszul_syz_sigs fs)" by (simp add: Koszul_syz_sigs_def assms(1))
    thus ?thesis using adds_term_refl ..
  next
    fix v
    assume "v  set (filter_min (addst) ?ks)"
    hence "v  set (Koszul_syz_sigs fs)" by (simp add: Koszul_syz_sigs_def assms(1))
    moreover assume "v addst ?l"
    ultimately show ?thesis ..
  qed
qed

lemma lt_Koszul_syz_init:
  assumes "0  set fs" and "i < j" and "j < length fs"
  shows "lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i))) =
          ord_term_lin.max (term_of_pair (punit.lt (fs ! i), j)) (term_of_pair (punit.lt (fs ! j), i))"
            (is "lt (?p - ?q) = ?r")
proof -
  from assms(2, 3) have "i < length fs" by simp
  with assms(1) have lt_i: "lt ?p = term_of_pair (punit.lt (fs ! i), j)" by (rule lt_Koszul_syz_comp)
  from assms(1, 3) have lt_j: "lt ?q = term_of_pair (punit.lt (fs ! j), i)" by (rule lt_Koszul_syz_comp)
  from assms(2) have "component_of_term (lt ?p)  component_of_term (lt ?q)"
    by (simp add: lt_i lt_j component_of_term_of_pair)
  hence "lt ?p  lt ?q" by auto
  hence "lt (?p - ?q) = ord_term_lin.max (lt ?p) (lt ?q)" by (rule lt_minus_distinct_eq_max)
  also have "... = ?r" by (simp only: lt_i lt_j)
  finally show ?thesis .
qed

corollary Koszul_syz_sigs_auxE_lt_Koszul_syz:
  assumes "0  set fs" and "v  set (Koszul_syz_sigs_aux fs 0)"
  obtains i j where "i < j" and "j < length fs"
    and "v = lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i)))"
proof -
  from assms(2) obtain i j where "i < j" and "j < length fs"
    and "v = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), 0 + j))
                        (term_of_pair (punit.lt (fs ! j), 0 + i))"
    by (rule Koszul_syz_sigs_auxE)
  with assms(1) have "v = lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) -
                                (fs ! j)  monomial 1 (term_of_pair (0, i)))"
    by (simp add: lt_Koszul_syz_init)
  with i < j j < length fs show ?thesis ..
qed

corollary Koszul_syz_sigs_is_syz_sig:
  assumes "dickson_grading dgrad" and "distinct fs" and "0  set fs" and "v  set (Koszul_syz_sigs fs)"
  shows "is_syz_sig dgrad v"
proof -
  from assms(4) have "v  set (Koszul_syz_sigs_aux fs 0)"
    using filter_min_subset by (fastforce simp: Koszul_syz_sigs_def)
  with assms(3) obtain i j where "i < j" and "j < length fs"
    and v': "v = lt ((fs ! i)  monomial 1 (term_of_pair (0, j)) - (fs ! j)  monomial 1 (term_of_pair (0, i)))"
          (is "v = lt (?p - ?q)")
    by (rule Koszul_syz_sigs_auxE_lt_Koszul_syz)
  let ?a = "monomial (1::'b) (term_of_pair (0, i))"
  let ?b = "monomial (1::'b) (term_of_pair (0, j))"
  from i < j j < length fs have "i < length fs" by simp
  with assms(2) have a: "rep_list ?a = fs ! i" by (simp add: rep_list_monomial term_simps)
  from assms(2) j < length fs have b: "rep_list ?b = fs ! j" by (simp add: rep_list_monomial term_simps)
  note v'
  also have "lt (?p - ?q) = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), j)) (term_of_pair (punit.lt (fs ! j), i))"
    using assms(3) i < j j < length fs by (rule lt_Koszul_syz_init)
  also have "... = ord_term_lin.max (punit.lt (rep_list ?a)  lt ?b) (punit.lt (rep_list ?b)  lt ?a)"
    by (simp add: a b lt_monomial splus_def term_simps)
  finally have v: "v = ord_term_lin.max (punit.lt (rep_list ?a)  lt ?b) (punit.lt (rep_list ?b)  lt ?a)" .
  show ?thesis unfolding v using assms(1)
  proof (rule Koszul_syz_is_syz_sig)
    show "?a  dgrad_sig_set dgrad"
      by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 i < length fs)
  next
    show "?b  dgrad_sig_set dgrad"
      by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 j < length fs)
  next
    from i < length fs have "fs ! i  set fs" by (rule nth_mem)
    with assms(3) show "rep_list ?a  0" by (fastforce simp: a)
  next
    from j < length fs have "fs ! j  set fs" by (rule nth_mem)
    with assms(3) show "rep_list ?b  0" by (fastforce simp: b)
  next
    from i < j show "component_of_term (lt ?a) < component_of_term (lt ?b)"
      by (simp add: lt_monomial component_of_term_of_pair)
  qed
qed

lemma Koszul_syz_sigs_minimal:
  assumes "u  set (Koszul_syz_sigs fs)" and "v  set (Koszul_syz_sigs fs)" and "u addst v"
  shows "u = v"
proof -
  from assms(1, 2) have "u  set (filter_min (addst) (Koszul_syz_sigs_aux fs 0))"
    and "v  set (filter_min (addst) (Koszul_syz_sigs_aux fs 0))" by (simp_all add: Koszul_syz_sigs_def)
  with _ show ?thesis using assms(3)
  proof (rule filter_min_minimal)
    show "transp (addst)" by (rule transpI, drule adds_term_trans)
  qed
qed

lemma Koszul_syz_sigs_distinct: "distinct (Koszul_syz_sigs fs)"
proof -
  from adds_term_refl have "reflp (addst)" by (rule reflpI)
  thus ?thesis by (simp add: Koszul_syz_sigs_def filter_min_distinct)
qed

subsubsection ‹Algorithms›

definition spair_spp :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  ('t × ('a 0 'b))"
  where "spair_spp p q = (let t1 = punit.lt (snd p); t2 = punit.lt (snd q); l = lcs t1 t2 in
                          (ord_term_lin.max ((l - t1)  fst p) ((l - t2)  fst q),
                            punit.monom_mult (1 / punit.lc (snd p)) (l - t1) (snd p) -
                            punit.monom_mult (1 / punit.lc (snd q)) (l - t2) (snd q)))"

definition is_regular_spair_spp :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "is_regular_spair_spp p q 
                    (snd p  0  snd q  0  punit.lt (snd q)  fst p  punit.lt (snd p)  fst q)"

definition spair_sigs :: "('t 0 'b)  ('t 0 'b)  ('t × 't)"
  where "spair_sigs p q =
                  (let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
                    ((l - t1)  lt p, (l - t2)  lt q))"

definition spair_sigs_spp :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  ('t × 't)"
  where "spair_sigs_spp p q =
                  (let t1 = punit.lt (snd p); t2 = punit.lt (snd q); l = lcs t1 t2 in
                    ((l - t1)  fst p, (l - t2)  fst q))"

fun poly_of_pair :: "((('t 0 'b) × ('t 0 'b)) + nat)  ('t 0 'b)"
  where
    "poly_of_pair (Inl (p, q)) = spair p q" |
    "poly_of_pair (Inr j) = monomial 1 (term_of_pair (0, j))"

fun spp_of_pair :: "((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat)  ('t × ('a 0 'b))"
  where
    "spp_of_pair (Inl (p, q)) = spair_spp p q" |
    "spp_of_pair (Inr j) = (term_of_pair (0, j), fs ! j)"

fun sig_of_pair :: "((('t 0 'b) × ('t 0 'b)) + nat)  't"
  where
    "sig_of_pair (Inl (p, q)) = (let (u, v) = spair_sigs p q in ord_term_lin.max u v)" |
    "sig_of_pair (Inr j) = term_of_pair (0, j)"

fun sig_of_pair_spp :: "((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat)  't"
  where
    "sig_of_pair_spp (Inl (p, q)) = (let (u, v) = spair_sigs_spp p q in ord_term_lin.max u v)" |
    "sig_of_pair_spp (Inr j) = term_of_pair (0, j)"

definition pair_ord :: "((('t 0 'b) × ('t 0 'b)) + nat)  ((('t 0 'b) × ('t 0 'b)) + nat)  bool"
  where "pair_ord x y  (sig_of_pair x t sig_of_pair y)"

definition pair_ord_spp :: "((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) 
                            ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat)  bool"
  where "pair_ord_spp x y  (sig_of_pair_spp x t sig_of_pair_spp y)"

primrec new_spairs :: "('t 0 'b) list  ('t 0 'b)  ((('t 0 'b) × ('t 0 'b)) + nat) list" where
  "new_spairs [] p = []" |
  "new_spairs (b # bs) p =
    (if is_regular_spair p b then insort_wrt pair_ord (Inl (p, b)) (new_spairs bs p) else new_spairs bs p)"

primrec new_spairs_spp :: "('t × ('a 0 'b)) list  ('t × ('a 0 'b)) 
                            ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list" where
  "new_spairs_spp [] p = []" |
  "new_spairs_spp (b # bs) p =
      (if is_regular_spair_spp p b then
        insort_wrt pair_ord_spp (Inl (p, b)) (new_spairs_spp bs p)
      else new_spairs_spp bs p)"

definition add_spairs :: "((('t 0 'b) × ('t 0 'b)) + nat) list  ('t 0 'b) list  ('t 0 'b) 
                            ((('t 0 'b) × ('t 0 'b)) + nat) list"
  where "add_spairs ps bs p = merge_wrt pair_ord (new_spairs bs p) ps"

definition add_spairs_spp :: "((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list 
                              ('t × ('a 0 'b)) list  ('t × ('a 0 'b)) 
                              ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list"
  where "add_spairs_spp ps bs p = merge_wrt pair_ord_spp (new_spairs_spp bs p) ps"

lemma spair_alt_spair_sigs:
  "spair p q = monom_mult (1 / punit.lc (rep_list p)) (pp_of_term (fst (spair_sigs p q)) - lp p) p -
                monom_mult (1 / punit.lc (rep_list q)) (pp_of_term (snd (spair_sigs p q)) - lp q) q"
  by (simp add: spair_def spair_sigs_def Let_def term_simps)

lemma sig_of_spair:
  assumes "is_regular_spair p q"
  shows "sig_of_pair (Inl (p, q)) = lt (spair p q)"
proof -
  from assms have "rep_list p  0" by (rule is_regular_spairD1)
  hence 1: "punit.lc (rep_list p)  0" and "p  0" by (rule punit.lc_not_0, auto simp: rep_list_zero)
  from assms have "rep_list q  0" by (rule is_regular_spairD2)
  hence 2: "punit.lc (rep_list q)  0" and "q  0" by (rule punit.lc_not_0, auto simp: rep_list_zero)
  let ?t1 = "punit.lt (rep_list p)"
  let ?t2 = "punit.lt (rep_list q)"
  let ?l = "lcs ?t1 ?t2"
  from assms have "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p) 
                   lt (monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q)"
    by (rule is_regular_spairD3)
  hence *: "lt (monom_mult (1 / punit.lc (rep_list p)) (pp_of_term (fst (spair_sigs p q)) - lp p) p) 
            lt (monom_mult (1 / punit.lc (rep_list q)) (pp_of_term (snd (spair_sigs p q)) - lp q) q)"
    by (simp add: spair_sigs_def Let_def term_simps)
  from 1 2 p  0 q  0 show ?thesis
    by (simp add: spair_alt_spair_sigs lt_monom_mult lt_minus_distinct_eq_max[OF *],
        simp add: spair_sigs_def Let_def term_simps)
qed

lemma sig_of_spair_commute: "sig_of_pair (Inl (p, q)) = sig_of_pair (Inl (q, p))"
  by (simp add: spair_sigs_def Let_def lcs_comm ord_term_lin.max.commute)

lemma in_new_spairsI:
  assumes "b  set bs" and "is_regular_spair p b"
  shows "Inl (p, b)  set (new_spairs bs p)"
  using assms(1)
proof (induct bs)
  case Nil
  thus ?case by simp
next
  case (Cons a bs)
  from Cons(2) have "b = a  b  set bs" by simp
  thus ?case
  proof
    assume "b = a"
    from assms(2) show ?thesis by (simp add: b = a)
  next
    assume "b  set bs"
    hence "Inl (p, b)  set (new_spairs bs p)" by (rule Cons(1))
    thus ?thesis by simp
  qed
qed

lemma in_new_spairsD:
  assumes "Inl (a, b)  set (new_spairs bs p)"
  shows "a = p" and "b  set bs" and "is_regular_spair p b"
proof -
  from assms have "a = p  b  set bs  is_regular_spair p b"
  proof (induct bs)
  case Nil
  thus ?case by simp
  next
    case (Cons c bs)
    from Cons(2) have "(is_regular_spair p c  Inl (a, b) = Inl (p, c))  Inl (a, b)  set (new_spairs bs p)"
      by (simp split: if_split_asm)
    thus ?case
    proof
      assume "is_regular_spair p c  Inl (a, b) = Inl (p, c)"
      hence "is_regular_spair p c" and "a = p" and "b = c" by simp_all
      thus ?thesis by simp
    next
      assume "Inl (a, b)  set (new_spairs bs p)"
      hence "a = p  b  set bs  is_regular_spair p b" by (rule Cons(1))
      thus ?thesis by simp
    qed
  qed
  thus "a = p" and "b  set bs" and "is_regular_spair p b" by simp_all
qed

corollary in_new_spairs_iff:
  "Inl (p, b)  set (new_spairs bs p)  (b  set bs  is_regular_spair p b)"
  by (auto intro: in_new_spairsI dest: in_new_spairsD)

lemma Inr_not_in_new_spairs: "Inr j  set (new_spairs bs p)"
  by (induct bs, simp_all)

lemma sum_prodE:
  assumes "a b. p = Inl (a, b)  thesis" and "j. p = Inr j  thesis"
  shows thesis
  using _ assms(2)
proof (rule sumE)
  fix x
  assume "p = Inl x"
  moreover obtain a b where "x = (a, b)" by fastforce
  ultimately have "p = Inl (a, b)" by simp
  thus ?thesis by (rule assms(1))
qed

corollary in_new_spairsE:
  assumes "q  set (new_spairs bs p)"
  obtains b where "b  set bs" and "is_regular_spair p b" and "q = Inl (p, b)"
proof (rule sum_prodE)
  fix a b
  assume q: "q = Inl (a, b)"
  from assms have "a = p" and "b  set bs" and "is_regular_spair p b"
    unfolding q by (rule in_new_spairsD)+
  note this(2, 3)
  moreover have "q = Inl (p, b)" by (simp only: q a = p)
  ultimately show ?thesis ..
next
  fix j
  assume "q = Inr j"
  with assms show ?thesis by (simp add: Inr_not_in_new_spairs)
qed

lemma new_spairs_sorted: "sorted_wrt pair_ord (new_spairs bs p)"
proof (induct bs)
  case Nil
  show ?case by simp
next
  case (Cons a bs)
  moreover have "transp pair_ord" by (rule transpI, simp add: pair_ord_def)
  moreover have "pair_ord x y  pair_ord y x" for x y by (simp add: pair_ord_def ord_term_lin.linear)
  ultimately show ?case by (simp add: sorted_wrt_insort_wrt)
qed

lemma sorted_add_spairs:
  assumes "sorted_wrt pair_ord ps"
  shows "sorted_wrt pair_ord (add_spairs ps bs p)"
  unfolding add_spairs_def using _ _ new_spairs_sorted assms
proof (rule sorted_merge_wrt)
  show "transp pair_ord" by (rule transpI, simp add: pair_ord_def)
next
  fix x y
  show "pair_ord x y  pair_ord y x" by (simp add: pair_ord_def ord_term_lin.linear)
qed

context
  fixes rword_strict :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"   ―‹Must be a @{emph ‹strict›} rewrite order.›
begin

qualified definition rword :: "('t × ('a 0 'b))  ('t × ('a 0 'b))  bool"
  where "rword x y  ¬ rword_strict y x"

definition is_pred_syz :: "'t list  't  bool"
  where "is_pred_syz ss u = (vset ss. v addst u)"

definition is_rewritable :: "('t 0 'b) list  ('t 0 'b)  't  bool"
  where "is_rewritable bs p u = (bset bs. b  0  lt b addst u  rword_strict (spp_of p) (spp_of b))"

definition is_rewritable_spp :: "('t × ('a 0 'b)) list  ('t × ('a 0 'b))  't  bool"
  where "is_rewritable_spp bs p u = (bset bs. fst b addst u  rword_strict p b)"

fun sig_crit :: "('t 0 'b) list  't list  ((('t 0 'b) × ('t 0 'b)) + nat)  bool"
  where
    "sig_crit bs ss (Inl (p, q)) =
      (let (u, v) = spair_sigs p q in
        is_pred_syz ss u  is_pred_syz ss v  is_rewritable bs p u  is_rewritable bs q v)" |
    "sig_crit bs ss (Inr j) = is_pred_syz ss (term_of_pair (0, j))"

fun sig_crit' :: "('t 0 'b) list  ((('t 0 'b) × ('t 0 'b)) + nat)  bool"
  where
    "sig_crit' bs (Inl (p, q)) =
      (let (u, v) = spair_sigs p q in
        is_syz_sig dgrad u  is_syz_sig dgrad v  is_rewritable bs p u  is_rewritable bs q v)" |
    "sig_crit' bs (Inr j) = is_syz_sig dgrad (term_of_pair (0, j))"

fun sig_crit_spp :: "('t × ('a 0 'b)) list  't list  ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat)  bool"
  where
    "sig_crit_spp bs ss (Inl (p, q)) =
      (let (u, v) = spair_sigs_spp p q in
        is_pred_syz ss u  is_pred_syz ss v  is_rewritable_spp bs p u  is_rewritable_spp bs q v)" |
    "sig_crit_spp bs ss (Inr j) = is_pred_syz ss (term_of_pair (0, j))"

text @{const sig_crit} is used in algorithms, @{const sig_crit'} is only needed for proving.›

fun rb_spp_body ::
      "((('t × ('a 0 'b)) list × 't list × ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list) × nat) 
       ((('t × ('a 0 'b)) list × 't list × ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list) × nat)"
  where
  "rb_spp_body ((bs, ss, []), z) = ((bs, ss, []), z)" |
  "rb_spp_body ((bs, ss, p # ps), z) =
    (let ss' = new_syz_sigs_spp ss bs p in
      if sig_crit_spp bs ss' p then
          ((bs, ss', ps), z)
       else
          let p' = sig_trd_spp bs (spp_of_pair p) in
          if snd p' = 0 then
            ((bs, fst p' # ss', ps), Suc z)
          else
            ((p' # bs, ss', add_spairs_spp ps bs p'), z))"

definition rb_spp_aux ::
      "((('t × ('a 0 'b)) list × 't list × ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list) × nat) 
       ((('t × ('a 0 'b)) list × 't list × ((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat) list) × nat)"
  where rb_spp_aux_def [code del]: "rb_spp_aux = tailrec.fun (λx. snd (snd (fst x)) = []) (λx. x) rb_spp_body"

lemma rb_spp_aux_Nil [code]: "rb_spp_aux ((bs, ss, []), z) = ((bs, ss, []), z)"
  by (simp add: rb_spp_aux_def tailrec.simps)

lemma rb_spp_aux_Cons [code]:
  "rb_spp_aux ((bs, ss, p # ps), z) = rb_spp_aux (rb_spp_body ((bs, ss, p # ps), z))"
  by (simp add: rb_spp_aux_def tailrec.simps)

text ‹The last parameter / return value of @{const rb_spp_aux}, @{term z}, counts the number of
  zero-reductions. Below we will prove that this number remains $0$ under certain conditions.›

context
  assumes rword_is_strict_rewrite_ord: "is_strict_rewrite_ord rword_strict"
  assumes dgrad: "dickson_grading dgrad"
begin

lemma rword: "is_rewrite_ord rword"
  unfolding rword_def using rword_is_strict_rewrite_ord by (rule is_strict_rewrite_ordD)

lemma sig_crit'_sym: "sig_crit' bs (Inl (p, q))  sig_crit' bs (Inl (q, p))"
  by (auto simp: spair_sigs_def Let_def lcs_comm)

lemma is_rewritable_ConsD:
  assumes "is_rewritable (b # bs) p u" and "u t lt b"
  shows "is_rewritable bs p u"
proof -
  from assms(1) obtain b' where "b'  set (b # bs)" and "b'  0" and "lt b' addst u"
    and "rword_strict (spp_of p) (spp_of b')" unfolding is_rewritable_def by blast
  from this(3) have "lt b' t u" by (rule ord_adds_term)
  with assms(2) have "b'  b" by auto
  with b'  set (b # bs) have "b'  set bs" by simp
  with b'  0 lt b' addst u rword_strict (spp_of p) (spp_of b') show ?thesis
    by (auto simp: is_rewritable_def)
qed

lemma sig_crit'_ConsD:
  assumes "sig_crit' (b # bs) p" and "sig_of_pair p t lt b"
  shows "sig_crit' bs p"
proof (rule sum_prodE)
  fix x y
  assume p: "p = Inl (x, y)"
  define u where "u = fst (spair_sigs x y)"
  define v where "v = snd (spair_sigs x y)"
  have sigs: "spair_sigs x y = (u, v)" by (simp add: u_def v_def)
  have "u t sig_of_pair p" and "v t sig_of_pair p" by (simp_all add: p sigs)
  hence "u t lt b" and "v t lt b" using assms(2) by simp_all
  with assms(1) show ?thesis by (auto simp: p sigs dest: is_rewritable_ConsD)
next
  fix j
  assume p: "p = Inr j"
  from assms show ?thesis by (simp add: p)
qed

definition rb_aux_inv1 :: "('t 0 'b) list  bool"
  where "rb_aux_inv1 bs =
               (set bs  dgrad_sig_set dgrad  0  rep_list ` set bs 
                sorted_wrt (λx y. lt y t lt x) bs 
                (i<length bs. ¬ is_sig_red (≺t) (≼) (set (drop (Suc i) bs)) (bs ! i)) 
                (i<length bs.
    ((j<length fs. lt (bs ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) 
            punit.lt (rep_list (bs ! i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
    (pset bs. qset bs. is_regular_spair p q  rep_list (spair p q)  0 
            lt (bs ! i) = lt (spair p q)  punit.lt (rep_list (bs ! i))  punit.lt (rep_list (spair p q))))) 
                (i<length bs. is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))))"

fun rb_aux_inv :: "(('t 0 'b) list × 't list × ((('t 0 'b) × ('t 0 'b)) + nat) list)  bool"
  where "rb_aux_inv (bs, ss, ps) =
          (rb_aux_inv1 bs 
            (uset ss. is_syz_sig dgrad u) 
            (p q. Inl (p, q)  set ps  (is_regular_spair p q  p  set bs  q  set bs)) 
            (j. Inr j  set ps  (j < length fs  (bset bs. lt b t term_of_pair (0, j))) 
                              length (filter (λq. sig_of_pair q = term_of_pair (0, j)) ps)  1) 
            (sorted_wrt pair_ord ps) 
            (pset ps. (b1set bs. b2set bs. is_regular_spair b1 b2 
                          sig_of_pair p t lt (spair b1 b2)  (Inl (b1, b2)  set ps  Inl (b2, b1)  set ps)) 
                        (j<length fs. sig_of_pair p t term_of_pair (0, j)  Inr j  set ps)) 
            (bset bs. pset ps. lt b t sig_of_pair p) 
            (aset bs. bset bs. is_regular_spair a b  Inl (a, b)  set ps  Inl (b, a)  set ps 
                ¬ is_RB_in dgrad rword (set bs) (lt (spair a b)) 
                (pset ps. sig_of_pair p = lt (spair a b)  ¬ sig_crit' bs p)) 
            (j<length fs. Inr j  set ps  (is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) 
                rep_list (monomial (1::'b) (term_of_pair (0, j)))  ideal (rep_list ` set bs))))"

lemmas [simp del] = rb_aux_inv.simps

lemma rb_aux_inv1_D1: "rb_aux_inv1 bs  set bs  dgrad_sig_set dgrad"
  by (simp add: rb_aux_inv1_def)

lemma rb_aux_inv1_D2: "rb_aux_inv1 bs  0  rep_list ` set bs"
  by (simp add: rb_aux_inv1_def)

lemma rb_aux_inv1_D3: "rb_aux_inv1 bs  sorted_wrt (λx y. lt y t lt x) bs"
  by (simp add: rb_aux_inv1_def)

lemma rb_aux_inv1_D4:
  "rb_aux_inv1 bs  i < length bs  ¬ is_sig_red (≺t) (≼) (set (drop (Suc i) bs)) (bs ! i)"
  by (simp add: rb_aux_inv1_def)

lemma rb_aux_inv1_D5:
  "rb_aux_inv1 bs  i < length bs  is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))"
  by (simp add: rb_aux_inv1_def)

lemma rb_aux_inv1_E:
  assumes "rb_aux_inv1 bs" and "i < length bs"
    and "j. j < length fs  lt (bs ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) 
            punit.lt (rep_list (bs ! i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))  thesis"
    and "p q. p  set bs  q  set bs  is_regular_spair p q  rep_list (spair p q)  0 
            lt (bs ! i) = lt (spair p q)  punit.lt (rep_list (bs ! i))  punit.lt (rep_list (spair p q))  thesis"
  shows thesis
  using assms unfolding rb_aux_inv1_def by blast

lemmas rb_aux_inv1_D = rb_aux_inv1_D1 rb_aux_inv1_D2 rb_aux_inv1_D3 rb_aux_inv1_D4
                           rb_aux_inv1_D5

lemma rb_aux_inv1_distinct_lt:
  assumes "rb_aux_inv1 bs"
  shows "distinct (map lt bs)"
proof (rule distinct_sorted_wrt_irrefl)
  show "irreflp (≻t)" by (simp add: irreflp_def)
next
  show "transp (≻t)" by (auto simp: transp_def)
next
  from assms show "sorted_wrt (≻t) (map lt bs)"
    unfolding sorted_wrt_map conversep_iff by (rule rb_aux_inv1_D3)
qed

corollary rb_aux_inv1_lt_inj_on:
  assumes "rb_aux_inv1 bs"
  shows "inj_on lt (set bs)"
proof
  fix a b
  assume "a  set bs"
  then obtain i where i: "i < length bs" and a: "a = bs ! i" by (metis in_set_conv_nth)
  assume "b  set bs"
  then obtain j where j: "j < length bs" and b: "b = bs ! j" by (metis in_set_conv_nth)
  assume "lt a = lt b"
  with i j have "(map lt bs) ! i = (map lt bs) ! j" by (simp add: a b)
  moreover from assms have "distinct (map lt bs)" by (rule rb_aux_inv1_distinct_lt)
  moreover from i have "i < length (map lt bs)" by simp
  moreover from j have "j < length (map lt bs)" by simp
  ultimately have "i = j" by (simp only: nth_eq_iff_index_eq)
  thus "a = b" by (simp add: a b)
qed

lemma canon_rewriter_unique:
  assumes "rb_aux_inv1 bs" and "is_canon_rewriter rword (set bs) u a"
    and "is_canon_rewriter rword (set bs) u b"
  shows "a = b"
proof -
  from assms(1) have "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
  moreover from rword(1) assms(2, 3) have "lt a = lt b" by (rule is_rewrite_ord_canon_rewriterD2)
  moreover from assms(2) have "a  set bs" by (rule is_canon_rewriterD1)
  moreover from assms(3) have "b  set bs" by (rule is_canon_rewriterD1)
  ultimately show ?thesis by (rule inj_onD)
qed

lemma rb_aux_inv_D1: "rb_aux_inv (bs, ss, ps)  rb_aux_inv1 bs"
  by (simp add: rb_aux_inv.simps)

lemma rb_aux_inv_D2: "rb_aux_inv (bs, ss, ps)  u  set ss  is_syz_sig dgrad u"
  by (simp add: rb_aux_inv.simps)

lemma rb_aux_inv_D3:
  assumes "rb_aux_inv (bs, ss, ps)" and "Inl (p, q)  set ps"
  shows "p  set bs" and "q  set bs" and "is_regular_spair p q"
  using assms by (simp_all add: rb_aux_inv.simps)

lemma rb_aux_inv_D4:
  assumes "rb_aux_inv (bs, ss, ps)" and "Inr j  set ps"
  shows "j < length fs" and "b. b  set bs  lt b t term_of_pair (0, j)"
    and "length (filter (λq. sig_of_pair q = term_of_pair (0, j)) ps)  1"
  using assms by (simp_all add: rb_aux_inv.simps)

lemma rb_aux_inv_D5: "rb_aux_inv (bs, ss, ps)  sorted_wrt pair_ord ps"
  by (simp add: rb_aux_inv.simps)

lemma rb_aux_inv_D6_1:
  assumes "rb_aux_inv (bs, ss, ps)" and "p  set ps" and "b1  set bs" and "b2  set bs"
    and "is_regular_spair b1 b2" and "sig_of_pair p t lt (spair b1 b2)"
  obtains "Inl (b1, b2)  set ps" | "Inl (b2, b1)  set ps"
  using assms unfolding rb_aux_inv.simps by blast

lemma rb_aux_inv_D6_2:
  "rb_aux_inv (bs, ss, ps)  p  set ps  j < length fs  sig_of_pair p t term_of_pair (0, j) 
    Inr j  set ps"
  by (simp add: rb_aux_inv.simps)

lemma rb_aux_inv_D7: "rb_aux_inv (bs, ss, ps)  b  set bs  p  set ps  lt b t sig_of_pair p"
  by (simp add: rb_aux_inv.simps)

lemma rb_aux_inv_D8:
  assumes "rb_aux_inv (bs, ss, ps)" and "a  set bs" and "b  set bs" and "is_regular_spair a b"
    and "Inl (a, b)  set ps" and "Inl (b, a)  set ps" and "¬ is_RB_in dgrad rword (set bs) (lt (spair a b))"
  obtains p where "p  set ps" and "sig_of_pair p = lt (spair a b)" and "¬ sig_crit' bs p"
  using assms unfolding rb_aux_inv.simps by meson

lemma rb_aux_inv_D9:
  assumes "rb_aux_inv (bs, ss, ps)" and "j < length fs" and "Inr j  set ps"
  shows "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
    and "rep_list (monomial (1::'b) (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
  using assms by (simp_all add: rb_aux_inv.simps)

lemma rb_aux_inv_is_RB_upt:
  assumes "rb_aux_inv (bs, ss, ps)" and "p. p  set ps  u t sig_of_pair p"
  shows "is_RB_upt dgrad rword (set bs) u"
proof -
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  from dgrad rword(1) show ?thesis
  proof (rule is_RB_upt_finite)
    from inv1 show "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  next
    from inv1 show "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
  next
    show "finite (set bs)" by (fact finite_set)
  next
    fix g1 g2
    assume 1: "g1  set bs" and 2: "g2  set bs" and 3: "is_regular_spair g1 g2"
      and 4: "lt (spair g1 g2) t u"
    have 5: "p  set ps" if "sig_of_pair p = lt (spair g1 g2)" for p
    proof
      assume "p  set ps"
      hence "u t sig_of_pair p" by (rule assms(2))
      also have "... t u" unfolding that by (fact 4)
      finally show False ..
    qed
    show "is_RB_in dgrad rword (set bs) (lt (spair g1 g2))"
    proof (rule ccontr)
      note assms(1) 1 2 3
      moreover have "Inl (g1, g2)  set ps" by (rule 5, rule sig_of_spair, fact 3)
      moreover have "Inl (g2, g1)  set ps"
        by (rule 5, simp only: sig_of_spair_commute, rule sig_of_spair, fact 3)
      moreover assume "¬ is_RB_in dgrad rword (set bs) (lt (spair g1 g2))"
      ultimately obtain p where "p  set ps" and "sig_of_pair p = lt (spair g1 g2)"
        by (rule rb_aux_inv_D8)
      from this(2) have "p  set ps" by (rule 5)
      thus False using p  set ps ..
    qed
  next
    fix j
    assume 1: "term_of_pair (0, j) t u"
    note assms(1)
    moreover assume "j < length fs"
    moreover have "Inr j  set ps"
    proof
      assume "Inr j  set ps"
      hence "u t sig_of_pair (Inr j)" by (rule assms(2))
      also have "... t u" by (simp add: 1)
      finally show False ..
    qed
    ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (rule rb_aux_inv_D9)
  qed
qed

lemma rb_aux_inv_is_RB_upt_Cons:
  assumes "rb_aux_inv (bs, ss, p # ps)"
  shows "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
  using assms
proof (rule rb_aux_inv_is_RB_upt)
  fix q
  assume "q  set (p # ps)"
  hence "q = p  q  set ps" by simp
  thus "sig_of_pair p t sig_of_pair q"
  proof
    assume "q = p"
    thus ?thesis by simp
  next
    assume "q  set ps"
    moreover from assms have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    ultimately show ?thesis by (simp add: pair_ord_def)
  qed
qed

lemma Inr_in_tailD:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "Inr j  set ps"
  shows "sig_of_pair p  term_of_pair (0, j)"
proof
  assume eq: "sig_of_pair p = term_of_pair (0, j)"
  from assms(2) have "Inr j  set (p # ps)" by simp
  let ?P = "λq. sig_of_pair q = term_of_pair (0, j)"
  from assms(2) obtain i1 where "i1 < length ps" and Inrj: "Inr j = ps ! i1"
    by (metis in_set_conv_nth)
  from assms(1) Inr j  set (p # ps) have "length (filter ?P (p # ps))  1"
    by (rule rb_aux_inv_D4)
  moreover from i1 < length ps have "Suc i1 < length (p # ps)" by simp
  moreover have "0 < length (p # ps)" by simp
  moreover have "?P ((p # ps) ! Suc i1)" by (simp add: Inrj[symmetric])
  moreover have "?P ((p # ps) ! 0)" by (simp add: eq)
  ultimately have "Suc i1 = 0" by (rule length_filter_le_1)
  thus False ..
qed

lemma pair_list_aux:
  assumes "rb_aux_inv (bs, ss, ps)" and "p  set ps"
  shows "sig_of_pair p = lt (poly_of_pair p)  poly_of_pair p  0  poly_of_pair p  dgrad_sig_set dgrad"
proof (rule sum_prodE)
  fix a b
  assume p: "p = Inl (a, b)"
  from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence bs_sub: "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  from assms have "is_regular_spair a b" unfolding p by (rule rb_aux_inv_D3)
  hence "sig_of_pair p = lt (poly_of_pair p)" and "poly_of_pair p  0"
    unfolding p poly_of_pair.simps by (rule sig_of_spair, rule is_regular_spair_nonzero)
  moreover from dgrad have "poly_of_pair p  dgrad_sig_set dgrad" unfolding p poly_of_pair.simps
  proof (rule dgrad_sig_set_closed_spair)
    from assms have "a  set bs" unfolding p by (rule rb_aux_inv_D3)
    thus "a  dgrad_sig_set dgrad" using bs_sub ..
  next
    from assms have "b  set bs" unfolding p by (rule rb_aux_inv_D3)
    thus "b  dgrad_sig_set dgrad" using bs_sub ..
  qed
  ultimately show ?thesis by simp
next
  fix j
  assume "p = Inr j"
  from assms have "j < length fs" unfolding p = Inr j by (rule rb_aux_inv_D4)
  have "monomial 1 (term_of_pair (0, j))  dgrad_sig_set dgrad"
    by (rule dgrad_sig_set_closed_monomial, simp add: pp_of_term_of_pair dgrad_max_0,
        simp add: component_of_term_of_pair j < length fs)
  thus ?thesis by (simp add: p = Inr j lt_monomial monomial_0_iff)
qed

corollary pair_list_sig_of_pair:
  "rb_aux_inv (bs, ss, ps)  p  set ps  sig_of_pair p = lt (poly_of_pair p)"
  by (simp add: pair_list_aux)

corollary pair_list_nonzero: "rb_aux_inv (bs, ss, ps)  p  set ps  poly_of_pair p  0"
  by (simp add: pair_list_aux)

corollary pair_list_dgrad_sig_set:
  "rb_aux_inv (bs, ss, ps)  p  set ps  poly_of_pair p  dgrad_sig_set dgrad"
  by (simp add: pair_list_aux)

lemma is_rewritableI_is_canon_rewriter:
  assumes "rb_aux_inv1 bs" and "b  set bs" and "b  0" and "lt b addst u"
    and "¬ is_canon_rewriter rword (set bs) u b"
  shows "is_rewritable bs b u"
proof -
  from assms(2-5) obtain b' where "b'  set bs" and "b'  0" and "lt b' addst u"
    and 1: "¬ rword (spp_of b') (spp_of b)" by (auto simp: is_canon_rewriter_def)
  show ?thesis unfolding is_rewritable_def
  proof (intro bexI conjI)
    from rword(1) have 2: "rword (spp_of b) (spp_of b')"
    proof (rule is_rewrite_ordD3)
      assume "rword (spp_of b') (spp_of b)"
      with 1 show ?thesis ..
    qed
    from rword(1) 1 have "b  b'" by (auto dest: is_rewrite_ordD1)
    have "lt b  lt b'"
    proof
      assume "lt b = lt b'"
      with rb_aux_inv1_lt_inj_on[OF assms(1)] have "b = b'" using assms(2) b'  set bs
        by (rule inj_onD)
      with b  b' show False ..
    qed
    hence "fst (spp_of b)  fst (spp_of b')" by (simp add: spp_of_def)
    with rword_is_strict_rewrite_ord 2 show "rword_strict (spp_of b) (spp_of b')"
      by (auto simp: rword_def dest: is_strict_rewrite_ord_antisym)
  qed fact+
qed

lemma is_rewritableD_is_canon_rewriter:
  assumes "rb_aux_inv1 bs" and "is_rewritable bs b u"
  shows "¬ is_canon_rewriter rword (set bs) u b"
proof
  assume "is_canon_rewriter rword (set bs) u b"
  hence "b  set bs" and "b  0" and "lt b addst u"
    and 1: "a. a  set bs  a  0  lt a addst u  rword (spp_of a) (spp_of b)"
    by (rule is_canon_rewriterD)+
  from assms(2) obtain b' where "b'  set bs" and "b'  0" and "lt b' addst u"
    and 2: "rword_strict (spp_of b) (spp_of b')" unfolding is_rewritable_def by blast
  from this(1, 2, 3) have "rword (spp_of b') (spp_of b)" by (rule 1)
  moreover from rword_is_strict_rewrite_ord 2 have "rword (spp_of b) (spp_of b')"
    unfolding rword_def by (rule is_strict_rewrite_ord_asym)
  ultimately have "fst (spp_of b') = fst (spp_of b)" by (rule is_rewrite_ordD4[OF rword])
  hence "lt b' = lt b" by (simp add: spp_of_def)
  with rb_aux_inv1_lt_inj_on[OF assms(1)] have "b' = b" using b'  set bs b  set bs
    by (rule inj_onD)
  from rword_is_strict_rewrite_ord have "¬ rword_strict (spp_of b) (spp_of b')"
    unfolding b' = b by (rule is_strict_rewrite_ord_irrefl)
  thus False using 2 ..
qed

lemma lemma_12:
  assumes "rb_aux_inv (bs, ss, ps)" and "is_RB_upt dgrad rword (set bs) u"
    and "dgrad (pp_of_term u)  dgrad_max dgrad" and "is_canon_rewriter rword (set bs) u a"
    and "¬ is_syz_sig dgrad u" and "is_sig_red (≺t) (=) (set bs) (monom_mult 1 (pp_of_term u - lp a) a)"
  obtains p q where "p  set bs" and "q  set bs" and "is_regular_spair p q" and "lt (spair p q) = u"
    and "¬ sig_crit' bs (Inl (p, q))"
proof -
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence inj: "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
  from assms(4) have "lt a addst u" by (rule is_canon_rewriterD3)
  hence "lp a adds pp_of_term u" and comp_a: "component_of_term (lt a) = component_of_term u"
    by (simp_all add: adds_term_def)
  let ?s = "pp_of_term u - lp a"
  let ?a = "monom_mult 1 ?s a"
  from assms(4) have "a  set bs" by (rule is_canon_rewriterD1)
  from assms(6) have "rep_list ?a  0" using is_sig_red_top_addsE by blast
  hence "rep_list a  0" by (auto simp: rep_list_monom_mult)
  hence "a  0" by (auto simp: rep_list_zero)
  hence "lt ?a = ?s  lt a" by (simp add: lt_monom_mult)
  also from lp a adds pp_of_term u have eq0: "... = u"
    by (simp add: splus_def comp_a adds_minus term_simps)
  finally have "lt ?a = u" .
  note dgrad rword(1)
  moreover from assms(2) have "is_RB_upt dgrad rword (set bs) (lt ?a)" by (simp only: lt ?a = u)
  moreover from dgrad have "?a  dgrad_sig_set dgrad"
  proof (rule dgrad_sig_set_closed_monom_mult)
    from dgrad lp a adds pp_of_term u have "dgrad (pp_of_term u - lp a)  dgrad (pp_of_term u)"
      by (rule dickson_grading_minus)
    thus "dgrad (pp_of_term u - lp a)  dgrad_max dgrad" using assms(3) by (rule le_trans)
  next
    from inv1 have "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
    with a  set bs show "a  dgrad_sig_set dgrad" ..
  qed
  ultimately obtain v b where "v t lt ?a" and "dgrad (pp_of_term v)  dgrad_max dgrad"
    and "component_of_term v < length fs" and ns: "¬ is_syz_sig dgrad v"
    and v: "v = (punit.lt (rep_list ?a) - punit.lt (rep_list b))  lt b"
    and cr: "is_canon_rewriter rword (set bs) v b" and "is_sig_red (≺t) (=) {b} ?a"
    using assms(6) by (rule lemma_11)
  from this(6) have "b  set bs" by (rule is_canon_rewriterD1)
  with a  set bs show ?thesis
  proof
    from dgrad rword(1) assms(2) inj assms(5, 4) b  set bs is_sig_red (≺t) (=) {b} ?a assms(3)
    show "is_regular_spair a b" by (rule lemma_9(3))
  next
    from dgrad rword(1) assms(2) inj assms(5, 4) b  set bs is_sig_red (≺t) (=) {b} ?a assms(3)
    show "lt (spair a b) = u" by (rule lemma_9(4))
  next
    from rep_list a  0 have v': "v = (?s + punit.lt (rep_list a) - punit.lt (rep_list b))  lt b"
      by (simp add: v rep_list_monom_mult punit.lt_monom_mult)
    moreover from dgrad rword(1) assms(2) inj assms(5, 4) b  set bs is_sig_red (≺t) (=) {b} ?a assms(3)
    have "lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list a) = ?s"
      and "lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list b) =
            ?s + punit.lt (rep_list a) - punit.lt (rep_list b)"
      by (rule lemma_9)+
    ultimately have eq1: "spair_sigs a b = (u, v)" by (simp add: spair_sigs_def eq0)
    show "¬ sig_crit' bs (Inl (a, b))"
    proof (simp add: eq1 assms(5) ns, intro conjI notI)
      assume "is_rewritable bs a u"
      with inv1 have "¬ is_canon_rewriter rword (set bs) u a" by (rule is_rewritableD_is_canon_rewriter)
      thus False using assms(4) ..
    next
      assume "is_rewritable bs b v"
      with inv1 have "¬ is_canon_rewriter rword (set bs) v b" by (rule is_rewritableD_is_canon_rewriter)
      thus False using cr ..
    qed
  qed
qed

lemma is_canon_rewriterI_eq_sig:
  assumes "rb_aux_inv1 bs" and "b  set bs"
  shows "is_canon_rewriter rword (set bs) (lt b) b"
proof -
  from assms(2) have "rep_list b  rep_list ` set bs" by (rule imageI)
  moreover from assms(1) have "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
  ultimately have "b  0" by (auto simp: rep_list_zero)
  with assms(2) show ?thesis
  proof (rule is_canon_rewriterI)
    fix a
    assume "a  set bs" and "a  0" and "lt a addst lt b"
    from assms(2) obtain i where "i < length bs" and b: "b = bs ! i" by (metis in_set_conv_nth)
    from assms(1) this(1) have "is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))"
      by (rule rb_aux_inv1_D5)
    with dgrad have "is_sig_GB_upt dgrad (set (drop (Suc i) bs)) (lt (bs ! i))"
      by (rule is_RB_upt_is_sig_GB_upt)
    hence "is_sig_GB_upt dgrad (set (drop (Suc i) bs)) (lt b)" by (simp only: b)
    moreover have "set (drop (Suc i) bs)  set bs" by (rule set_drop_subset)
    moreover from assms(1) have "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
    ultimately have "is_sig_GB_upt dgrad (set bs) (lt b)" by (rule is_sig_GB_upt_mono)
    with rword(1) dgrad show "rword (spp_of a) (spp_of b)"
    proof (rule is_rewrite_ordD5)
      from assms(1) i < length bs have "¬ is_sig_red (≺t) (≼) (set (drop (Suc i) bs)) (bs ! i)"
        by (rule rb_aux_inv1_D4)
      hence "¬ is_sig_red (≺t) (=) (set (drop (Suc i) bs)) b" by (simp add: b is_sig_red_top_tail_cases)
      moreover have "¬ is_sig_red (≺t) (=) (set (take (Suc i) bs)) b"
      proof
        assume "is_sig_red (≺t) (=) (set (take (Suc i) bs)) b"
        then obtain f where f_in: "f  set (take (Suc i) bs)" and "is_sig_red (≺t) (=) {f} b"
          by (rule is_sig_red_singletonI)
        from this(2) have "lt f t lt b" by (rule is_sig_red_regularD_lt)
        from i < length bs have take_eq: "take (Suc i) bs = (take i bs) @ [b]"
          unfolding b by (rule take_Suc_conv_app_nth)
        from assms(1) have "sorted_wrt (λx y. lt y t lt x) ((take (Suc i) bs) @ (drop (Suc i) bs))"
          unfolding append_take_drop_id by (rule rb_aux_inv1_D3)
        hence 1: "y. y  set (take i bs)  lt b t lt y"
          by (simp add: sorted_wrt_append take_eq del: append_take_drop_id)
        from f_in have "f = b  f  set (take i bs)" by (simp add: take_eq)
        hence "lt b t lt f"
        proof
          assume "f  set (take i bs)"
          hence "lt b t lt f" by (rule 1)
          thus ?thesis by simp
        qed simp
        with lt f t lt b show False by simp
      qed
      ultimately have "¬ is_sig_red (≺t) (=) (set (take (Suc i) bs)  set (drop (Suc i) bs)) b"
        by (simp add: is_sig_red_Un)
      thus "¬ is_sig_red (≺t) (=) (set bs) b" by (metis append_take_drop_id set_append)
    qed fact+
  qed (simp add: term_simps)
qed

lemma not_sig_crit:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "¬ sig_crit bs (new_syz_sigs ss bs p) p" and "b  set bs"
  shows "lt b t sig_of_pair p"
proof (rule sum_prodE)
  fix x y
  assume p: "p = Inl (x, y)"
  have "p  set (p # ps)" by simp
  hence "Inl (x, y)  set (p # ps)" by (simp only: p)
  define t1 where "t1 = punit.lt (rep_list x)"
  define t2 where "t2 = punit.lt (rep_list y)"
  define u where "u = fst (spair_sigs x y)"
  define v where "v = snd (spair_sigs x y)"
  have u: "u = (lcs t1 t2 - t1)  lt x" by (simp add: u_def spair_sigs_def t1_def t2_def Let_def)
  have v: "v = (lcs t1 t2 - t2)  lt y" by (simp add: v_def spair_sigs_def t1_def t2_def Let_def)
  have spair_sigs: "spair_sigs x y = (u, v)" by (simp add: u_def v_def)
  with assms(2) have "¬ is_rewritable bs x u" and "¬ is_rewritable bs y v"
    by (simp_all add: p)
  from assms(1) Inl (x, y)  set (p # ps) have x_in: "x  set bs" and y_in: "y  set bs"
    and "is_regular_spair x y" by (rule rb_aux_inv_D3)+
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  from inv1 have "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
  with x_in y_in have "rep_list x  0" and "rep_list y  0" by auto
  hence "x  0" and "y  0" by (auto simp: rep_list_zero)
  from inv1 have sorted: "sorted_wrt (λx y. lt y t lt x) bs" by (rule rb_aux_inv1_D3)
  from x_in obtain i1 where "i1 < length bs" and x: "x = bs ! i1" by (metis in_set_conv_nth)
  from y_in obtain i2 where "i2 < length bs" and y: "y = bs ! i2" by (metis in_set_conv_nth)
  have "lt b  sig_of_pair p"
  proof
    assume lt_b: "lt b = sig_of_pair p"
    from inv1 have crw: "is_canon_rewriter rword (set bs) (lt b) b" using assms(3)
      by (rule is_canon_rewriterI_eq_sig)
    show False
    proof (rule ord_term_lin.linorder_cases)
      assume "u t v"
      hence "lt b = v" by (auto simp: lt_b p spair_sigs ord_term_lin.max_def)
      with crw have crw_b: "is_canon_rewriter rword (set bs) v b" by simp
      from v have "lt y addst v" by (rule adds_termI)
      hence "is_canon_rewriter rword (set bs) v y"
        using inv1 y_in y  0 ¬ is_rewritable bs y v is_rewritableI_is_canon_rewriter by blast
      with inv1 crw_b have "b = y" by (rule canon_rewriter_unique)
      with lt b = v have "lt y = v" by simp
      from inv1 i2 < length bs have "¬ is_sig_red (≺t) (≼) (set (drop (Suc i2) bs)) (bs ! i2)"
        by (rule rb_aux_inv1_D4)
      moreover have "is_sig_red (≺t) (≼) (set (drop (Suc i2) bs)) (bs ! i2)"
      proof (rule is_sig_red_singletonD)
        have "is_sig_red (≺t) (=) {x} y"
        proof (rule is_sig_red_top_addsI)
          from lt y = v have "(lcs t1 t2 - t2)  lt y = lt y" by (simp only: v)
          also have "... = 0  lt y" by (simp only: term_simps)
          finally have "lcs t1 t2 - t2 = 0" by (simp only: splus_right_canc)
          hence "lcs t1 t2 = t2" by (metis (full_types) add.left_neutral adds_minus adds_lcs_2)
          with adds_lcs[of t1 t2] show "punit.lt (rep_list x) adds punit.lt (rep_list y)"
            by (simp only: t1_def t2_def)
        next
          from u t v show "punit.lt (rep_list y)  lt x t punit.lt (rep_list x)  lt y"
            by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
        qed (simp|fact)+
        thus "is_sig_red (≺t) (≼) {x} (bs ! i2)" by (simp add: y is_sig_red_top_tail_cases)
      next
        have "lt x t 0  lt x" by (simp only: term_simps)
        also have "... t u" unfolding u using zero_min by (rule splus_mono_left)
        also have "... t v" by fact
        finally have *: "lt (bs ! i1) t lt (bs ! i2)" by (simp only: lt y = v x y[symmetric])
        have "i2 < i1"
        proof (rule linorder_cases)
          assume "i1 < i2"
          with sorted have "lt (bs ! i2) t lt (bs ! i1)" using i2 < length bs
            by (rule sorted_wrt_nth_less)
          with * show ?thesis by simp
        next
          assume "i1 = i2"
          with * show ?thesis by simp
        qed
        hence "Suc i2  i1" by simp
        thus "x  set (drop (Suc i2) bs)" unfolding x using i1 < length bs by (rule nth_in_set_dropI)
      qed
      ultimately show ?thesis ..
    next
      assume "v t u"
      hence "lt b = u" by (auto simp: lt_b p spair_sigs ord_term_lin.max_def)
      with crw have crw_b: "is_canon_rewriter rword (set bs) u b" by simp
      from u have "lt x addst u" by (rule adds_termI)
      hence "is_canon_rewriter rword (set bs) u x"
        using inv1 x_in x  0 ¬ is_rewritable bs x u is_rewritableI_is_canon_rewriter by blast
      with inv1 crw_b have "b = x" by (rule canon_rewriter_unique)
      with lt b = u have "lt x = u" by simp
      from inv1 i1 < length bs have "¬ is_sig_red (≺t) (≼) (set (drop (Suc i1) bs)) (bs ! i1)"
        by (rule rb_aux_inv1_D4)
      moreover have "is_sig_red (≺t) (≼) (set (drop (Suc i1) bs)) (bs ! i1)"
      proof (rule is_sig_red_singletonD)
        have "is_sig_red (≺t) (=) {y} x"
        proof (rule is_sig_red_top_addsI)
          from lt x = u have "(lcs t1 t2 - t1)  lt x = lt x" by (simp only: u)
          also have "... = 0  lt x" by (simp only: term_simps)
          finally have "lcs t1 t2 - t1 = 0" by (simp only: splus_right_canc)
          hence "lcs t1 t2 = t1" by (metis (full_types) add.left_neutral adds_minus adds_lcs)
          with adds_lcs_2[of t2 t1] show "punit.lt (rep_list y) adds punit.lt (rep_list x)"
            by (simp only: t1_def t2_def)
        next
          from v t u show "punit.lt (rep_list x)  lt y t punit.lt (rep_list y)  lt x"
            by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
        qed (simp|fact)+
        thus "is_sig_red (≺t) (≼) {y} (bs ! i1)" by (simp add: x is_sig_red_top_tail_cases)
      next
        have "lt y t 0  lt y" by (simp only: term_simps)
        also have "... t v" unfolding v using zero_min by (rule splus_mono_left)
        also have "... t u" by fact
        finally have *: "lt (bs ! i2) t lt (bs ! i1)" by (simp only: lt x = u y x[symmetric])
        have "i1 < i2"
        proof (rule linorder_cases)
          assume "i2 < i1"
          with sorted have "lt (bs ! i1) t lt (bs ! i2)" using i1 < length bs
            by (rule sorted_wrt_nth_less)
          with * show ?thesis by simp
        next
          assume "i2 = i1"
          with * show ?thesis by simp
        qed
        hence "Suc i1  i2" by simp
        thus "y  set (drop (Suc i1) bs)" unfolding y using i2 < length bs by (rule nth_in_set_dropI)
      qed
      ultimately show ?thesis ..
    next
      assume "u = v"
      hence "punit.lt (rep_list x)  lt y = punit.lt (rep_list y)  lt x"
        by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
      moreover from is_regular_spair x y
      have "punit.lt (rep_list y)  lt x  punit.lt (rep_list x)  lt y" by (rule is_regular_spairD3)
      ultimately show ?thesis by simp
    qed
  qed
  moreover from assms(1, 3) p  set (p # ps) have "lt b t sig_of_pair p" by (rule rb_aux_inv_D7)
  ultimately show ?thesis by simp
next
  fix j
  assume p: "p = Inr j"
  have "Inr j  set (p # ps)" by (simp add: p)
  with assms(1) have "lt b t term_of_pair (0, j)" using assms(3) by (rule rb_aux_inv_D4)
  thus ?thesis by (simp add: p)
qed

context
  assumes fs_distinct: "distinct fs"
  assumes fs_nonzero: "0  set fs"
begin

lemma rep_list_monomial': "rep_list (monomial 1 (term_of_pair (0, j))) = ((fs ! j) when j < length fs)"
  by (simp add: rep_list_monomial fs_distinct term_simps)

lemma new_syz_sigs_is_syz_sig:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "v  set (new_syz_sigs ss bs p)"
  shows "is_syz_sig dgrad v"
proof (rule sum_prodE)
  fix a b
  assume "p = Inl (a, b)"
  with assms(2) have "v  set ss" by simp
  with assms(1) show ?thesis by (rule rb_aux_inv_D2)
next
  fix j
  assume p: "p = Inr j"
  let ?f = "λb. term_of_pair (punit.lt (rep_list b), j)"
  let ?a = "monomial (1::'b) (term_of_pair (0, j))"
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  have "Inr j  set (p # ps)" by (simp add: p)
  with assms(1) have "j < length fs" by (rule rb_aux_inv_D4)
  hence a: "rep_list ?a = fs ! j" by (simp add: rep_list_monomial')
  show ?thesis
  proof (cases "is_pot_ord")
    case True
    with assms(2) have "v  set (filter_min_append (addst) ss (filter_min (addst) (map ?f bs)))"
      by (simp add: p)
    hence "v  set ss  ?f ` set bs" using filter_min_append_subset filter_min_subset by fastforce
    thus ?thesis
    proof
      assume "v  set ss"
      with assms(1) show ?thesis by (rule rb_aux_inv_D2)
    next
      assume "v  ?f ` set bs"
      then obtain b where "b  set bs" and "v = ?f b" ..
      have comp_b: "component_of_term (lt b) < component_of_term (lt ?a)"
      proof (rule ccontr)
        have *: "pp_of_term (term_of_pair (0, j))  pp_of_term (lt b)"
          by (simp add: pp_of_term_of_pair zero_min)
        assume "¬ component_of_term (lt b) < component_of_term (lt ?a)"
        hence "component_of_term (term_of_pair (0, j))  component_of_term (lt b)"
          by (simp add: lt_monomial)
        with * have "term_of_pair (0, j) t lt b" by (rule ord_termI)
        moreover from assms(1) Inr j  set (p # ps) b  set bs have "lt b t term_of_pair (0, j)"
          by (rule rb_aux_inv_D4)
        ultimately show False by simp
      qed
      have "v = punit.lt (rep_list b)  lt ?a"
        by (simp add: v = ?f b a lt_monomial splus_def term_simps)
      also have "... = ord_term_lin.max (punit.lt (rep_list b)  lt ?a) (punit.lt (rep_list ?a)  lt b)"
      proof -
        have "component_of_term (punit.lt (rep_list ?a)  lt b) = component_of_term (lt b)"
          by (simp only: term_simps)
        also have "... < component_of_term (lt ?a)" by (fact comp_b)
        also have "... = component_of_term (punit.lt (rep_list b)  lt ?a)"
          by (simp only: term_simps)
        finally have "component_of_term (punit.lt (rep_list ?a)  lt b) <
                      component_of_term (punit.lt (rep_list b)  lt ?a)" .
        with True have "punit.lt (rep_list ?a)  lt b t punit.lt (rep_list b)  lt ?a"
          by (rule is_pot_ordD)
        thus ?thesis by (auto simp: ord_term_lin.max_def)
      qed
      finally have v: "v = ord_term_lin.max (punit.lt (rep_list b)  lt ?a) (punit.lt (rep_list ?a)  lt b)" .
      show ?thesis unfolding v using dgrad
      proof (rule Koszul_syz_is_syz_sig)
        from inv1 have "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
        with b  set bs show "b  dgrad_sig_set dgrad" ..
      next
        show "?a  dgrad_sig_set dgrad"
          by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 j < length fs)
      next
        from inv1 have "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
        with b  set bs show "rep_list b  0" by fastforce
      next
        from j < length fs have "fs ! j  set fs" by (rule nth_mem)
        with fs_nonzero show "rep_list ?a  0" by (auto simp: a)
      qed (fact comp_b)
    qed
  next
    case False
    with assms(2) have "v  set ss" by (simp add: p)
    with assms(1) show ?thesis by (rule rb_aux_inv_D2)
  qed
qed

lemma new_syz_sigs_minimal:
  assumes "u' v'. u'  set ss  v'  set ss  u' addst v'  u' = v'"
  assumes "u  set (new_syz_sigs ss bs p)" and "v  set (new_syz_sigs ss bs p)" and "u addst v"
  shows "u = v"
proof (rule sum_prodE)
  fix a b
  assume p: "p = Inl (a, b)"
  from assms(2, 3) have "u  set ss" and "v  set ss" by (simp_all add: p)
  thus ?thesis using assms(4) by (rule assms(1))
next
  fix j
  assume p: "p = Inr j"
  show ?thesis
  proof (cases is_pot_ord)
    case True
    have "transp (addst)" by (rule transpI, drule adds_term_trans)
    define ss' where "ss' = filter_min (addst) (map (λb. term_of_pair (punit.lt (rep_list b), j)) bs)"
    note assms(1)
    moreover have "u' = v'" if "u'  set ss'" and "v'  set ss'" and "u' addst v'" for u' v'
      using transp (addst) that unfolding ss'_def by (rule filter_min_minimal)
    moreover from True assms(2, 3) have "u  set (filter_min_append (addst) ss ss')"
      and "v  set (filter_min_append (addst) ss ss')" by (simp_all add: p ss'_def)
    ultimately show ?thesis using assms(4) by (rule filter_min_append_minimal)
  next
    case False
    with assms(2, 3) have "u  set ss" and "v  set ss" by (simp_all add: p)
    thus ?thesis using assms(4) by (rule assms(1))
  qed
qed

lemma new_syz_sigs_distinct:
  assumes "distinct ss"
  shows "distinct (new_syz_sigs ss bs p)"
proof (rule sum_prodE)
  fix a b
  assume "p = Inl (a, b)"
  with assms show ?thesis by simp
next
  fix j
  assume p: "p = Inr j"
  show ?thesis
  proof (cases is_pot_ord)
    case True
    define ss' where "ss' = filter_min (addst) (map (λb. term_of_pair (punit.lt (rep_list b), j)) bs)"
    from adds_term_refl have "reflp (addst)" by (rule reflpI)
    moreover note assms
    moreover have "distinct ss'" unfolding ss'_def using reflp (addst) by (rule filter_min_distinct)
    ultimately have "distinct (filter_min_append (addst) ss ss')" by (rule filter_min_append_distinct)
    thus ?thesis by (simp add: p ss'_def True)
  next
    case False
    with assms show ?thesis by (simp add: p)
  qed
qed

lemma sig_crit'I_sig_crit:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
  shows "sig_crit' bs p"
proof -
  have rl: "is_syz_sig dgrad u"
    if "is_pred_syz (new_syz_sigs ss bs p) u" and "dgrad (pp_of_term u)  dgrad_max dgrad" for u
  proof -
    from that(1) obtain s where "s  set (new_syz_sigs ss bs p)" and adds: "s addst u"
      unfolding is_pred_syz_def ..
    from assms(1) this(1) have "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
    with dgrad show ?thesis using adds that(2) by (rule is_syz_sig_adds)
  qed
  from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence bs_sub: "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  show ?thesis
  proof (rule sum_prodE)
    fix a b
    assume p: "p = Inl (a, b)"
    hence "Inl (a, b)  set (p # ps)" by simp
    with assms(1) have "a  set bs" and "b  set bs" by (rule rb_aux_inv_D3)+
    with bs_sub have a_in: "a  dgrad_sig_set dgrad" and b_in: "b  dgrad_sig_set dgrad" by fastforce+
    define t1 where "t1 = punit.lt (rep_list a)"
    define t2 where "t2 = punit.lt (rep_list b)"
    define u where "u = fst (spair_sigs a b)"
    define v where "v = snd (spair_sigs a b)"
    from dgrad a_in have "dgrad t1  dgrad_max dgrad" unfolding t1_def by (rule dgrad_sig_setD_rep_list_lt)
    moreover from dgrad b_in have "dgrad t2  dgrad_max dgrad"
      unfolding t2_def by (rule dgrad_sig_setD_rep_list_lt)
    ultimately have "ord_class.max (dgrad t1) (dgrad t2)  dgrad_max dgrad" by simp
    with dickson_grading_lcs[OF dgrad] have "dgrad (lcs t1 t2)  dgrad_max dgrad" by (rule le_trans)
    have u: "u = (lcs t1 t2 - t1)  lt a" by (simp add: u_def spair_sigs_def t1_def t2_def Let_def)
    have v: "v = (lcs t1 t2 - t2)  lt b" by (simp add: v_def spair_sigs_def t1_def t2_def Let_def)
    have 1: "spair_sigs a b = (u, v)" by (simp add: u_def v_def)
    from assms(2) have "(is_pred_syz (new_syz_sigs ss bs p) u  is_pred_syz (new_syz_sigs ss bs p) v) 
                        (is_rewritable bs a u  is_rewritable bs b v)" by (simp add: p 1)
    thus ?thesis
    proof
      assume "is_pred_syz (new_syz_sigs ss bs p) u  is_pred_syz (new_syz_sigs ss bs p) v"
      thus ?thesis
      proof
        assume "is_pred_syz (new_syz_sigs ss bs p) u"
        moreover have "dgrad (pp_of_term u)  dgrad_max dgrad"
        proof (simp add: u term_simps dickson_gradingD1[OF dgrad], rule)
          from dgrad adds_lcs have "dgrad (lcs t1 t2 - t1)  dgrad (lcs t1 t2)"
            by (rule dickson_grading_minus)
          also have "...  dgrad_max dgrad" by fact
          finally show "dgrad (lcs t1 t2 - t1)  dgrad_max dgrad" .
        next
          from a_in show "dgrad (lp a)  dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
        qed
        ultimately have "is_syz_sig dgrad u" by (rule rl)
        thus ?thesis by (simp add: p 1)
      next
        assume "is_pred_syz (new_syz_sigs ss bs p) v"
        moreover have "dgrad (pp_of_term v)  dgrad_max dgrad"
        proof (simp add: v term_simps dickson_gradingD1[OF dgrad], rule)
          from dgrad adds_lcs_2 have "dgrad (lcs t1 t2 - t2)  dgrad (lcs t1 t2)"
            by (rule dickson_grading_minus)
          also have "...  dgrad_max dgrad" by fact
          finally show "dgrad (lcs t1 t2 - t2)  dgrad_max dgrad" .
        next
          from b_in show "dgrad (lp b)  dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
        qed
        ultimately have "is_syz_sig dgrad v" by (rule rl)
        thus ?thesis by (simp add: p 1)
      qed
    next
      assume "is_rewritable bs a u  is_rewritable bs b v"
      thus ?thesis by (simp add: p 1)
    qed
  next
    fix j
    assume "p = Inr j"
    with assms(2) have "is_pred_syz (new_syz_sigs ss bs p) (term_of_pair (0, j))" by simp
    moreover have "dgrad (pp_of_term (term_of_pair (0, j)))  dgrad_max dgrad"
      by (simp add: pp_of_term_of_pair dgrad_max_0)
    ultimately have "is_syz_sig dgrad (term_of_pair (0, j))" by (rule rl)
    thus ?thesis by (simp add: p = Inr j)
  qed
qed

lemma rb_aux_inv_preserved_0:
  assumes "rb_aux_inv (bs, ss, p # ps)"
    and "s. s  set ss'  is_syz_sig dgrad s"
    and "a b. a  set bs  b  set bs  is_regular_spair a b  Inl (a, b)  set ps 
           Inl (b, a)  set ps  ¬ is_RB_in dgrad rword (set bs) (lt (spair a b)) 
           qset ps. sig_of_pair q = lt (spair a b)  ¬ sig_crit' bs q"
    and "j. j < length fs  p = Inr j  Inr j  set ps  is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) 
            rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
  shows "rb_aux_inv (bs, ss', ps)"
proof -
  from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  show ?thesis unfolding rb_aux_inv.simps
  proof (intro conjI ballI allI impI)
    fix s
    assume "s  set ss'"
    thus "is_syz_sig dgrad s" by (rule assms(2))
  next
    fix q1 q2
    assume "Inl (q1, q2)  set ps"
    hence "Inl (q1, q2)  set (p # ps)" by simp
    with assms(1) show "is_regular_spair q1 q2" and "q1  set bs" and "q2  set bs"
      by (rule rb_aux_inv_D3)+
  next
    fix j
    assume "Inr j  set ps"
    hence "Inr j  set (p # ps)" by simp
    with assms(1) have "j < length fs" and "length (filter (λq. sig_of_pair q = term_of_pair (0, j)) (p # ps))  1"
      by (rule rb_aux_inv_D4)+
    have "length (filter (λq. sig_of_pair q = term_of_pair (0, j)) ps) 
          length (filter (λq. sig_of_pair q = term_of_pair (0, j)) (p # ps))" by simp
    also have "...  1" by fact
    finally show "length (filter (λq. sig_of_pair q = term_of_pair (0, j)) ps)  1" .
    show "j < length fs" by fact

    fix b
    assume "b  set bs"
    with assms(1) Inr j  set (p # ps) show "lt b t term_of_pair (0, j)" by (rule rb_aux_inv_D4)
  next
    from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    thus "sorted_wrt pair_ord ps" by simp
  next
    fix q
    assume "q  set ps"
    from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    hence "p'. p'  set ps  sig_of_pair p t sig_of_pair p'" by (simp add: pair_ord_def)
    with q  set ps have 1: "sig_of_pair p t sig_of_pair q" by blast
    {
      fix b1 b2
      note assms(1)
      moreover from q  set ps have "q  set (p # ps)" by simp
      moreover assume "b1  set bs" and "b2  set bs" and "is_regular_spair b1 b2"
        and 2: "sig_of_pair q t lt (spair b1 b2)"
      ultimately show "Inl (b1, b2)  set ps  Inl (b2, b1)  set ps"
      proof (rule rb_aux_inv_D6_1)
        assume "Inl (b1, b2)  set (p # ps)"
        moreover from 1 2 have "sig_of_pair p t lt (spair b1 b2)" by simp
        ultimately have "Inl (b1, b2)  set ps"
          by (auto simp: sig_of_spair is_regular_spair b1 b2 simp del: sig_of_pair.simps)
        thus ?thesis ..
      next
        assume "Inl (b2, b1)  set (p # ps)"
        moreover from 1 2 have "sig_of_pair p t lt (spair b1 b2)" by simp
        ultimately have "Inl (b2, b1)  set ps"
          by (auto simp: sig_of_spair is_regular_spair b1 b2 sig_of_spair_commute simp del: sig_of_pair.simps)
        thus ?thesis ..
      qed
    }
    {
      fix j
      note assms(1)
      moreover from q  set ps have "q  set (p # ps)" by simp
      moreover assume "j < length fs" and 2: "sig_of_pair q t term_of_pair (0, j)"
      ultimately have "Inr j  set (p # ps)" by (rule rb_aux_inv_D6_2)
      moreover from 1 2 have "sig_of_pair p t sig_of_pair (Inr j)" by simp
      ultimately show "Inr j  set ps" by auto
    }
  next
    fix b q
    assume "b  set bs" and "q  set ps"
    hence "b  set bs" and "q  set (p # ps)" by simp_all
    with assms(1) show "lt b t sig_of_pair q" by (rule rb_aux_inv_D7)
  next
    fix j
    assume "j < length fs" and "Inr j  set ps"
    have "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) 
            rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
    proof (cases "p = Inr j")
      case True
      with j < length fs show ?thesis using Inr j  set ps by (rule assms(4))
    next
      case False
      with Inr j  set ps have "Inr j  set (p # ps)" by simp
      with assms(1) j < length fs rb_aux_inv_D9 show ?thesis by blast
    qed
    thus "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
      and "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)" by simp_all
  qed (fact, rule assms(3))
qed

lemma rb_aux_inv_preserved_1:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
  shows "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
proof -
  from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence bs_sub: "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  from assms(1, 2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
  from assms(1) show ?thesis
  proof (rule rb_aux_inv_preserved_0)
    fix s
    assume "s  set (new_syz_sigs ss bs p)"
    with assms(1) show "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
  next
    fix a b
    assume 1: "a  set bs" and 2: "b  set bs" and 3: "is_regular_spair a b" and 4: "Inl (a, b)  set ps"
      and 5: "Inl (b, a)  set ps" and 6: "¬ is_RB_in dgrad rword (set bs) (lt (spair a b))"
    from assms(1, 2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
    show "qset ps. sig_of_pair q = lt (spair a b)  ¬ sig_crit' bs q"
    proof (cases "p = Inl (a, b)  p = Inl (b, a)")
      case True
      hence sig_of_p: "lt (spair a b) = sig_of_pair p"
      proof
        assume p: "p = Inl (a, b)"
        from 3 show ?thesis by (simp only: p sig_of_spair)
      next
        assume p: "p = Inl (b, a)"
        from 3 have "is_regular_spair b a" by (rule is_regular_spair_sym)
        thus ?thesis by (simp only: p sig_of_spair spair_comm[of a] lt_uminus)
      qed
      note assms(1)
      moreover have "is_RB_upt dgrad rword (set bs) (lt (spair a b))" unfolding sig_of_p
        using assms(1) by (rule rb_aux_inv_is_RB_upt_Cons)
      moreover have "dgrad (lp (spair a b))  dgrad_max dgrad"
      proof (rule dgrad_sig_setD_lp, rule dgrad_sig_set_closed_spair, fact dgrad)
        from a  set bs bs_sub show "a  dgrad_sig_set dgrad" ..
      next
        from b  set bs bs_sub show "b  dgrad_sig_set dgrad" ..
      qed
      moreover obtain c where crw: "is_canon_rewriter rword (set bs) (lt (spair a b)) c"
      proof (rule ord_term_lin.linorder_cases)
        from 3 have "rep_list b  0" by (rule is_regular_spairD2)
        moreover assume "punit.lt (rep_list b)  lt a t punit.lt (rep_list a)  lt b"
        ultimately have "lt (spair b a) = (lcs (punit.lt (rep_list b)) (punit.lt (rep_list a)) - punit.lt (rep_list b))  lt b"
          by (rule lt_spair)
        hence "lt (spair a b) = (lcs (punit.lt (rep_list b)) (punit.lt (rep_list a)) - punit.lt (rep_list b))  lt b"
          by (simp add: spair_comm[of a])
        hence "lt b addst lt (spair a b)" by (rule adds_termI)
        from rep_list b  0 have "b  0" by (auto simp: rep_list_zero)
        show ?thesis by (rule is_rewrite_ord_finite_canon_rewriterE, fact rword, fact finite_set, fact+)
      next
        from 3 have "rep_list a  0" by (rule is_regular_spairD1)
        moreover assume "punit.lt (rep_list a)  lt b t punit.lt (rep_list b)  lt a"
        ultimately have "lt (spair a b) = (lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list a))  lt a"
          by (rule lt_spair)
        hence "lt a addst lt (spair a b)" by (rule adds_termI)
        from rep_list a  0 have "a  0" by (auto simp: rep_list_zero)
        show ?thesis by (rule is_rewrite_ord_finite_canon_rewriterE, fact rword, fact finite_set, fact+)
      next
        from 3 have "punit.lt (rep_list b)  lt a  punit.lt (rep_list a)  lt b"
          by (rule is_regular_spairD3)
        moreover assume "punit.lt (rep_list b)  lt a = punit.lt (rep_list a)  lt b"
        ultimately show ?thesis ..
      qed
      moreover from 6 have "¬ is_syz_sig dgrad (lt (spair a b))" by (simp add: is_RB_in_def)
      moreover from 6 crw have "is_sig_red (≺t) (=) (set bs) (monom_mult 1 (lp (spair a b) - lp c) c)"
        by (simp add: is_RB_in_def)
      ultimately obtain x y where 7: "x  set bs" and 8: "y  set bs" and 9: "is_regular_spair x y"
        and 10: "lt (spair x y) = lt (spair a b)" and 11: "¬ sig_crit' bs (Inl (x, y))"
        by (rule lemma_12)
      from this(5) sig_crit' bs p have "Inl (x, y)  p" and "Inl (y, x)  p"
        by (auto simp only: sig_crit'_sym)
      show ?thesis
      proof (cases "Inl (x, y)  set ps  Inl (y, x)  set ps")
        case True
        thus ?thesis
        proof
          assume "Inl (x, y)  set ps"
          show ?thesis
          proof (intro bexI conjI)
            show "sig_of_pair (Inl (x, y)) = lt (spair a b)" by (simp only: sig_of_spair 9 10)
          qed fact+
        next
          assume "Inl (y, x)  set ps"
          show ?thesis
          proof (intro bexI conjI)
            from 9 have "is_regular_spair y x" by (rule is_regular_spair_sym)
            thus "sig_of_pair (Inl (y, x)) = lt (spair a b)"
              by (simp only: sig_of_spair spair_comm[of y] lt_uminus 10)
          next
            from 11 show "¬ sig_crit' bs (Inl (y, x))" by (auto simp only: sig_crit'_sym)
          qed fact
        qed
      next
        case False
        note assms(1) 7 8 9
        moreover from False Inl (x, y)  p Inl (y, x)  p have "Inl (x, y)  set (p # ps)"
          and "Inl (y, x)  set (p # ps)" by auto
        moreover from 6 have "¬ is_RB_in dgrad rword (set bs) (lt (spair x y))" by (simp add: 10)
        ultimately obtain q where 12: "q  set (p # ps)" and 13: "sig_of_pair q = lt (spair x y)"
          and 14: "¬ sig_crit' bs q" by (rule rb_aux_inv_D8)
        from 12 14 sig_crit' bs p have "q  set ps" by auto
        with 13 14 show ?thesis unfolding 10 by blast
      qed
    next
      case False
      with 4 5 have "Inl (a, b)  set (p # ps)" and "Inl (b, a)  set (p # ps)" by auto
      with assms(1) 1 2 3 obtain q where 7: "q  set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)"
        and 9: "¬ sig_crit' bs q" using 6 by (rule rb_aux_inv_D8)
      from 7 9 sig_crit' bs p have "q  set ps" by auto
      with 8 9 show ?thesis by blast
    qed
  next
    fix j
    assume "j < length fs"
    assume p: "p = Inr j"
    with sig_crit' bs p have "is_syz_sig dgrad (term_of_pair (0, j))" by simp
    hence "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (rule is_RB_inI2)
    moreover have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
    proof (rule sig_red_zero_idealI, rule syzygy_crit)
      from assms(1) have "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
        by (rule rb_aux_inv_is_RB_upt_Cons)
      with dgrad have "is_sig_GB_upt dgrad (set bs) (sig_of_pair p)"
        by (rule is_RB_upt_is_sig_GB_upt)
      thus "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))" by (simp add: p)
    next
      show "monomial 1 (term_of_pair (0, j))  dgrad_sig_set dgrad"
        by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 j < length fs)
    next
      show "lt (monomial (1::'b) (term_of_pair (0, j))) = term_of_pair (0, j)" by (simp add: lt_monomial)
    qed (fact dgrad, fact)
    ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) 
                      rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)" ..
  qed
qed

lemma rb_aux_inv_preserved_2:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "rep_list (sig_trd bs (poly_of_pair p)) = 0"
  shows "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
proof -
  let ?p = "sig_trd bs (poly_of_pair p)"
  have 0: "(sig_red (≺t) (≼) (set bs))** (poly_of_pair p) ?p"
    by (rule sig_trd_red_rtrancl)
  hence eq: "lt ?p = lt (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lt)
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  have *: "is_syz_sig dgrad (lt (poly_of_pair p))"
  proof (rule is_syz_sigI)
    have "poly_of_pair p  0" by (rule pair_list_nonzero, fact, simp)
    hence "lc (poly_of_pair p)  0" by (rule lc_not_0)
    moreover from 0 have "lc ?p = lc (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lc)
    ultimately have "lc ?p  0" by simp
    thus "?p  0" by (simp add: lc_eq_zero_iff)
  next
    note dgrad(1)
    moreover from inv1 have "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
    moreover have "poly_of_pair p  dgrad_sig_set dgrad" by (rule pair_list_dgrad_sig_set, fact, simp)
    ultimately show "?p  dgrad_sig_set dgrad" using 0 by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  qed (fact eq, fact assms(2))
  hence rb: "is_RB_in dgrad rword (set bs) (lt (poly_of_pair p))" by (rule is_RB_inI2)
  from assms(1) show ?thesis
  proof (rule rb_aux_inv_preserved_0)
    fix s
    assume "s  set (lt ?p # new_syz_sigs ss bs p)"
    hence "s = lt (poly_of_pair p)  s  set (new_syz_sigs ss bs p)" by (simp add: eq)
    thus "is_syz_sig dgrad s"
    proof
      assume "s = lt (poly_of_pair p)"
      with * show ?thesis by simp
    next
      assume "s  set (new_syz_sigs ss bs p)"
      with assms(1) show ?thesis by (rule new_syz_sigs_is_syz_sig)
    qed
  next
    fix a b
    assume 1: "a  set bs" and 2: "b  set bs" and 3: "is_regular_spair a b" and 4: "Inl (a, b)  set ps"
    and 5: "Inl (b, a)  set ps" and 6: "¬ is_RB_in dgrad rword (set bs) (lt (spair a b))"
    have "p  set (p # ps)" by simp
    with assms(1) have sig_of_p: "sig_of_pair p = lt (poly_of_pair p)" by (rule pair_list_sig_of_pair)
    from rb 6 have neq: "lt (poly_of_pair p)  lt (spair a b)" by auto
    hence "p  Inl (a, b)" and "p  Inl (b, a)" by (auto simp: spair_comm[of a])
    with 4 5 have "Inl (a, b)  set (p # ps)" and "Inl (b, a)  set (p # ps)" by auto
    with assms(1) 1 2 3 obtain q where 7: "q  set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)"
      and 9: "¬ sig_crit' bs q" using 6 by (rule rb_aux_inv_D8)
    from this(1, 2) neq have "q  set ps" by (auto simp: sig_of_p)
    thus "qset ps. sig_of_pair q = lt (spair a b)  ¬ sig_crit' bs q" using 8 9 by blast
  next
    fix j
    assume "j < length fs"
    assume p: "p = Inr j"
    from rb have "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (simp add: p lt_monomial)
    moreover have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
    proof (rule sig_red_zero_idealI, rule sig_red_zeroI)
      from 0 show "(sig_red (≺t) (≼) (set bs))** (monomial 1 (term_of_pair (0, j))) ?p" by (simp add: p)
    qed fact
    ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) 
                      rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)" ..
  qed
qed

lemma rb_aux_inv_preserved_3:
  assumes "rb_aux_inv (bs, ss, p # ps)" and "¬ sig_crit bs (new_syz_sigs ss bs p) p"
    and "rep_list (sig_trd bs (poly_of_pair p))  0"
  shows "rb_aux_inv ((sig_trd bs (poly_of_pair p)) # bs, new_syz_sigs ss bs p,
                          add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
    and "lt (sig_trd bs (poly_of_pair p))  lt ` set bs"
proof -
  have "p  set (p # ps)" by simp
  with assms(1) have sig_of_p: "sig_of_pair p = lt (poly_of_pair p)"
    and p_in: "poly_of_pair p  dgrad_sig_set dgrad"
    by (rule pair_list_sig_of_pair, rule pair_list_dgrad_sig_set)
  define p' where "p' = sig_trd bs (poly_of_pair p)"
  from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence bs_sub: "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  have p_red: "(sig_red (≺t) (≼) (set bs))** (poly_of_pair p) p'"
    and p'_irred: "¬ is_sig_red (≺t) (≼) (set bs) p'"
    unfolding p'_def by (rule sig_trd_red_rtrancl, rule sig_trd_irred)
  from dgrad bs_sub p_in p_red have p'_in: "p'  dgrad_sig_set dgrad"
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  from p_red have lt_p': "lt p' = lt (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lt)
  have sig_merge: "sig_of_pair p t sig_of_pair q" if "q  set (add_spairs ps bs p')" for q
    using that unfolding add_spairs_def set_merge_wrt
  proof
    assume "q  set (new_spairs bs p')"
    then obtain b0 where "is_regular_spair p' b0" and "q = Inl (p', b0)" by (rule in_new_spairsE)
    hence sig_of_q: "sig_of_pair q = lt (spair p' b0)" by (simp only: sig_of_spair)
    show ?thesis unfolding sig_of_q sig_of_p lt_p'[symmetric] by (rule is_regular_spair_lt_ge_1, fact)
  next
    assume "q  set ps"
    moreover from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    ultimately show ?thesis by (simp add: pair_ord_def)
  qed
  have sig_of_p_less: "sig_of_pair p t term_of_pair (0, j)" if "Inr j  set ps" for j
  proof (intro ord_term_lin.le_neq_trans)
    from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    with Inr j  set ps show "sig_of_pair p t term_of_pair (0, j)"
      by (auto simp: pair_ord_def)
  next
    from assms(1) that show "sig_of_pair p  term_of_pair (0, j)" by (rule Inr_in_tailD)
  qed
  have lt_p_gr: "lt b t lt (poly_of_pair p)" if "b  set bs" for b unfolding sig_of_p[symmetric]
    using assms(1, 2) that by (rule not_sig_crit)
  have inv1: "rb_aux_inv1 (p' # bs)" unfolding rb_aux_inv1_def
  proof (intro conjI impI allI)
    from bs_sub p'_in show "set (p' # bs)  dgrad_sig_set dgrad" by simp
  next
    from inv1 have "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
    with assms(3) show "0  rep_list ` set (p' # bs)" by (simp add: p'_def)
  next
    from inv1 have "sorted_wrt (λx y. lt y t lt x) bs" by (rule rb_aux_inv1_D3)
    with lt_p_gr show "sorted_wrt (λx y. lt y t lt x) (p' # bs)" by (simp add: lt_p')
  next
    fix i
    assume "i < length (p' # bs)"
    have "(¬ is_sig_red (≺t) (≼) (set (drop (Suc i) (p' # bs))) ((p' # bs) ! i)) 
          ((j<length fs. lt ((p' # bs) ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) 
             punit.lt (rep_list ((p' # bs) ! i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
         (pset (p' # bs). qset (p' # bs). is_regular_spair p q  rep_list (spair p q)  0 
                lt ((p' # bs) ! i) = lt (spair p q) 
                punit.lt (rep_list ((p' # bs) ! i))  punit.lt (rep_list (spair p q)))) 
          is_RB_upt dgrad rword (set (drop (Suc i) (p' # bs))) (lt ((p' # bs) ! i))"
      (is "?thesis1  ?thesis2  ?thesis3")
    proof (cases i)
      case 0
      show ?thesis
      proof (simp add: i = 0 p'_irred del: bex_simps, rule conjI)
        show "(j<length fs. lt p' = lt (monomial (1::'b) (term_of_pair (0, j))) 
                  punit.lt (rep_list p')  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
              (pinsert p' (set bs). qinsert p' (set bs). is_regular_spair p q  rep_list (spair p q)  0 
                     lt p' = lt (spair p q)  punit.lt (rep_list p')  punit.lt (rep_list (spair p q)))"
        proof (rule sum_prodE)
          fix a b
          assume p: "p = Inl (a, b)"
          have "Inl (a, b)  set (p # ps)" by (simp add: p)
          with assms(1) have "a  set bs" and "b  set bs" and "is_regular_spair a b"
            by (rule rb_aux_inv_D3)+
          from p_red have p'_red: "(sig_red (≺t) (≼) (set bs))** (spair a b) p'" by (simp add: p)
          hence "(punit.red (rep_list ` set bs))** (rep_list (spair a b)) (rep_list p')"
            by (rule sig_red_red_rtrancl)
          moreover from assms(3) have "rep_list p'  0" by (simp add: p'_def)
          ultimately have "rep_list (spair a b)  0" by (auto dest: punit.rtrancl_0)
          moreover from p'_red have "lt p' = lt (spair a b)"
            and "punit.lt (rep_list p')  punit.lt (rep_list (spair a b))"
            by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
          ultimately show ?thesis using a  set bs b  set bs is_regular_spair a b by blast
        next
          fix j
          assume "p = Inr j"
          hence "Inr j  set (p # ps)" by simp
          with assms(1) have "j < length fs" by (rule rb_aux_inv_D4)
          from p_red have "(sig_red (≺t) (≼) (set bs))** (monomial 1 (term_of_pair (0, j))) p'"
            by (simp add: p = Inr j)
          hence "lt p' = lt (monomial (1::'b) (term_of_pair (0, j)))"
            and "punit.lt (rep_list p')  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))"
            by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
          with j < length fs show ?thesis by blast
        qed
      next
        from assms(1) show "is_RB_upt dgrad rword (set bs) (lt p')" unfolding lt_p' sig_of_p[symmetric]
          by (rule rb_aux_inv_is_RB_upt_Cons)
      qed
    next
      case (Suc i')
      with i < length (p' # bs) have i': "i' < length bs" by simp
      show ?thesis
      proof (simp add: i = Suc i' del: bex_simps, intro conjI)
        from inv1 i' show "¬ is_sig_red (≺t) (≼) (set (drop (Suc i') bs)) (bs ! i')"
          by (rule rb_aux_inv1_D4)
      next
        from inv1 i'
        show "(j<length fs. lt (bs ! i') = lt (monomial (1::'b) (term_of_pair (0, j))) 
                punit.lt (rep_list (bs ! i'))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
            (pinsert p' (set bs). qinsert p' (set bs). is_regular_spair p q  rep_list (spair p q)  0 
                lt (bs ! i') = lt (spair p q)  punit.lt (rep_list (bs ! i'))  punit.lt (rep_list (spair p q)))"
          by (auto elim!: rb_aux_inv1_E)
      next
        from inv1 i' show "is_RB_upt dgrad rword (set (drop (Suc i') bs)) (lt (bs ! i'))"
          by (rule rb_aux_inv1_D5)
      qed
    qed
    thus ?thesis1 and ?thesis2 and ?thesis3 by simp_all
  qed
  have rb: "is_RB_in dgrad rword (set (p' # bs)) (sig_of_pair p)"
  proof (rule is_RB_inI1)
    have "p'  set (p' # bs)" by simp
    with inv1 have "is_canon_rewriter rword (set (p' # bs)) (lt p') p'"
      by (rule is_canon_rewriterI_eq_sig)
    thus "is_canon_rewriter rword (set (p' # bs)) (sig_of_pair p) p'" by (simp add: lt_p' sig_of_p)
  next
    from p'_irred have "¬ is_sig_red (≺t) (=) (set bs) p'"
      by (simp add: is_sig_red_top_tail_cases)
    with sig_irred_regular_self have "¬ is_sig_red (≺t) (=) ({p'}  set bs) p'"
      by (simp add: is_sig_red_Un del: Un_insert_left)
    thus "¬ is_sig_red (≺t) (=) (set (p' # bs)) (monom_mult 1 (pp_of_term (sig_of_pair p) - lp p') p')"
      by (simp add: lt_p' sig_of_p)
  qed
  show "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
    unfolding rb_aux_inv.simps
  proof (intro conjI ballI allI impI)
    show "rb_aux_inv1 (p' # bs)" by (fact inv1)
  next
    fix s
    assume "s  set (new_syz_sigs ss bs p)"
    with assms(1) show "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
  next
    fix q1 q2
    assume "Inl (q1, q2)  set (add_spairs ps bs p')"
    hence "Inl (q1, q2)  set (new_spairs bs p')  Inl (q1, q2)  set (p # ps)"
      by (auto simp: add_spairs_def set_merge_wrt)
    hence "is_regular_spair q1 q2  q1  set (p' # bs)  q2  set (p' # bs)"
    proof
      assume "Inl (q1, q2)  set (new_spairs bs p')"
      hence "q1 = p'" and "q2  set bs" and "is_regular_spair p' q2" by (rule in_new_spairsD)+
      thus ?thesis by simp
    next
      assume "Inl (q1, q2)  set (p # ps)"
      with assms(1) have "is_regular_spair q1 q2" and "q1  set bs" and "q2  set bs"
        by (rule rb_aux_inv_D3)+
      thus ?thesis by simp
    qed
    thus "is_regular_spair q1 q2" and "q1  set (p' # bs)" and "q2  set (p' # bs)" by simp_all
  next
    fix j
    assume "Inr j  set (add_spairs ps bs p')"
    hence "Inr j  set ps" by (simp add: add_spairs_def set_merge_wrt Inr_not_in_new_spairs)
    hence "Inr j  set (p # ps)" by simp
    with assms(1) show "j < length fs" by (rule rb_aux_inv_D4)

    fix b
    assume "b  set (p' # bs)"
    hence "b = p'  b  set bs" by simp
    thus "lt b t term_of_pair (0, j)"
    proof
      assume "b = p'"
      hence "lt b = sig_of_pair p" by (simp only: lt_p' sig_of_p)
      also from Inr j  set ps have "... t term_of_pair (0, j)" by (rule sig_of_p_less)
      finally show ?thesis .
    next
      assume "b  set bs"
      with assms(1) Inr j  set (p # ps) show ?thesis by (rule rb_aux_inv_D4)
    qed
  next
    fix j
    assume "Inr j  set (add_spairs ps bs p')"
    hence "Inr j  set ps" by (simp add: add_spairs_def set_merge_wrt Inr_not_in_new_spairs)
    hence "Inr j  set (p # ps)" by simp
    let ?P = "λq. sig_of_pair q = term_of_pair (0, j)"
    have "filter ?P (add_spairs ps bs p') = filter ?P ps" unfolding add_spairs_def
    proof (rule filter_merge_wrt_2)
      fix q
      assume "q  set (new_spairs bs p')"
      then obtain b where "b  set bs" and "is_regular_spair p' b" and "q = Inl (p', b)"
        by (rule in_new_spairsE)
      moreover assume "sig_of_pair q = term_of_pair (0, j)"
      ultimately have "lt (spair p' b) = term_of_pair (0, j)"
        by (simp add: sig_of_spair del: sig_of_pair.simps)
      hence eq: "component_of_term (lt (spair p' b)) = j" by (simp add: component_of_term_of_pair)
      have "component_of_term (lt p') < j"
      proof (rule ccontr)
        assume "¬ component_of_term (lt p') < j"
        hence "component_of_term (term_of_pair (0, j))  component_of_term (lt p')"
          by (simp add: component_of_term_of_pair)
        moreover have "pp_of_term (term_of_pair (0, j))  pp_of_term (lt p')"
          by (simp add: pp_of_term_of_pair zero_min)
        ultimately have "term_of_pair (0, j) t lt p'" using ord_termI by blast
        moreover have "lt p' t term_of_pair (0, j)" unfolding lt_p' sig_of_p[symmetric]
          using Inr j  set ps by (rule sig_of_p_less)
        ultimately show False by simp
      qed
      moreover have "component_of_term (lt b) < j"
      proof (rule ccontr)
        assume "¬ component_of_term (lt b) < j"
        hence "component_of_term (term_of_pair (0, j))  component_of_term (lt b)"
          by (simp add: component_of_term_of_pair)
        moreover have "pp_of_term (term_of_pair (0, j))  pp_of_term (lt b)"
          by (simp add: pp_of_term_of_pair zero_min)
        ultimately have "term_of_pair (0, j) t lt b" using ord_termI by blast
        moreover from assms(1) Inr j  set (p # ps) b  set bs
        have "lt b t term_of_pair (0, j)" by (rule rb_aux_inv_D4)
        ultimately show False by simp
      qed
      ultimately have "component_of_term (lt (spair p' b)) < j"
        using is_regular_spair_component_lt_cases[OF is_regular_spair p' b] by auto
      thus False by (simp add: eq)
    qed
    hence "length (filter ?P (add_spairs ps bs p'))  length (filter ?P (p # ps))"
      by simp
    also from assms(1) Inr j  set (p # ps) have "...  1" by (rule rb_aux_inv_D4)
    finally show "length (filter ?P (add_spairs ps bs p'))  1" .
  next
    from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
    hence "sorted_wrt pair_ord ps" by simp
    thus "sorted_wrt pair_ord (add_spairs ps bs p')" by (rule sorted_add_spairs)
  next
    fix q b1 b2
    assume 1: "q  set (add_spairs ps bs p')" and 2: "is_regular_spair b1 b2"
      and 3: "sig_of_pair q t lt (spair b1 b2)"
    assume "b1  set (p' # bs)" and "b2  set (p' # bs)"
    hence "b1 = p'  b1  set bs" and "b2 = p'  b2  set bs" by simp_all
    thus "Inl (b1, b2)  set (add_spairs ps bs p')  Inl (b2, b1)  set (add_spairs ps bs p')"
    proof (elim disjE)
      assume "b1 = p'" and "b2 = p'"
      with 2 show ?thesis by (simp add: is_regular_spair_def)
    next
      assume "b1 = p'" and "b2  set bs"
      from this(2) 2 have "Inl (b1, b2)  set (new_spairs bs p')" unfolding b1 = p'
        by (rule in_new_spairsI)
      with 2 show ?thesis by (simp add: sig_of_spair add_spairs_def set_merge_wrt image_Un del: sig_of_pair.simps)
    next
      assume "b2 = p'" and "b1  set bs"
      note this(2)
      moreover from 2 have "is_regular_spair b2 b1" by (rule is_regular_spair_sym)
      ultimately have "Inl (b2, b1)  set (new_spairs bs p')" unfolding b2 = p'
        by (rule in_new_spairsI)
      with 2 show ?thesis
        by (simp add: sig_of_spair_commute sig_of_spair add_spairs_def set_merge_wrt image_Un del: sig_of_pair.simps)
    next
      note assms(1) p  set (p # ps)
      moreover assume "b1  set bs" and "b2  set bs"
      moreover note 2
      moreover have 4: "sig_of_pair p t lt (spair b1 b2)"
        by (rule ord_term_lin.le_less_trans, rule sig_merge, fact 1, fact 3)
      ultimately show ?thesis
      proof (rule rb_aux_inv_D6_1)
        assume "Inl (b1, b2)  set (p # ps)"
        with 4 have "Inl (b1, b2)  set ps"
          by (auto simp: sig_of_spair is_regular_spair b1 b2 simp del: sig_of_pair.simps)
        thus ?thesis by (simp add: add_spairs_def set_merge_wrt)
      next
        assume "Inl (b2, b1)  set (p # ps)"
        with 4 have "Inl (b2, b1)  set ps"
          by (auto simp: sig_of_spair sig_of_spair_commute is_regular_spair b1 b2 simp del: sig_of_pair.simps)
        thus ?thesis by (simp add: add_spairs_def set_merge_wrt)
      qed
    qed
  next
    fix q j
    assume "j < length fs"
    assume "q  set (add_spairs ps bs p')"
    hence "sig_of_pair p t sig_of_pair q" by (rule sig_merge)
    also assume "sig_of_pair q t term_of_pair (0, j)"
    finally have 1: "sig_of_pair p t term_of_pair (0, j)" .
    with assms(1) p  set (p # ps) j < length fs have "Inr j  set (p # ps)"
      by (rule rb_aux_inv_D6_2)
    with 1 show "Inr j  set (add_spairs ps bs p')" by (auto simp: add_spairs_def set_merge_wrt)
  next
    fix b q
    assume "b  set (p' # bs)" and q_in: "q  set (add_spairs ps bs p')"
    from this(1) have "b = p'  b  set bs" by simp
    hence "lt b t lt p'"
    proof
      note assms(1)
      moreover assume "b  set bs"
      moreover have "p  set (p # ps)" by simp
      ultimately have "lt b t sig_of_pair p" by (rule rb_aux_inv_D7)
      thus ?thesis by (simp only: lt_p' sig_of_p)
    qed simp
    also have "... = sig_of_pair p" by (simp only: sig_of_p lt_p')
    also from q_in have "... t sig_of_pair q" by (rule sig_merge)
    finally show "lt b t sig_of_pair q" .
  next
    fix a b
    assume 1: "a  set (p' # bs)" and 2: "b  set (p' # bs)" and 3: "is_regular_spair a b"
    assume 6: "¬ is_RB_in dgrad rword (set (p' # bs)) (lt (spair a b))"
    with rb have neq: "lt (spair a b)  lt (poly_of_pair p)" by (auto simp: sig_of_p)
    assume "Inl (a, b)  set (add_spairs ps bs p')"
    hence 40: "Inl (a, b)  set (new_spairs bs p')" and "Inl (a, b)  set ps"
      by (simp_all add: add_spairs_def set_merge_wrt)
    from this(2) neq have 4: "Inl (a, b)  set (p # ps)" by auto
    assume "Inl (b, a)  set (add_spairs ps bs p')"
    hence 50: "Inl (b, a)  set (new_spairs bs p')" and "Inl (b, a)  set ps"
      by (simp_all add: add_spairs_def set_merge_wrt)
    from this(2) neq have 5: "Inl (b, a)  set (p # ps)" by (auto simp: spair_comm[of a])
    have "a  p'"
    proof
      assume "a = p'"
      with 3 have "b  p'" by (auto simp: is_regular_spair_def)
      with 2 have "b  set bs" by simp
      hence "Inl (a, b)  set (new_spairs bs p')" using 3 unfolding a = p' by (rule in_new_spairsI)
      with 40 show False ..
    qed
    with 1 have "a  set bs" by simp
    have "b  p'"
    proof
      assume "b = p'"
      with 3 have "a  p'" by (auto simp: is_regular_spair_def)
      with 1 have "a  set bs" by simp
      moreover from 3 have "is_regular_spair b a" by (rule is_regular_spair_sym)
      ultimately have "Inl (b, a)  set (new_spairs bs p')" unfolding b = p' by (rule in_new_spairsI)
      with 50 show False ..
    qed
    with 2 have "b  set bs" by simp
    have lt_sp: "lt (spair a b) t lt p'"
    proof (rule ord_term_lin.linorder_cases)
      assume "lt (spair a b) = lt p'"
      with neq show ?thesis by (simp add: lt_p')
    next
      assume "lt p' t lt (spair a b)"
      hence "sig_of_pair p t lt (spair a b)" by (simp only: lt_p' sig_of_p)
      with assms(1) p  set (p # ps) a  set bs b  set bs 3 show ?thesis
      proof (rule rb_aux_inv_D6_1)
        assume "Inl (a, b)  set (p # ps)"
        with 4 show ?thesis ..
      next
        assume "Inl (b, a)  set (p # ps)"
        with 5 show ?thesis ..
      qed
    qed
    have "¬ is_RB_in dgrad rword (set bs) (lt (spair a b))"
    proof
      assume "is_RB_in dgrad rword (set bs) (lt (spair a b))"
      hence "is_RB_in dgrad rword (set (p' # bs)) (lt (spair a b))" unfolding set_simps using lt_sp
        by (rule is_RB_in_insertI)
      with 6 show False ..
    qed
    with assms(1) a  set bs b  set bs 3 4 5
    obtain q where "q  set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)" and 9: "¬ sig_crit' bs q"
      by (rule rb_aux_inv_D8)
    from this(1, 2) lt_sp have "q  set ps" by (auto simp: lt_p' sig_of_p)
    show "qset (add_spairs ps bs p'). sig_of_pair q = lt (spair a b)  ¬ sig_crit' (p' # bs) q"
    proof (intro bexI conjI)
      show "¬ sig_crit' (p' # bs) q"
      proof
        assume "sig_crit' (p' # bs) q"
        moreover from lt_sp have "sig_of_pair q t lt p'" by (simp only: 8)
        ultimately have "sig_crit' bs q" by (rule sig_crit'_ConsD)
        with 9 show False ..
      qed
    next
      from q  set ps show "q  set (add_spairs ps bs p')" by (simp add: add_spairs_def set_merge_wrt)
    qed fact
  next
    fix j
    assume "j < length fs"
    assume "Inr j  set (add_spairs ps bs p')"
    hence "Inr j  set ps" by (simp add: add_spairs_def set_merge_wrt)

    show "is_RB_in dgrad rword (set (p' # bs)) (term_of_pair (0, j))"
    proof (cases "term_of_pair (0, j) = sig_of_pair p")
      case True
      with rb show ?thesis by simp
    next
      case False
      with Inr j  set ps have "Inr j  set (p # ps)" by auto
      with assms(1) j < length fs have rb': "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
        by (rule rb_aux_inv_D9)
      have "term_of_pair (0, j) t lt p'"
      proof (rule ord_term_lin.linorder_cases)
        assume "term_of_pair (0, j) = lt p'"
        with False show ?thesis by (simp add: lt_p' sig_of_p)
      next
        assume "lt p' t term_of_pair (0, j)"
        hence "sig_of_pair p t term_of_pair (0, j)" by (simp only: lt_p' sig_of_p)
        with assms(1) p  set (p # ps) j < length fs have "Inr j  set (p # ps)"
          by (rule rb_aux_inv_D6_2)
        with Inr j  set (p # ps) show ?thesis ..
      qed
      with rb' show ?thesis unfolding set_simps by (rule is_RB_in_insertI)
    qed

    show "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set (p' # bs))"
    proof (cases "p = Inr j")
      case True
      show ?thesis
      proof (rule sig_red_zero_idealI, rule sig_red_zeroI)
        from p_red have "(sig_red (≺t) (≼) (set bs))** (monomial 1 (term_of_pair (0, j))) p'"
          by (simp add: True)
        moreover have "set bs  set (p' # bs)" by fastforce
        ultimately have "(sig_red (≺t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) p'"
          by (rule sig_red_rtrancl_mono)
        hence "(sig_red (≼t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) p'"
          by (rule sig_red_rtrancl_sing_regI)
        also have "sig_red (≼t) (≼) (set (p' # bs)) p' 0" unfolding sig_red_def
        proof (intro exI bexI)
          from assms(3) have "rep_list p'  0" by (simp add: p'_def)
          show "sig_red_single (≼t) (≼) p' 0 p' 0"
          proof (rule sig_red_singleI)
            show "rep_list p'  0" by fact
          next
            from rep_list p'  0 have "punit.lt (rep_list p')  keys (rep_list p')"
              by (rule punit.lt_in_keys)
            thus "0 + punit.lt (rep_list p')  keys (rep_list p')" by simp
          next
            from rep_list p'  0 have "punit.lc (rep_list p')  0" by (rule punit.lc_not_0)
            thus "0 = p' - monom_mult (lookup (rep_list p') (0 + punit.lt (rep_list p')) / punit.lc (rep_list p')) 0 p'"
              by (simp add: punit.lc_def[symmetric])
          qed (simp_all add: term_simps)
        qed simp
        finally show "(sig_red (≼t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) 0" .
      qed (fact rep_list_zero)
    next
      case False
      with Inr j  set ps have "Inr j  set (p # ps)" by simp
      with assms(1) j < length fs
      have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` set bs)"
        by (rule rb_aux_inv_D9)
      also have "...  ideal (rep_list ` set (p' # bs))" by (rule ideal.span_mono, fastforce)
      finally show ?thesis .
    qed
  qed

  show "lt p'  lt ` set bs" unfolding lt_p'
  proof
    assume "lt (poly_of_pair p)  lt ` set bs"
    then obtain b where "b  set bs" and "lt (poly_of_pair p) = lt b" ..
    note this(2)
    also from b  set bs have "lt b t lt (poly_of_pair p)" by (rule lt_p_gr)
    finally show False ..
  qed
qed

lemma rb_aux_inv_init: "rb_aux_inv ([], Koszul_syz_sigs fs, map Inr [0..<length fs])"
proof (simp add: rb_aux_inv.simps rb_aux_inv1_def o_def, intro conjI ballI allI impI)
  fix v
  assume "v  set (Koszul_syz_sigs fs)"
  with dgrad fs_distinct fs_nonzero show "is_syz_sig dgrad v" by (rule Koszul_syz_sigs_is_syz_sig)
next
  fix p q :: "'t 0 'b"
  show "Inl (p, q)  Inr ` {0..<length fs}" by blast
next
  fix j
  assume "Inr j  Inr ` {0..<length fs}"
  thus "j < length fs" by fastforce
next
  fix j
  have eq: "(term_of_pair (0, i) = term_of_pair (0, j))  (j = i)" for i
    by (auto dest: term_of_pair_injective)
  show "length (filter (λi. term_of_pair (0, i) = term_of_pair (0, j)) [0..<length fs])  Suc 0"
    by (simp add: eq)
next
  show "sorted_wrt pair_ord (map Inr [0..<length fs])"
  proof (simp add: sorted_wrt_map pair_ord_def sorted_wrt_upt_iff, intro allI impI)
    fix i j :: nat
    assume "i < j"
    hence "i  j" by simp
    show "term_of_pair (0, i) t term_of_pair (0, j)" by (rule ord_termI, simp_all add: term_simps i  j)
  qed
qed

corollary rb_aux_inv_init_fst:
  "rb_aux_inv (fst (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))"
  using rb_aux_inv_init by simp

function (domintros) rb_aux :: "((('t 0 'b) list × 't list × ((('t 0 'b) × ('t 0 'b)) + nat) list) × nat) 
                                    ((('t 0 'b) list × 't list × ((('t 0 'b) × ('t 0 'b)) + nat) list) × nat)"
  where
    "rb_aux ((bs, ss, []), z) = ((bs, ss, []), z)" |
    "rb_aux ((bs, ss, p # ps), z) =
      (let ss' = new_syz_sigs ss bs p in
        if sig_crit bs ss' p then
          rb_aux ((bs, ss', ps), z)
        else
          let p' = sig_trd bs (poly_of_pair p) in
            if rep_list p' = 0 then
              rb_aux ((bs, lt p' # ss', ps), Suc z)
            else
              rb_aux ((p' # bs, ss', add_spairs ps bs p'), z))"
  by pat_completeness auto

definition rb :: "('t 0 'b) list × nat"
  where "rb = (let ((bs, _, _), z) = rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0) in (bs, z))"

text @{const rb} is only an auxiliary function used for stating some theorems about rewrite bases
  and their computation in a readable way. Actual computations (of Gr\"obner bases) are performed
  by function sig_gb›, defined below.
  The second return value of @{const rb} is the number of zero-reductions. It is only needed for
  proving that under certain assumptions, there are no such zero-reductions.›

text ‹Termination›

qualified definition "rb_aux_term1  {(x, y). z. x = z # y}"

qualified definition "rb_aux_term2  {(x, y). (fst x, fst y)  rb_aux_term1 
                    (fst x = fst y  length (snd (snd x)) < length (snd (snd y)))}"

qualified definition "rb_aux_term  rb_aux_term2  {(x, y). rb_aux_inv x  rb_aux_inv y}"

lemma wfp_on_rb_aux_term1: "wfp_on (λx y. (x, y)  rb_aux_term1) (Collect rb_aux_inv1)"
proof (rule wfp_onI_chain, rule, elim exE)
  fix seq'
  assume "i. seq' i  Collect rb_aux_inv1  (seq' (Suc i), seq' i)  rb_aux_term1"
  hence inv: "rb_aux_inv1 (seq' j)" and cons: "b. seq' (Suc j) = b # seq' j" for j
    by (simp_all add: rb_aux_term1_def)
  from this(2) have 1: thesis0 if "j. i < length (seq' j)  thesis0" for i thesis0
    using that by (rule list_seq_indexE_length)

  define seq where "seq = (λi. let j = (SOME k. i < length (seq' k)) in rev (seq' j) ! i)"
  have 2: "seq i = rev (seq' j) ! i" if "i < length (seq' j)" for i j
  proof -
    define k where "k = (SOME k. i < length (seq' k))"
    from that have "i < length (seq' k)" unfolding k_def by (rule someI)
    with cons that have "rev (seq' k) ! i = rev (seq' j) ! i" by (rule list_seq_nth')
    thus ?thesis by (simp add: seq_def k_def[symmetric])
  qed
  have 3: "seq i  set (seq' j)" if "i < length (seq' j)" for i j
  proof -
    from that have "i < length (rev (seq' j))" by simp
    moreover from that have "seq i = rev (seq' j) ! i" by (rule 2)
    ultimately have "seq i  set (rev (seq' j))" by (metis nth_mem)
    thus ?thesis by simp
  qed
  have 4: "seq ` {0..<i} = set (take i (rev (seq' j)))" if "i < length (seq' j)" for i j
  proof -
    from refl have "seq ` {0..<i} = (!) (rev (seq' j)) ` {0..<i}"
    proof (rule image_cong)
      fix i'
      assume "i'  {0..<i}"
      hence "i' < i" by simp
      hence "i' < length (seq' j)" using that by simp
      thus "seq i' = rev (seq' j) ! i'" by (rule 2)
    qed
    also have "... = set (take i (rev (seq' j)))" by (rule nth_image, simp add: that less_imp_le_nat)
    finally show ?thesis .
  qed
  from dgrad show False
  proof (rule rb_termination)
    have "seq i  dgrad_sig_set dgrad" for i
    proof -
      obtain j where "i < length (seq' j)" by (rule 1)
      hence "seq i  set (seq' j)" by (rule 3)
      moreover from inv have "set (seq' j)  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
      ultimately show ?thesis ..
    qed
    thus "range seq  dgrad_sig_set dgrad" by blast
  next
    have "rep_list (seq i)  0" for i
    proof -
      obtain j where "i < length (seq' j)" by (rule 1)
      hence "seq i  set (seq' j)" by (rule 3)
      moreover from inv have "0  rep_list ` set (seq' j)" by (rule rb_aux_inv1_D2)
      ultimately show ?thesis by auto
    qed
    thus "0  rep_list ` range seq" by fastforce
  next
    fix i1 i2 :: nat
    assume "i1 < i2"
    also obtain j where i2: "i2 < length (seq' j)" by (rule 1)
    finally have i1: "i1 < length (seq' j)" .
    from i1 have s1: "seq i1 = rev (seq' j) ! i1" by (rule 2)
    from i2 have s2: "seq i2 = rev (seq' j) ! i2" by (rule 2)
    from inv have "sorted_wrt (λx y. lt y t lt x) (seq' j)" by (rule rb_aux_inv1_D3)
    hence "sorted_wrt (λx y. lt x t lt y) (rev (seq' j))" by (simp add: sorted_wrt_rev)
    moreover note i1 < i2
    moreover from i2 have "i2 < length (rev (seq' j))" by simp
    ultimately have "lt (rev (seq' j) ! i1) t lt (rev (seq' j) ! i2)" by (rule sorted_wrt_nth_less)
    thus "lt (seq i1) t lt (seq i2)" by (simp only: s1 s2)
  next
    fix i
    obtain j where i: "i < length (seq' j)" by (rule 1)
    hence eq1: "seq i = rev (seq' j) ! i" and eq2: "seq ` {0..<i} = set (take i (rev (seq' j)))"
      by (rule 2, rule 4)
    let ?i = "length (seq' j) - Suc i"
    from i have "?i < length (seq' j)" by simp
    with inv have "¬ is_sig_red (≺t) (≼) (set (drop (Suc ?i) (seq' j))) ((seq' j) ! ?i)"
      by (rule rb_aux_inv1_D4)
    thus "¬ is_sig_red (≺t) (≼) (seq ` {0..<i}) (seq i)"
      using i by (simp add: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
 
    from inv ?i < length (seq' j)
    show "(j<length fs. lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j))) 
             punit.lt (rep_list (seq i))  punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) 
         (j k. is_regular_spair (seq j) (seq k)  rep_list (spair (seq j) (seq k))  0 
                lt (seq i) = lt (spair (seq j) (seq k)) 
                punit.lt (rep_list (seq i))  punit.lt (rep_list (spair (seq j) (seq k))))" (is "?l  ?r")
    proof (rule rb_aux_inv1_E)
      fix j0
      assume "j0 < length fs"
        and "lt (seq' j ! (length (seq' j) - Suc i)) = lt (monomial (1::'b) (term_of_pair (0, j0)))"
        and "punit.lt (rep_list (seq' j ! (length (seq' j) - Suc i))) 
             punit.lt (rep_list (monomial 1 (term_of_pair (0, j0))))"
      hence ?l using i by (auto simp: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
      thus ?thesis ..
    next
      fix p q
      assume "p  set (seq' j)"
      then obtain pi where "pi < length (seq' j)" and "p = (seq' j) ! pi" by (metis in_set_conv_nth)
      hence p: "p = seq (length (seq' j) - Suc pi)"
        by (metis "2" p  set (seq' j) diff_Suc_less length_pos_if_in_set length_rev rev_nth rev_rev_ident)
      assume "q  set (seq' j)"
      then obtain qi where "qi < length (seq' j)" and "q = (seq' j) ! qi" by (metis in_set_conv_nth)
      hence q: "q = seq (length (seq' j) - Suc qi)"
        by (metis "2" q  set (seq' j) diff_Suc_less length_pos_if_in_set length_rev rev_nth rev_rev_ident)
      assume "is_regular_spair p q" and "rep_list (spair p q)  0"
        and "lt (seq' j ! (length (seq' j) - Suc i)) = lt (spair p q)"
        and "punit.lt (rep_list (seq' j ! (length (seq' j) - Suc i)))  punit.lt (rep_list (spair p q))"
      hence ?r using i by (auto simp: eq1 eq2 p q rev_nth take_rev Suc_diff_Suc)
      thus ?thesis ..
    qed

    from inv ?i < length (seq' j)
    have "is_RB_upt dgrad rword (set (drop (Suc ?i) (seq' j))) (lt ((seq' j) ! ?i))"
      by (rule rb_aux_inv1_D5)
    with dgrad have "is_sig_GB_upt dgrad (set (drop (Suc ?i) (seq' j))) (lt ((seq' j) ! ?i))"
      by (rule is_RB_upt_is_sig_GB_upt)
    thus "is_sig_GB_upt dgrad (seq ` {0..<i}) (lt (seq i))"
      using i by (simp add: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
  qed
qed

lemma wfp_on_rb_aux_term2: "wfp_on (λx y. (x, y)  rb_aux_term2) (Collect rb_aux_inv)"
proof (rule wfp_onI_min)
  fix x Q
  assume "x  Q" and Q_sub: "Q  Collect rb_aux_inv"
  from this(1) have "fst x  fst ` Q" by (rule imageI)
  have "fst ` Q  Collect rb_aux_inv1"
  proof
    fix y
    assume "y  fst ` Q"
    then obtain z where "z  Q" and y: "y = fst z" by fastforce
    obtain bs ss ps where z: "z = (bs, ss, ps)" by (rule rb_aux_inv.cases)
    from z  Q Q_sub have "rb_aux_inv z" by blast
    thus "y  Collect rb_aux_inv1" by (simp add: y z rb_aux_inv.simps)
  qed
  with wfp_on_rb_aux_term1 fst x  fst ` Q obtain z' where "z'  fst ` Q"
    and z'_min: "y. (y, z')  rb_aux_term1  y  fst ` Q" by (rule wfp_onE_min) blast
  from this(1) obtain z0 where "z0  Q" and z': "z' = fst z0" by fastforce
  define Q0 where "Q0 = {z. z  Q  fst z = fst z0}"
  from z0  Q have "z0  Q0" by (simp add: Q0_def)
  hence "length (snd (snd z0))  length ` snd ` snd ` Q0" by (intro imageI)
  with wf_less obtain n where n1: "n  length ` snd ` snd ` Q0"
    and n2: "n'. n' < n  n'  length ` snd ` snd ` Q0" by (rule wfE_min, blast)
  from n1 obtain z where "z  Q0" and n3: "n = length (snd (snd z))" by fastforce
  have z_min: "y  Q0" if "length (snd (snd y)) < length (snd (snd z))" for y
  proof
    assume "y  Q0"
    hence "length (snd (snd y))  length ` snd ` snd ` Q0" by (intro imageI)
    with n2 have "¬ length (snd (snd y)) < length (snd (snd z))" unfolding n3[symmetric] by blast
    thus False using that ..
  qed
  show "zQ. yCollect rb_aux_inv. (y, z)  rb_aux_term2  y  Q"
  proof (intro bexI ballI impI)
    fix y
    assume "y  Collect rb_aux_inv"
    assume "(y, z)  rb_aux_term2"
    hence "(fst y, fst z)  rb_aux_term1  (fst y = fst z  length (snd (snd y)) < length (snd (snd z)))"
      by (simp add: rb_aux_term2_def)
    thus "y  Q"
    proof
      assume "(fst y, fst z)  rb_aux_term1"
      moreover from z  Q0 have "fst z = fst z0" by (simp add: Q0_def)
      ultimately have "(fst y, z')  rb_aux_term1" by (simp add: rb_aux_term1_def z')
      hence "fst y  fst ` Q" by (rule z'_min)
      thus ?thesis by blast
    next
      assume "fst y = fst z  length (snd (snd y)) < length (snd (snd z))"
      hence "fst y = fst z" and "length (snd (snd y)) < length (snd (snd z))" by simp_all
      from this(2) have "y  Q0" by (rule z_min)
      moreover from z  Q0 have "fst y = fst z0" by (simp add: Q0_def fst y = fst z)
      ultimately show ?thesis by (simp add: Q0_def)
    qed
  next
    from z  Q0 show "z  Q" by (simp add: Q0_def)
  qed
qed

corollary wf_rb_aux_term: "wf rb_aux_term"
proof (rule wfI_min)
  fix x::"('t 0 'b) list × 't list × ((('t 0 'b) × ('t 0 'b)) + nat) list" and Q
  assume "x  Q"
  show "zQ. y. (y, z)  rb_aux_term  y  Q"
  proof (cases "rb_aux_inv x")
    case True
    let ?Q = "Q  Collect rb_aux_inv"
    note wfp_on_rb_aux_term2
    moreover from x  Q True have "x  ?Q" by simp
    moreover have "?Q  Collect rb_aux_inv" by simp
    ultimately obtain z where "z  ?Q" and z_min: "y. (y, z)  rb_aux_term2  y  ?Q"
      by (rule wfp_onE_min) blast
    show ?thesis
    proof (intro bexI allI impI)
      fix y
      assume "(y, z)  rb_aux_term"
      hence "(y, z)  rb_aux_term2" and "rb_aux_inv y" by (simp_all add: rb_aux_term_def)
      from this(1) have "y  ?Q" by (rule z_min)
      with rb_aux_inv y show "y  Q" by simp
    next
      from z  ?Q show "z  Q" by simp
    qed
  next
    case False
    show ?thesis
    proof (intro bexI allI impI)
      fix y
      assume "(y, x)  rb_aux_term"
      hence "rb_aux_inv x" by (simp add: rb_aux_term_def)
      with False show "y  Q" ..
    qed fact
  qed
qed

lemma rb_aux_domI:
  assumes "rb_aux_inv (fst args)"
  shows "rb_aux_dom args"
proof -
  let ?rel = "rb_aux_term <*lex*> ({}::(nat × nat) set)"
  from wf_rb_aux_term wf_empty have "wf ?rel" ..
  thus ?thesis using assms
  proof (induct args)
    case (less args)
    obtain bs ss ps0 z where args: "args = ((bs, ss, ps0), z)" using prod.exhaust by metis
    show ?case
    proof (cases ps0)
      case Nil
      show ?thesis unfolding args Nil by (rule rb_aux.domintros)
    next
      case (Cons p ps)
      from less(1) have 1: "y. (y, ((bs, ss, p # ps), z))  ?rel  rb_aux_inv (fst y)  rb_aux_dom y"
        by (simp only: args Cons)
      from less(2) have 2: "rb_aux_inv (bs, ss, p # ps)" by (simp only: args Cons fst_conv)
      show ?thesis unfolding args Cons
      proof (rule rb_aux.domintros)
        assume "sig_crit bs (new_syz_sigs ss bs p) p"
        with 2 have a: "rb_aux_inv (bs, (new_syz_sigs ss bs p), ps)" by (rule rb_aux_inv_preserved_1)
        with 2 have "((bs, (new_syz_sigs ss bs p), ps), bs, ss, p # ps)  rb_aux_term"
          by (simp add: rb_aux_term_def rb_aux_term2_def)
        hence "(((bs, (new_syz_sigs ss bs p), ps), z), (bs, ss, p # ps), z)  ?rel" by simp
        moreover from a have "rb_aux_inv (fst ((bs, (new_syz_sigs ss bs p), ps), z))"
          by (simp only: fst_conv)
        ultimately show "rb_aux_dom ((bs, (new_syz_sigs ss bs p), ps), z)" by (rule 1)
      next
        assume "rep_list (sig_trd bs (poly_of_pair p)) = 0"
        with 2 have a: "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
          by (rule rb_aux_inv_preserved_2)
        with 2 have "((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), bs, ss, p # ps) 
                      rb_aux_term"
          by (simp add: rb_aux_term_def rb_aux_term2_def)
        hence "(((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z), (bs, ss, p # ps), z) 
                      ?rel" by simp
        moreover from a have "rb_aux_inv (fst ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z))"
          by (simp only: fst_conv)
        ultimately show "rb_aux_dom ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)"
          by (rule 1)
      next
        let ?args = "(sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p, add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
        assume "¬ sig_crit bs (new_syz_sigs ss bs p) p" and "rep_list (sig_trd bs (poly_of_pair p))  0"
        with 2 have a: "rb_aux_inv ?args" by (rule rb_aux_inv_preserved_3)
        with 2 have "(?args, bs, ss, p # ps)  rb_aux_term"
          by (simp add: rb_aux_term_def rb_aux_term2_def rb_aux_term1_def)
        hence "((?args, z), (bs, ss, p # ps), z)  ?rel" by simp
        moreover from a have "rb_aux_inv (fst (?args, z))" by (simp only: fst_conv)
        ultimately show "rb_aux_dom (?args, z)" by (rule 1)
      qed
    qed
  qed
qed

text ‹Invariant›

lemma rb_aux_inv_invariant:
  assumes "rb_aux_inv (fst args)"
  shows "rb_aux_inv (fst (rb_aux args))"
proof -
  from assms have "rb_aux_dom args" by (rule rb_aux_domI)
  thus ?thesis using assms
  proof (induct args rule: rb_aux.pinduct)
    case (1 bs ss z)
    thus ?case by (simp only: rb_aux.psimps(1))
  next
    case (2 bs ss p ps z)
    from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
    show ?case
    proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
      assume a: "sig_crit bs (new_syz_sigs ss bs p) p"
      with * have "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
        by (rule rb_aux_inv_preserved_1)
      hence "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))" by (simp only: fst_conv)
      with refl a show "rb_aux_inv (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))" by (rule 2(2))
      thus "rb_aux_inv (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))" .
    next
      assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p"
      assume b: "rep_list (sig_trd bs (poly_of_pair p)) = 0"
      with * have "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
        by (rule rb_aux_inv_preserved_2)
      hence "rb_aux_inv (fst ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z))"
        by (simp only: fst_conv)
      with refl a refl b
      show "rb_aux_inv (fst (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)))"
        by (rule 2(3))
    next
      let ?args = "(sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
                       add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
      assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p))  0"
      with * have "rb_aux_inv ?args" by (rule rb_aux_inv_preserved_3)
      hence "rb_aux_inv (fst (?args, z))" by (simp only: fst_conv)
      with refl a refl b
      show "rb_aux_inv (fst (rb_aux (?args, z)))"
        by (rule 2(4))
    qed
  qed
qed

lemma rb_aux_inv_last_Nil:
  assumes "rb_aux_dom args"
  shows "snd (snd (fst (rb_aux args))) = []"
  using assms
proof (induct args rule: rb_aux.pinduct)
  case (1 bs ss z)
  thus ?case by (simp add: rb_aux.psimps(1))
next
  case (2 bs ss p ps z)
  show ?case
  proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
    assume "sig_crit bs (new_syz_sigs ss bs p) p"
    with refl show "snd (snd (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) = []"
      and "snd (snd (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) = []"
      by (rule 2(2))+
  next
    assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p)) = 0"
    from refl a refl b
    show "snd (snd (fst (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)))) = []"
      by (rule 2(3))
  next
    assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p))  0"
    from refl a refl b
    show "snd (snd (fst (rb_aux ((sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
                                add_spairs ps bs (sig_trd bs (poly_of_pair p))), z)))) = []"
      by (rule 2(4))
  qed
qed

corollary rb_aux_shape:
  assumes "rb_aux_dom args"
  obtains bs ss z where "rb_aux args = ((bs, ss, []), z)"
proof -
  obtain bs ss ps z where "rb_aux args = ((bs, ss, ps), z)" using prod.exhaust by metis
  moreover from assms have "snd (snd (fst (rb_aux args))) = []" by (rule rb_aux_inv_last_Nil)
  ultimately have "rb_aux args = ((bs, ss, []), z)" by simp
  thus ?thesis ..
qed

lemma rb_aux_is_RB_upt:
  "is_RB_upt dgrad rword (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
proof -
  let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
  from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
  then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
  moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
    by (rule rb_aux_inv_invariant)
  ultimately have "rb_aux_inv (bs, ss, [])" by simp
  have "is_RB_upt dgrad rword (set bs) u" by (rule rb_aux_inv_is_RB_upt, fact, simp)
  thus ?thesis by (simp add: eq)
qed

corollary rb_is_RB_upt: "is_RB_upt dgrad rword (set (fst rb)) u"
  using rb_aux_is_RB_upt[of 0 u] by (auto simp add: rb_def split: prod.split)

corollary rb_aux_is_sig_GB_upt:
  "is_sig_GB_upt dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
  using dgrad rb_aux_is_RB_upt by (rule is_RB_upt_is_sig_GB_upt)

corollary rb_aux_is_sig_GB_in:
  "is_sig_GB_in dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
proof -
  let ?u = "term_of_pair (pp_of_term u, Suc (component_of_term u))"
  have "u t ?u"
  proof (rule ord_term_lin.le_neq_trans)
    show "u t ?u" by (rule ord_termI, simp_all add: term_simps)
  next
    show "u  ?u"
    proof
      assume "u = ?u"
      hence "component_of_term u = component_of_term ?u" by simp
      thus False by (simp add: term_simps)
    qed
  qed
  with rb_aux_is_sig_GB_upt show ?thesis by (rule is_sig_GB_uptD2)
qed

corollary rb_aux_is_Groebner_basis:
  assumes "hom_grading dgrad"
  shows "punit.is_Groebner_basis (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))))"
proof -
  let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
  from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
  then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
  moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
    by (rule rb_aux_inv_invariant)
  ultimately have "rb_aux_inv (bs, ss, [])" by simp
  hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  hence "set (fst (fst (rb_aux ?args)))  dgrad_max_set dgrad" by (simp add: eq dgrad_sig_set'_def)
  with dgrad assms have "punit.is_Groebner_basis (rep_list ` set (fst (fst (rb_aux ?args))))"
    using rb_aux_is_sig_GB_in by (rule is_sig_GB_is_Groebner_basis)
  thus ?thesis by simp
qed

lemma ideal_rb_aux:
  "ideal (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))) =
   ideal (set fs)" (is "ideal ?l = ideal ?r")
proof
  show "ideal ?l  ideal ?r" by (rule ideal.span_subset_spanI, auto simp: rep_list_in_ideal)
next
  show "ideal ?r  ideal ?l"
  proof (rule ideal.span_subset_spanI, rule subsetI)
    fix f
    assume "f  set fs"
    then obtain j where "j < length fs" and f: "f = fs ! j" by (metis in_set_conv_nth)
    let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
    from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
    then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
    moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
      by (rule rb_aux_inv_invariant)
    ultimately have "rb_aux_inv (bs, ss, [])" by simp
    moreover note j < length fs
    moreover have "Inr j  set []" by simp
    ultimately have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal ?l"
      unfolding eq set_map fst_conv by (rule rb_aux_inv_D9)
    thus "f  ideal ?l" by (simp add: rep_list_monomial' j < length fs f)
  qed
qed

corollary ideal_rb: "ideal (rep_list ` set (fst rb)) = ideal (set fs)"
proof -
  have "ideal (rep_list ` set (fst rb)) =
        ideal (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0))))))"
    by (auto simp: rb_def split: prod.splits)
  also have "... = ideal (set fs)" by (fact ideal_rb_aux)
  finally show ?thesis .
qed

lemma
  shows dgrad_max_set_closed_rb_aux:
    "set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) 
      punit_dgrad_max_set dgrad" (is ?thesis1)
  and rb_aux_nonzero:
    "0  set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
      (is ?thesis2)
proof -
  let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
  from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
  then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
  moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
    by (rule rb_aux_inv_invariant)
  ultimately have "rb_aux_inv (bs, ss, [])" by simp
  hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence "set bs  dgrad_sig_set dgrad" and *: "0  rep_list ` set bs"
    by (rule rb_aux_inv1_D1, rule rb_aux_inv1_D2)
  from this(1) have "set bs  dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
  with dgrad show ?thesis1 by (simp add: eq dgrad_max_3)
  from * show ?thesis2 by (simp add: eq)
qed

subsubsection ‹Minimality of the Computed Basis›

lemma rb_aux_top_irred':
  assumes "rword_strict = rw_rat_strict" and "rb_aux_inv (bs, ss, p # ps)"
    and "¬ sig_crit bs (new_syz_sigs ss bs p) p"
  shows "¬ is_sig_red (≼t) (=) (set bs) (sig_trd bs (poly_of_pair p))"
proof -
  have "rword = rw_rat" by (intro ext, simp only: rword_def rw_rat_alt, simp add: assms(1))

  have lt_p: "sig_of_pair p = lt (poly_of_pair p)" by (rule pair_list_sig_of_pair, fact, simp)
  define p' where "p' = sig_trd bs (poly_of_pair p)"
  have red_p: "(sig_red (≺t) (≼) (set bs))** (poly_of_pair p) p'"
    unfolding p'_def by (rule sig_trd_red_rtrancl)
  hence lt_p': "lt p' = sig_of_pair p"
    and lt_p'': "punit.lt (rep_list p')  punit.lt (rep_list (poly_of_pair p))"
    unfolding lt_p by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
  have "¬ is_sig_red (=) (=) (set bs) p'"
  proof
    assume "is_sig_red (=) (=) (set bs) p'"
    then obtain b where "b  set bs" and "rep_list b  0" and "rep_list p'  0"
      and 1: "punit.lt (rep_list b) adds punit.lt (rep_list p')"
      and 2: "punit.lt (rep_list p')  lt b = punit.lt (rep_list b)  lt p'"
      by (rule is_sig_red_top_addsE)
    note this(3)
    moreover from red_p have "(punit.red (rep_list ` set bs))** (rep_list (poly_of_pair p)) (rep_list p')"
      by (rule sig_red_red_rtrancl)
    ultimately have "rep_list (poly_of_pair p)  0" by (auto simp: punit.rtrancl_0)

    define x where "x = punit.lt (rep_list p') - punit.lt (rep_list b)"
    from 1 2 have x1: "x  lt b = lt p'" by (simp add: term_is_le_rel_minus x_def)
    from this[symmetric] have "lt b addst sig_of_pair p" unfolding lt_p' by (rule adds_termI)
    from 1 have x2: "x + punit.lt (rep_list b) = punit.lt (rep_list p')" by (simp add: x_def adds_minus)
    from rep_list b  0 have "b  0" by (auto simp: rep_list_zero)

    show False
    proof (rule sum_prodE)
      fix a0 b0
      assume p: "p = Inl (a0, b0)"
      hence "Inl (a0, b0)  set (p # ps)" by simp
      with assms(2) have reg: "is_regular_spair a0 b0" and "a0  set bs" and "b0  set bs"
        by (rule rb_aux_inv_D3)+
      from assms(2) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
      hence "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
      with a0  set bs b0  set bs have "rep_list a0  0" and "rep_list b0  0" by fastforce+
      hence "a0  0" and "b0  0" by (auto simp: rep_list_zero)

      let ?t1 = "punit.lt (rep_list a0)"
      let ?t2 = "punit.lt (rep_list b0)"
      let ?l = "lcs ?t1 ?t2"
      from rep_list (poly_of_pair p)  0 have "punit.spoly (rep_list a0) (rep_list b0)  0"
        by (simp add: p rep_list_spair)
      with rep_list a0  0 rep_list b0  0
      have "punit.lt (punit.spoly (rep_list a0) (rep_list b0))  ?l"
        by (rule punit.lt_spoly_less_lcs[simplified])

      obtain b' where 3: "is_canon_rewriter rword (set bs) (sig_of_pair p) b'"
        and 4: "punit.lt (rep_list (poly_of_pair p)) 
                      (pp_of_term (sig_of_pair p) - lp b') + punit.lt (rep_list b')"
      proof (cases "(?l - ?t1)  lt a0 t (?l - ?t2)  lt b0")
        case True
        have "sig_of_pair p = lt (spair a0 b0)" unfolding lt_p by (simp add: p)
        also from reg have "... = (?l - ?t2)  lt b0"
          by (simp add: True is_regular_spair_lt ord_term_lin.max_def)
        finally have eq1: "sig_of_pair p = (?l - ?t2)  lt b0" .
        hence "lt b0 addst sig_of_pair p" by (rule adds_termI)
        moreover from assms(3) have "¬ is_rewritable bs b0 ((?l - ?t2)  lt b0)"
          by (simp add: p spair_sigs_def Let_def)
        ultimately have "is_canon_rewriter rword (set bs) (sig_of_pair p) b0"
          unfolding eq1[symmetric] using inv1 b0  set bs b0  0 is_rewritableI_is_canon_rewriter
          by blast
        thus ?thesis
        proof
          have "punit.lt (rep_list (poly_of_pair p)) = punit.lt (punit.spoly (rep_list a0) (rep_list b0))"
            by (simp add: p rep_list_spair)
          also have "...  ?l" by fact
          also have "... = (?l - ?t2) + ?t2" by (simp only: adds_minus adds_lcs_2)
          also have "... = (pp_of_term (sig_of_pair p) - lp b0) + ?t2"
            by (simp only: eq1 pp_of_term_splus add_diff_cancel_right')
          finally show "punit.lt (rep_list (poly_of_pair p))  pp_of_term (sig_of_pair p) - lp b0 + ?t2" .
        qed
      next
        case False
        have "sig_of_pair p = lt (spair a0 b0)" unfolding lt_p by (simp add: p)
        also from reg have "... = (?l - ?t1)  lt a0"
          by (simp add: False is_regular_spair_lt ord_term_lin.max_def)
        finally have eq1: "sig_of_pair p = (?l - ?t1)  lt a0" .
        hence "lt a0 addst sig_of_pair p" by (rule adds_termI)
        moreover from assms(3) have "¬ is_rewritable bs a0 ((?l - ?t1)  lt a0)"
          by (simp add: p spair_sigs_def Let_def)
        ultimately have "is_canon_rewriter rword (set bs) (sig_of_pair p) a0"
          unfolding eq1[symmetric] using inv1 a0  set bs a0  0 is_rewritableI_is_canon_rewriter
          by blast
        thus ?thesis
        proof
          have "punit.lt (rep_list (poly_of_pair p)) = punit.lt (punit.spoly (rep_list a0) (rep_list b0))"
            by (simp add: p rep_list_spair)
          also have "...  ?l" by fact
          also have "... = (?l - ?t1) + ?t1" by (simp only: adds_minus adds_lcs)
          also have "... = (pp_of_term (sig_of_pair p) - lp a0) + ?t1"
            by (simp only: eq1 pp_of_term_splus add_diff_cancel_right')
          finally show "punit.lt (rep_list (poly_of_pair p))  pp_of_term (sig_of_pair p) - lp a0 + ?t1" .
        qed
      qed

      define y where "y = pp_of_term (sig_of_pair p) - lp b'"
      from lt_p'' 4 have y2: "punit.lt (rep_list p')  y + punit.lt (rep_list b')"
        unfolding y_def by (rule ordered_powerprod_lin.le_less_trans)
      from 3 have "lt b' addst sig_of_pair p" by (rule is_canon_rewriterD3)
      hence "lp b' adds lp p'" and "component_of_term (lt b') = component_of_term (lt p')"
        by (simp_all add: adds_term_def lt_p')
      hence y1: "y  lt b' = lt p'" by (simp add: y_def splus_def lt_p' adds_minus term_simps)
      from 3 b  set bs b  0 lt b addst sig_of_pair p
      have "rword (spp_of b) (spp_of b')" by (rule is_canon_rewriterD)
      hence "punit.lt (rep_list b')  lt b t punit.lt (rep_list b)  lt b'"
        by (auto simp: rword = rw_rat rw_rat_def Let_def spp_of_def)
      hence "(x + y)  (punit.lt (rep_list b')  lt b) t (x + y)  (punit.lt (rep_list b)  lt b')"
        by (rule splus_mono)
      hence "(y + punit.lt (rep_list b'))  (x  lt b) t (x + punit.lt (rep_list b))  (y  lt b')"
        by (simp add: ac_simps)
      hence "(y + punit.lt (rep_list b'))  lt p' t punit.lt (rep_list p')  lt p'"
        by (simp only: x1 x2 y1)
      hence "y + punit.lt (rep_list b')  punit.lt (rep_list p')" by (rule ord_term_canc_left)
      with y2 show ?thesis by simp
    next
      fix j
      assume p: "p = Inr j"
      hence "lt p' = term_of_pair (0, j)" by (simp add: lt_p')
      with x1 term_of_pair_pair[of "lt b"] have "lt b = term_of_pair (0, j)"
        by (auto simp: splus_def dest!: term_of_pair_injective plus_eq_zero_2)
      moreover have "lt b t term_of_pair (0, j)" by (rule rb_aux_inv_D4, fact, simp add: p, fact)
      ultimately show ?thesis by simp
    qed
  qed
  moreover have "¬ is_sig_red (≺t) (=) (set bs) p'"
  proof
    assume "is_sig_red (≺t) (=) (set bs) p'"
    hence "is_sig_red (≺t) (≼) (set bs) p'" by (simp add: is_sig_red_top_tail_cases)
    with sig_trd_irred show False unfolding p'_def ..
  qed
  ultimately show ?thesis by (simp add: p'_def is_sig_red_sing_reg_cases)
qed

lemma rb_aux_top_irred:
  assumes "rword_strict = rw_rat_strict" and "rb_aux_inv (fst args)" and "b  set (fst (fst (rb_aux args)))"
    and "b0. b0  set (fst (fst args))  ¬ is_sig_red (≼t) (=) (set (fst (fst args)) - {b0}) b0"
  shows "¬ is_sig_red (≼t) (=) (set (fst (fst (rb_aux args))) - {b}) b"
proof -
  from assms(2) have "rb_aux_dom args" by (rule rb_aux_domI)
  thus ?thesis using assms(2, 3, 4)
  proof (induct args rule: rb_aux.pinduct)
    case (1 bs ss z)
    let ?nil = "[]::((('t 0 'b) × ('t 0 'b)) + nat) list"
    from 1(3) have "b  set (fst (fst ((bs, ss, ?nil), z)))" by (simp add: rb_aux.psimps(1)[OF 1(1)])
    hence "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, ss, ?nil), z))) - {b}) b" by (rule 1(4))
    thus ?case by (simp add: rb_aux.psimps(1)[OF 1(1)])
  next
    case (2 bs ss p ps z)
    from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
    define p' where "p' = sig_trd bs (poly_of_pair p)"
    from 2(6) show ?case
    proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def p'_def[symmetric] split: if_splits)
      note refl
      moreover assume "sig_crit bs (new_syz_sigs ss bs p) p"
      moreover from * this have "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))"
        unfolding fst_conv by (rule rb_aux_inv_preserved_1)
      moreover assume "b  set (fst (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z))))"
      ultimately show "¬ is_sig_red (≼t) (=) (set (fst (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) - {b}) b"
      proof (rule 2(2))
        fix b0
        assume "b0  set (fst (fst ((bs, new_syz_sigs ss bs p, ps), z)))"
        hence "b0  set (fst (fst ((bs, ss, p # ps), z)))" by simp
        hence "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
        thus "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, new_syz_sigs ss bs p, ps), z))) - {b0}) b0"
          by simp
      qed
    next
      note refl
      moreover assume "¬ sig_crit bs (new_syz_sigs ss bs p) p"
      moreover note refl
      moreover assume "rep_list p' = 0"
      moreover from * this have "rb_aux_inv (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))"
        unfolding p'_def fst_conv by (rule rb_aux_inv_preserved_2)
      moreover assume "b  set (fst (fst (rb_aux ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))))"
      ultimately show "¬ is_sig_red (≼t) (=) (set (fst (fst (rb_aux ((bs,
                                                        lt p' # new_syz_sigs ss bs p, ps), Suc z)))) - {b}) b"
      proof (rule 2(3)[simplified p'_def[symmetric]])
        fix b0
        assume "b0  set (fst (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z)))"
        hence "b0  set (fst (fst ((bs, ss, p # ps), z)))" by simp
        hence "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
        thus "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))) - {b0}) b0"
          by simp
      qed
    next
      note refl
      moreover assume "¬ sig_crit bs (new_syz_sigs ss bs p) p"
      moreover note refl
      moreover assume "rep_list p'  0"
      moreover from * ¬ sig_crit bs (new_syz_sigs ss bs p) p this
      have inv: "rb_aux_inv (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))"
        unfolding p'_def fst_conv by (rule rb_aux_inv_preserved_3)
      moreover assume "b  set (fst (fst (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))))"
      ultimately show "¬ is_sig_red (≼t) (=) (set (fst (fst (rb_aux ((p' # bs,
                                              new_syz_sigs ss bs p, add_spairs ps bs p'), z)))) - {b}) b"
      proof (rule 2(4)[simplified p'_def[symmetric]])
        fix b0
        assume "b0  set (fst (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)))"
        hence "b0 = p'  b0  set bs" by simp
        hence "¬ is_sig_red (≼t) (=) (({p'} - {b0})  (set bs - {b0})) b0"
        proof
          assume "b0 = p'"
          have "¬ is_sig_red (≼t) (=) (set bs - {b0}) p'"
          proof
            assume "is_sig_red (≼t) (=) (set bs - {b0}) p'"
            moreover have "set bs - {b0}  set bs" by fastforce
            ultimately have "is_sig_red (≼t) (=) (set bs) p'" by (rule is_sig_red_mono)
            moreover have "¬ is_sig_red (≼t) (=) (set bs) p'" unfolding p'_def
              using assms(1) * ¬ sig_crit bs (new_syz_sigs ss bs p) p by (rule rb_aux_top_irred')
            ultimately show False by simp
          qed
          thus ?thesis by (simp add: b0 = p')
        next
          assume "b0  set bs"
          hence "b0  set (fst (fst ((bs, ss, p # ps), z)))" by simp
          hence "¬ is_sig_red (≼t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
          hence "¬ is_sig_red (≼t) (=) (set bs - {b0}) b0" by simp
          moreover have "¬ is_sig_red (≼t) (=) ({p'} - {b0}) b0"
          proof
            assume "is_sig_red (≼t) (=) ({p'} - {b0}) b0"
            moreover have "{p'} - {b0}  {p'}" by fastforce
            ultimately have "is_sig_red (≼t) (=) {p'} b0" by (rule is_sig_red_mono)
            hence "lt p' t lt b0" by (rule is_sig_redD_lt)

            from inv have "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
              by (simp only: fst_conv)
            hence "rb_aux_inv1 (p' # bs)" by (rule rb_aux_inv_D1)
            hence "sorted_wrt (λx y. lt y t lt x) (p' # bs)" by (rule rb_aux_inv1_D3)
            with b0  set bs have "lt b0 t lt p'" by simp
            with lt p' t lt b0 show False by simp
          qed
          ultimately show ?thesis by (simp add: is_sig_red_Un)
        qed
        thus "¬ is_sig_red (≼t) (=) (set (fst (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))) - {b0}) b0"
          by (simp add: Un_Diff[symmetric])
      qed
    qed
  qed
qed

corollary rb_aux_is_min_sig_GB:
  assumes "rword_strict = rw_rat_strict"
  shows "is_min_sig_GB dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
    (is "is_min_sig_GB _ (set (fst (fst (rb_aux ?args))))")
  unfolding is_min_sig_GB_def
proof (intro conjI allI ballI impI)
  from rb_aux_inv_init_fst have inv: "rb_aux_inv (fst (rb_aux ?args))"
    and "rb_aux_dom ?args"
    by (rule rb_aux_inv_invariant, rule rb_aux_domI)
  from this(2) obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')"
    by (rule rb_aux_shape)
  from inv have "rb_aux_inv (bs, ss, [])" by (simp only: eq fst_conv)
  hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
  thus "set (fst (fst (rb_aux ?args)))  dgrad_sig_set dgrad" by (simp add: eq)
next
  fix u
  show "is_sig_GB_in dgrad (set (fst (fst (rb_aux ?args)))) u" by (fact rb_aux_is_sig_GB_in)
next
  fix g
  assume "g  set (fst (fst (rb_aux ?args)))"
  with assms(1) rb_aux_inv_init_fst
  show "¬ is_sig_red (≼t) (=) (set (fst (fst (rb_aux ?args))) - {g}) g"
    by (rule rb_aux_top_irred) simp
qed

corollary rb_is_min_sig_GB:
  assumes "rword_strict = rw_rat_strict"
  shows "is_min_sig_GB dgrad (set (fst rb))"
  using rb_aux_is_min_sig_GB[OF assms, of 0] by (auto simp: rb_def split: prod.split)

subsubsection ‹No Zero-Reductions›

fun rb_aux_inv2 :: "(('t 0 'b) list × 't list × ((('t 0 'b) × ('t 0 'b)) + nat) list)  bool"
  where "rb_aux_inv2 (bs, ss, ps) =
          (rb_aux_inv (bs, ss, ps) 
           (j<length fs. Inr j  set ps 
               (fs ! j  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) bs)) 
               (bset bs. component_of_term (lt b) < j 
                  (sset ss. s addst term_of_pair (punit.lt (rep_list b), j))))))"

lemma rb_aux_inv2_D1: "rb_aux_inv2 args  rb_aux_inv args"
  by (metis prod.exhaust rb_aux_inv2.simps)

lemma rb_aux_inv2_D2:
  "rb_aux_inv2 (bs, ss, ps)  j < length fs  Inr j  set ps 
    fs ! j  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) bs))"
  by simp

lemma rb_aux_inv2_E:
  assumes "rb_aux_inv2 (bs, ss, ps)" and "j < length fs" and "Inr j  set ps" and "b  set bs"
    and "component_of_term (lt b) < j"
  obtains s where "s  set ss" and "s addst term_of_pair (punit.lt (rep_list b), j)"
  using assms by auto

context
  assumes pot: "is_pot_ord"
begin

lemma sig_red_zero_filter:
  assumes "sig_red_zero (≼t) (set bs) r" and "component_of_term (lt r) < j"
  shows "sig_red_zero (≼t) (set (filter (λb. component_of_term (lt b) < j) bs)) r"
proof -
  have "(≼t) = (≼t)  (≼t) = (≺t)" by simp
  with assms(1) have "sig_red_zero (≼t) {bset bs. lt b t lt r} r" by (rule sig_red_zero_subset)
  moreover have "{bset bs. lt b t lt r}  set (filter (λb. component_of_term (lt b) < j) bs)"
  proof
    fix b
    assume "b  {bset bs. lt b t lt r}"
    hence "b  set bs" and "lt b t lt r" by simp_all
    from pot this(2) have "component_of_term (lt b)  component_of_term (lt r)" by (rule is_pot_ordD2)
    also have "... < j" by (fact assms(2))
    finally have "component_of_term (lt b) < j" .
    with b  set bs show "b  set (filter (λb. component_of_term (lt b) < j) bs)" by simp
  qed
  ultimately show ?thesis by (rule sig_red_zero_mono)
qed

lemma rb_aux_inv2_preserved_0:
  assumes "rb_aux_inv2 (bs, ss, p # ps)" and "j < length fs" and "Inr j  set ps"
    and "b  set bs" and "component_of_term (lt b) < j"
  shows "sset (new_syz_sigs ss bs p). s addst term_of_pair (punit.lt (rep_list b), j)"
proof (rule sum_prodE)
  fix x y
  assume p: "p = Inl (x, y)"
  with assms(3) have "Inr j  set (p # ps)" by simp
  with assms(1, 2) obtain s where "s  set ss" and *: "s addst term_of_pair (punit.lt (rep_list b), j)"
    using assms(4, 5) by (rule rb_aux_inv2_E)
  from this(1) have "s  set (new_syz_sigs ss bs p)" by (simp add: p)
  with * show ?thesis ..
next
  fix i
  assume p: "p = Inr i"
  have trans: "transp (addst)" by (rule transpI, drule adds_term_trans)
  from adds_term_refl have refl: "reflp (addst)" by (rule reflpI)
  let ?v = "term_of_pair (punit.lt (rep_list b), j)"
  let ?f = "λb. term_of_pair (punit.lt (rep_list b), i)"
  define ss' where "ss' = filter_min (addst) (map ?f bs)"
  have eq: "new_syz_sigs ss bs p = filter_min_append (addst) ss ss'" by (simp add: p ss'_def pot)
  show ?thesis
  proof (cases "i = j")
    case True
    from b  set bs have "?v  ?f ` set bs" unfolding i = j by (rule imageI)
    hence "?v  set ss  set (map ?f bs)" by simp
    thus ?thesis
    proof
      assume "?v  set ss"
      hence "?v  set ss  set ss'" by simp
      with trans refl obtain s where "s  set (new_syz_sigs ss bs p)" and "s addst ?v"
        unfolding eq by (rule filter_min_append_relE)
      thus ?thesis ..
    next
      assume "?v  set (map ?f bs)"
      with trans refl obtain s where "s  set ss'" and "s addst ?v"
        unfolding ss'_def by (rule filter_min_relE)
      from this(1) have "s  set ss  set ss'" by simp
      with trans refl obtain s' where s': "s'  set (new_syz_sigs ss bs p)" and "s' addst s"
        unfolding eq by (rule filter_min_append_relE)
      from this(2) s addst ?v have "s' addst ?v" by (rule adds_term_trans)
      with s' show ?thesis ..
    qed
  next
    case False
    with assms(3) have "Inr j  set (p # ps)" by (simp add: p)
    with assms(1, 2) obtain s where "s  set ss" and "s addst ?v"
      using assms(4, 5) by (rule rb_aux_inv2_E)
    from this(1) have "s  set ss  set (map ?f bs)" by simp
    thus ?thesis
    proof
      assume "s  set ss"
      hence "s  set ss  set ss'" by simp
      with trans refl obtain s' where s': "s'  set (new_syz_sigs ss bs p)" and "s' addst s"
        unfolding eq by (rule filter_min_append_relE)
      from this(2) s addst ?v have "s' addst ?v" by (rule adds_term_trans)
      with s' show ?thesis ..
    next
      assume "s  set (map ?f bs)"
      with trans refl obtain s' where "s'  set ss'" and "s' addst s"
        unfolding ss'_def by (rule filter_min_relE)
      from this(1) have "s'  set ss  set ss'" by simp
      with trans refl obtain s'' where s'': "s''  set (new_syz_sigs ss bs p)" and "s'' addst s'"
        unfolding eq by (rule filter_min_append_relE)
      from this(2) s' addst s have "s'' addst s" by (rule adds_term_trans)
      hence "s'' addst ?v" using s addst ?v by (rule adds_term_trans)
      with s'' show ?thesis ..
    qed
  qed
qed

lemma rb_aux_inv2_preserved_1:
  assumes "rb_aux_inv2 (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
  shows "rb_aux_inv2 (bs, new_syz_sigs ss bs p, ps)"
  unfolding rb_aux_inv2.simps
proof (intro allI conjI impI ballI)
  from assms(1) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
  thus "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
    using assms(2) by (rule rb_aux_inv_preserved_1)

  fix j
  assume "j < length fs" and "Inr j  set ps"
  show "fs ! j  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) bs))"
  proof (cases "p = Inr j")
    case True
    with assms(2) have "is_pred_syz (new_syz_sigs ss bs p) (term_of_pair (0, j))" by simp
    let ?X = "set (filter (λb. component_of_term (lt b) < Suc j) bs)"
    have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` ?X)"
    proof (rule sig_red_zero_idealI)
      have "sig_red_zero (≺t) (set bs) (monomial 1 (term_of_pair (0, j)))"
      proof (rule syzygy_crit)
        from inv have "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
          by (rule rb_aux_inv_is_RB_upt_Cons)
        with dgrad have "is_sig_GB_upt dgrad (set bs) (sig_of_pair p)"
          by (rule is_RB_upt_is_sig_GB_upt)
        thus "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))" by (simp add: p = Inr j)
      next
        show "monomial 1 (term_of_pair (0, j))  dgrad_sig_set dgrad"
          by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 j < length fs)
      next
        show "lt (monomial (1::'b) (term_of_pair (0, j))) = term_of_pair (0, j)" by (simp add: lt_monomial)
      next
        from inv assms(2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
        thus "is_syz_sig dgrad (term_of_pair (0, j))" by (simp add: p = Inr j)
      qed (fact dgrad)
      hence "sig_red_zero (≼t) (set bs) (monomial 1 (term_of_pair (0, j)))"
        by (rule sig_red_zero_sing_regI)
      moreover have "component_of_term (lt (monomial (1::'b) (term_of_pair (0, j)))) < Suc j"
        by (simp add: lt_monomial component_of_term_of_pair)
      ultimately show "sig_red_zero (≼t) ?X (monomial 1 (term_of_pair (0, j)))"
        by (rule sig_red_zero_filter)
    qed
    thus ?thesis by (simp add: rep_list_monomial' j < length fs)
  next
    case False
    with Inr j  set ps have "Inr j  set (p # ps)" by simp
    with assms(1) j < length fs show ?thesis by (rule rb_aux_inv2_D2)
  qed
next
  fix j b
  assume "j < length fs" and "Inr j  set ps" and "b  set bs" and "component_of_term (lt b) < j"
  with assms(1) show "sset (new_syz_sigs ss bs p). s addst term_of_pair (punit.lt (rep_list b), j)"
    by (rule rb_aux_inv2_preserved_0)
qed

lemma rb_aux_inv2_preserved_3:
  assumes "rb_aux_inv2 (bs, ss, p # ps)" and "¬ sig_crit bs (new_syz_sigs ss bs p) p"
    and "rep_list (sig_trd bs (poly_of_pair p))  0"
  shows "rb_aux_inv2 (sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
                          add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
proof -
  from assms(1) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
  define p' where "p' = sig_trd bs (poly_of_pair p)"
  from sig_trd_red_rtrancl[of bs "poly_of_pair p"] have "lt p' = lt (poly_of_pair p)"
    unfolding p'_def by (rule sig_red_regular_rtrancl_lt)
  also have "... = sig_of_pair p" by (rule sym, rule pair_list_sig_of_pair, fact inv, simp)
  finally have lt_p': "lt p' = sig_of_pair p" .
  show ?thesis unfolding rb_aux_inv2.simps p'_def[symmetric]
  proof (intro allI conjI impI ballI)
    show "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
      unfolding p'_def using inv assms(2, 3) by (rule rb_aux_inv_preserved_3)
  next
    fix j
    assume "j < length fs" and *: "Inr j  set (add_spairs ps bs p')"
    show "fs ! j  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) (p' # bs)))"
    proof (cases "p = Inr j")
      case True
      let ?X = "set (filter (λb. component_of_term (lt b) < Suc j) (p' # bs))"
      have "rep_list (monomial 1 (term_of_pair (0, j)))  ideal (rep_list ` ?X)"
      proof (rule sig_red_zero_idealI)
        have "sig_red_zero (≼t) (set (p' # bs)) (monomial 1 (term_of_pair (0, j)))"
        proof (rule sig_red_zeroI)
          have "(sig_red (≺t) (≼) (set bs))** (monomial 1 (term_of_pair (0, j))) p'"
            using sig_trd_red_rtrancl[of bs "poly_of_pair p"] by (simp add: True p'_def)
          moreover have "set bs  set (p' # bs)" by fastforce
          ultimately have "(sig_red (≺t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) p'"
            by (rule sig_red_rtrancl_mono)
          hence "(sig_red (≼t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) p'"
            by (rule sig_red_rtrancl_sing_regI)
          also have "sig_red (≼t) (≼) (set (p' # bs)) p' 0" unfolding sig_red_def
          proof (intro exI bexI)
            from assms(3) have "rep_list p'  0" by (simp add: p'_def)
            show "sig_red_single (≼t) (≼) p' 0 p' 0"
            proof (rule sig_red_singleI)
              show "rep_list p'  0" by fact
            next
              from rep_list p'  0 have "punit.lt (rep_list p')  keys (rep_list p')"
                by (rule punit.lt_in_keys)
              thus "0 + punit.lt (rep_list p')  keys (rep_list p')" by simp
            next
              from rep_list p'  0 have "punit.lc (rep_list p')  0" by (rule punit.lc_not_0)
              thus "0 = p' - monom_mult (lookup (rep_list p') (0 + punit.lt (rep_list p')) / punit.lc (rep_list p')) 0 p'"
                by (simp add: punit.lc_def[symmetric])
            qed (simp_all add: term_simps)
          qed simp
          finally show "(sig_red (≼t) (≼) (set (p' # bs)))** (monomial 1 (term_of_pair (0, j))) 0" .
        qed (fact rep_list_zero)
        moreover have "component_of_term (lt (monomial (1::'b) (term_of_pair (0, j)))) < Suc j"
          by (simp add: lt_monomial component_of_term_of_pair)
        ultimately show "sig_red_zero (≼t) ?X (monomial 1 (term_of_pair (0, j)))"
          by (rule sig_red_zero_filter)
      qed
      thus ?thesis by (simp add: rep_list_monomial' j < length fs)
    next
      case False
      from * have "Inr j  set ps" by (simp add: add_spairs_def set_merge_wrt)
      hence "Inr j  set (p # ps)" using False by simp
      with assms(1) j < length fs
      have "fs ! j  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) bs))"
        by (rule rb_aux_inv2_D2)
      also have "...  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc j) (p' # bs)))"
        by (intro ideal.span_mono image_mono, fastforce)
      finally show ?thesis .
    qed
  next
    fix j and b::"'t 0 'b"
    assume "j < length fs" and *: "component_of_term (lt b) < j"
    assume "Inr j  set (add_spairs ps bs p')"
    hence "Inr j  set ps" by (simp add: add_spairs_def set_merge_wrt)
    assume "b  set (p' # bs)"
    hence "b = p'  b  set bs" by simp
    thus "sset (new_syz_sigs ss bs p). s addst term_of_pair (punit.lt (rep_list b), j)"
    proof
      assume "b = p'"
      with * have "component_of_term (sig_of_pair p) < component_of_term (term_of_pair (0, j))"
        by (simp only: lt_p' component_of_term_of_pair)
      with pot have **: "sig_of_pair p t term_of_pair (0, j)" by (rule is_pot_ordD)
      have "p  set (p # ps)" by simp
      with inv have "Inr j  set (p # ps)" using j < length fs ** by (rule rb_aux_inv_D6_2)
      with Inr j  set ps have "p = Inr j" by simp
      with ** show ?thesis by simp
    next
      assume "b  set bs"
      with assms(1) j < length fs Inr j  set ps show ?thesis
        using * by (rule rb_aux_inv2_preserved_0)
    qed
  qed
qed

lemma rb_aux_inv2_ideal_subset:
  assumes "rb_aux_inv2 (bs, ss, ps)" and "p0. p0  set ps  j  component_of_term (sig_of_pair p0)"
  shows "ideal (set (take j fs))  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < j) bs))"
          (is "ideal ?B  ideal ?A")
proof (intro ideal.span_subset_spanI subsetI)
  fix f
  assume "f  ?B"
  then obtain i where "i < length (take j fs)" and "f = (take j fs) ! i"
    by (metis in_set_conv_nth)
  hence "i < length fs" and "i < j" and f: "f = fs ! i" by auto
  from this(2) have "Suc i  j" by simp
  have "f  ideal (rep_list ` set (filter (λb. component_of_term (lt b) < Suc i) bs))"
    unfolding f using assms(1) i < length fs
  proof (rule rb_aux_inv2_D2)
    show "Inr i  set ps"
    proof
      assume "Inr i  set ps"
      hence "j  component_of_term (sig_of_pair (Inr i))" by (rule assms(2))
      hence "j  i" by (simp add: component_of_term_of_pair)
      with i < j show False by simp
    qed
  qed
  also have "...  ideal ?A"
    by (intro ideal.span_mono image_mono, auto dest: order_less_le_trans[OF _ Suc i  j])
  finally show "f  ideal ?A" .
qed

lemma rb_aux_inv_is_Groebner_basis:
  assumes "hom_grading dgrad" and "rb_aux_inv (bs, ss, ps)"
    and "p0. p0  set ps  j  component_of_term (sig_of_pair p0)"
  shows "punit.is_Groebner_basis (rep_list ` set (filter (λb. component_of_term (lt b) < j) bs))"
          (is "punit.is_Groebner_basis (rep_list ` set ?bs)")
  using dgrad assms(1)
proof (rule is_sig_GB_upt_is_Groebner_basis)
  show "set ?bs  dgrad_sig_set' j dgrad"
  proof
    fix b
    assume "b  set ?bs"
    hence "b  set bs" and "component_of_term (lt b) < j" by simp_all
    show "b  dgrad_sig_set' j dgrad" unfolding dgrad_sig_set'_def
    proof
      from assms(2) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
      hence "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
      with b  set bs have "b  dgrad_sig_set dgrad" ..
      thus "b  dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
    next
      show "b  sig_inv_set' j"
      proof (rule sig_inv_setI')
        fix v
        assume "v  keys b"
        hence "v t lt b" by (rule lt_max_keys)
        with pot have "component_of_term v  component_of_term (lt b)" by (rule is_pot_ordD2)
        also have "... < j" by fact
        finally show "component_of_term v < j" .
      qed
    qed
  qed
next
  fix u
  assume u: "component_of_term u < j"
  from dgrad have "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))"
  proof (rule is_RB_upt_is_sig_GB_upt)
    from assms(2) show "is_RB_upt dgrad rword (set bs) (term_of_pair (0, j))"
    proof (rule rb_aux_inv_is_RB_upt)
      fix p
      assume "p  set ps"
      hence "j  component_of_term (sig_of_pair p)" by (rule assms(3))
      with pot show "term_of_pair (0, j) t sig_of_pair p"
        by (auto simp: is_pot_ord term_simps zero_min)
    qed
  qed
  moreover from pot have "u t term_of_pair (0, j)"
    by (rule is_pot_ordD) (simp only: u component_of_term_of_pair)
  ultimately have 1: "is_sig_GB_in dgrad (set bs) u" by (rule is_sig_GB_uptD2)
  show "is_sig_GB_in dgrad (set ?bs) u"
  proof (rule is_sig_GB_inI)
    fix r :: "'t 0 'b"
    assume "lt r = u"
    assume "r  dgrad_sig_set dgrad"
    with 1 have "sig_red_zero (≼t) (set bs) r" using lt r = u by (rule is_sig_GB_inD)
    moreover from u have "component_of_term (lt r) < j" by (simp only: lt r = u)
    ultimately show "sig_red_zero (≼t) (set ?bs) r" by (rule sig_red_zero_filter)
  qed
qed

lemma rb_aux_inv2_no_zero_red:
  assumes "hom_grading dgrad" and "is_regular_sequence fs" and "rb_aux_inv2 (bs, ss, p # ps)"
    and "¬ sig_crit bs (new_syz_sigs ss bs p) p"
  shows "rep_list (sig_trd bs (poly_of_pair p))  0"
proof
  from assms(3) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
  moreover have "p  set (p # ps)" by simp
  ultimately have sig_p: "sig_of_pair p = lt (poly_of_pair p)" and "poly_of_pair p  0"
    and p_in: "poly_of_pair p  dgrad_sig_set dgrad"
    by (rule pair_list_sig_of_pair, rule pair_list_nonzero, rule pair_list_dgrad_sig_set)
  from this(2) have "lc (poly_of_pair p)  0" by (rule lc_not_0)
  from inv have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
  hence bs_sub: "set bs  dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)

  define p' where "p' = sig_trd bs (poly_of_pair p)"
  define j where "j = component_of_term (lt p')"
  define q where "q = lookup (vectorize_poly p') j"
  let ?bs = "filter (λb. component_of_term (lt b) < j) bs"
  let ?fs = "take (Suc j) fs"

  have "p'  dgrad_sig_set dgrad" unfolding p'_def using dgrad bs_sub p_in sig_trd_red_rtrancl
    by (rule dgrad_sig_set_closed_sig_red_rtrancl)
  hence "p'  sig_inv_set" by (simp add: dgrad_sig_set'_def)

  have lt_p': "lt p' = lt (poly_of_pair p)" and "lc p' = lc (poly_of_pair p)"
    unfolding p'_def using sig_trd_red_rtrancl
    by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
  from this(2) lc (poly_of_pair p)  0 have "p'  0" by (simp add: lc_eq_zero_iff[symmetric])
  hence "lt p'  keys p'" by (rule lt_in_keys)
  hence "j  keys (vectorize_poly p')" by (simp add: keys_vectorize_poly j_def)
  hence "q  0" by (simp add: q_def in_keys_iff)

  from p'  sig_inv_set lt p'  keys p' have "j < length fs"
    unfolding j_def by (rule sig_inv_setD')
  with le_refl have "fs ! j  set (drop j fs)" by (rule nth_in_set_dropI)
  with fs_distinct le_refl have 0: "fs ! j  set (take j fs)"
    by (auto dest: set_take_disj_set_drop_if_distinct)

  have 1: "j  component_of_term (sig_of_pair p0)" if "p0  set (p # ps)" for p0
  proof -
    from that have "p0 = p  p0  set ps" by simp
    thus ?thesis
    proof
      assume "p0 = p"
      thus ?thesis by (simp add: j_def lt_p' sig_p)
    next
      assume "p0  set ps"
      from inv have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
      hence "Ball (set ps) (pair_ord p)" by simp
      hence "pair_ord p p0" using p0  set ps ..
      hence "lt p' t sig_of_pair p0" by (simp add: pair_ord_def lt_p' sig_p)
      thus ?thesis using pot by (auto simp add: is_pot_ord j_def term_simps)
    qed
  qed
  with assms(1) inv have gb: "punit.is_Groebner_basis (rep_list ` set ?bs)"
    by (rule rb_aux_inv_is_Groebner_basis)

  have "p'  sig_inv_set' (Suc j)"
  proof (rule sig_inv_setI')
    fix v
    assume "v  keys p'"
    hence "v t lt p'" by (rule lt_max_keys)
    with pot have "component_of_term v  j" unfolding j_def by (rule is_pot_ordD2)
    thus "component_of_term v < Suc j" by simp
  qed
  hence 2: "keys (vectorize_poly p')  {0..<Suc j}" by (rule sig_inv_setD)
  moreover assume "rep_list p' = 0"
  ultimately have "0 = (kkeys (pm_of_idx_pm ?fs (vectorize_poly p')).
                              lookup (pm_of_idx_pm ?fs (vectorize_poly p')) k * k)"
    by (simp add: rep_list_def ideal.rep_def pm_of_idx_pm_take)
  also have "... = (kset ?fs. lookup (pm_of_idx_pm ?fs (vectorize_poly p')) k * k)"
    using finite_set keys_pm_of_idx_pm_subset by (rule sum.mono_neutral_left) (simp add: in_keys_iff)

  also from 2 have "... = (kset ?fs. lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
    by (simp only: pm_of_idx_pm_take)
  also have "... = lookup (pm_of_idx_pm fs (vectorize_poly p')) (fs ! j) * fs ! j +
                    (kset (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
    using j < length fs by (simp add: take_Suc_conv_app_nth q_def sum.insert[OF finite_set 0])
  also have "... = q * fs ! j + (kset (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
    using fs_distinct j < length fs by (simp only: lookup_pm_of_idx_pm_distinct q_def)
  finally have "- (q * fs ! j) =
                          (kset (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
    by (simp add: add_eq_0_iff)
  hence "- (q * fs ! j)  ideal (set (take j fs))" by (simp add: ideal.sum_in_spanI)
  hence "- (- (q * fs ! j))  ideal (set (take j fs))" by (rule ideal.span_neg)
  hence "q * fs ! j  ideal (set (take j fs))" by simp
  with assms(2) j < length fs have "q  ideal (set (take j fs))" by (rule is_regular_sequenceD)
  also from assms(3) 1 have "...  ideal (rep_list ` set ?bs)"
    by (rule rb_aux_inv2_ideal_subset)
  finally have "q  ideal (rep_list ` set ?bs)" .
  with gb obtain g where "g  rep_list ` set ?bs" and "g  0" and "punit.lt g adds punit.lt q"
    using q  0 by (rule punit.GB_adds_lt[simplified])
  from this(1) obtain b where "b  set bs" and "component_of_term (lt b) < j" and g: "g = rep_list b"
    by auto
  from assms(3) j < length fs _ this(1, 2)
  have "sset (new_syz_sigs ss bs p). s addst term_of_pair (punit.lt (rep_list b), j)"
  proof (rule rb_aux_inv2_preserved_0)
    show "Inr j  set ps"
    proof
      assume "Inr j  set ps"
      with inv have "sig_of_pair p  term_of_pair (0, j)" by (rule Inr_in_tailD)
      hence "lt p'  term_of_pair (0, j)" by (simp add: lt_p' sig_p)
      from inv have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
      hence "Ball (set ps) (pair_ord p)" by simp
      hence "pair_ord p (Inr j)" using Inr j  set ps ..
      hence "lt p' t term_of_pair (0, j)" by (simp add: pair_ord_def lt_p' sig_p)
      hence "lp p'  0" using pot by (simp add: is_pot_ord j_def term_simps)
      hence "lp p' = 0" using zero_min by (rule ordered_powerprod_lin.order_antisym)
      hence "lt p' = term_of_pair (0, j)" by (metis j_def term_of_pair_pair)
      with lt p'  term_of_pair (0, j) show False ..
    qed
  qed
  then obtain s where s_in: "s  set (new_syz_sigs ss bs p)" and "s addst term_of_pair (punit.lt g, j)"
    unfolding g ..
  from this(2) punit.lt g adds punit.lt q have "s addst term_of_pair (punit.lt q, j)"
    by (metis adds_minus_splus adds_term_splus component_of_term_of_pair pp_of_term_of_pair)
  also have "... = lt p'" by (simp only: q_def j_def lt_lookup_vectorize term_simps)
  finally have "s addst sig_of_pair p" by (simp only: lt_p' sig_p)
  with s_in have pred: "is_pred_syz (new_syz_sigs ss bs p) (sig_of_pair p)"
    by (auto simp: is_pred_syz_def)
  have "sig_crit bs (new_syz_sigs ss bs p) p"
  proof (rule sum_prodE)
    fix x y
    assume "p = Inl (x, y)"
    thus ?thesis using pred by (auto simp: ord_term_lin.max_def split: if_splits)
  next
    fix i
    assume "p = Inr i"
    thus ?thesis using pred by simp
  qed
  with assms(4) show False ..
qed

corollary rb_aux_no_zero_red':
  assumes "hom_grading dgrad" and "is_regular_sequence fs" and "rb_aux_inv2 (fst args)"
  shows "snd (rb_aux args) = snd args"
proof -
  from assms(3) have "rb_aux_inv (fst args)" by (rule rb_aux_inv2_D1)
  hence "rb_aux_dom args" by (rule rb_aux_domI)
  thus ?thesis using assms(3)
  proof (induct args rule: rb_aux.pinduct)
    case (1 bs ss z)
    show ?case by (simp only: rb_aux.psimps(1)[OF 1(1)])
  next
    case (2 bs ss p ps z)
    from 2(5) have *: "rb_aux_inv2 (bs, ss, p # ps)" by (simp only: fst_conv)
    show ?case
    proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
      note refl
      moreover assume "sig_crit bs (new_syz_sigs ss bs p) p"
      moreover from * this have "rb_aux_inv2 (fst ((bs, new_syz_sigs ss bs p, ps), z))"
        unfolding fst_conv by (rule rb_aux_inv2_preserved_1)
      ultimately have "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
                        snd ((bs, new_syz_sigs ss bs p, ps), z)" by (rule 2(2))
      thus "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) = z" by (simp only: snd_conv)
      thus "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) = z" .
    next
      assume "¬ sig_crit bs (new_syz_sigs ss bs p) p"
      with assms(1, 2) * have "rep_list (sig_trd bs (poly_of_pair p))  0"
        by (rule rb_aux_inv2_no_zero_red)
      moreover assume "rep_list (sig_trd bs (poly_of_pair p)) = 0"
      ultimately show "snd (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) #
                          new_syz_sigs ss bs p, ps), Suc z)) = z" ..
    next
      define p' where "p' = sig_trd bs (poly_of_pair p)"
      note refl
      moreover assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p"
      moreover note p'_def
      moreover assume b: "rep_list p'  0"
      moreover have "rb_aux_inv2 (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))"
        using * a b unfolding fst_conv p'_def by (rule rb_aux_inv2_preserved_3)
      ultimately have "snd (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)) =
            snd ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)"
        by (rule 2(4))
      thus "snd (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)) = z"
        by (simp only: snd_conv)
    qed
  qed
qed

corollary rb_aux_no_zero_red:
  assumes "hom_grading dgrad" and "is_regular_sequence fs"
  shows "snd (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)) = z"
proof -
  let ?args = "(([]::('t 0 'b) list, Koszul_syz_sigs fs,
                  (map Inr [0..<length fs])::((('t 0 'b) × ('t 0 'b)) + nat) list), z)"
  from rb_aux_inv_init have "rb_aux_inv2 (fst ?args)" by simp
  with assms have "snd (rb_aux ?args) = snd ?args" by (rule rb_aux_no_zero_red')
  thus ?thesis by (simp only: snd_conv)
qed

corollary rb_no_zero_red:
  assumes "hom_grading dgrad" and "is_regular_sequence fs"
  shows "snd rb = 0"
  using rb_aux_no_zero_red[OF assms, of 0] by (auto simp: rb_def split: prod.split)

end

subsection ‹Sig-Poly-Pairs›

text ‹We now prove that the algorithms defined for sig-poly-pairs (i.\,e. those whose names end with
  _spp›) behave exactly as those defined for module elements. More precisely, if A› is some
  algorithm defined for module elements, we prove something like
  @{prop "spp_of (A x) = A_spp (spp_of x)"}.›

fun spp_inv_pair :: "((('t × ('a 0 'b)) × ('t × ('a 0 'b))) + nat)  bool" where
  "spp_inv_pair (Inl (p, q)) = (spp_inv p  spp_inv q)" |
  "spp_inv_pair (Inr j) = True"

fun app_pair :: "('x  'y)  (('x × 'x) + nat)  (('y × 'y) + nat)" where
  "app_pair f (Inl (p, q)) = Inl (f p, f q)" |
  "app_pair f (Inr j) = Inr j"

fun app_args :: "('x  'y)  (('x list × 'z × ((('x × 'x) + nat) list)) × nat) 
                     (('y list × 'z × ((('y × 'y) + nat) list)) × nat)" where
  "app_args f ((as, bs, cs), n) = ((map f as, bs, map (app_pair f) cs), n)"

lemma app_pair_spp_of_vec_of:
  assumes "spp_inv_pair p"
  shows "app_pair spp_of (app_pair vec_of p) = p"
proof (rule sum_prodE)
  fix a b
  assume p: "p = Inl (a, b)"
  from assms have "spp_inv a" and "spp_inv b" by (simp_all add: p)
  thus ?thesis by (simp add: p spp_of_vec_of)
qed simp

lemma map_app_pair_spp_of_vec_of:
  assumes "list_all spp_inv_pair ps"
  shows "map (app_pair spp_of  app_pair vec_of) ps = ps"
proof (rule map_idI)
  fix p
  assume "p  set ps"
  with assms have "spp_inv_pair p" by (simp add: list_all_def)
  hence "app_pair spp_of (app_pair vec_of p) = p" by (rule app_pair_spp_of_vec_of)
  thus "(app_pair spp_of  app_pair vec_of) p = p" by simp
qed

lemma snd_app_args: "snd (app_args f args) = snd args"
  by (metis prod.exhaust app_args.simps snd_conv)

lemma new_syz_sigs_alt_spp:
  "new_syz_sigs ss bs p = new_syz_sigs_spp ss (map spp_of bs) (app_pair spp_of p)"
proof (rule sum_prodE)
  fix a b
  assume "p = Inl (a, b)"
  thus ?thesis by simp
next
  fix j
  assume "p = Inr j"
  thus ?thesis by (simp add: comp_def spp_of_def)
qed

lemma is_rewritable_alt_spp:
  assumes "0  set bs"
  shows "is_rewritable bs p u = is_rewritable_spp (map spp_of bs) (spp_of p) u"
proof -
  from assms have "b  set bs  b  0" for b by blast
  thus ?thesis by (auto simp: is_rewritable_def is_rewritable_spp_def fst_spp_of)
qed

lemma spair_sigs_alt_spp: "spair_sigs p q = spair_sigs_spp (spp_of p) (spp_of q)"
  by (simp add: spair_sigs_def spair_sigs_spp_def Let_def fst_spp_of snd_spp_of)

lemma sig_crit_alt_spp:
  assumes "0  set bs"
  shows "sig_crit bs ss p = sig_crit_spp (map spp_of bs) ss (app_pair spp_of p)"
proof (rule sum_prodE)
  fix a b
  assume p: "p = Inl (a, b)"
  from assms show ?thesis by (simp add: p spair_sigs_alt_spp is_rewritable_alt_spp)
qed simp

lemma spair_alt_spp:
  assumes "is_regular_spair p q"
  shows "spp_of (spair p q) = spair_spp (spp_of p) (spp_of q)"
proof -
  let ?t1 = "punit.lt (rep_list p)"
  let ?t2 = "punit.lt (rep_list q)"
  let ?l = "lcs ?t1 ?t2"
  from assms have p: "rep_list p  0" and q: "rep_list q  0"
    by (rule is_regular_spairD1, rule is_regular_spairD2)
  hence "p  0" and "q  0" and 1: "punit.lc (rep_list p)  0" and 2: "punit.lc (rep_list q)  0"
    by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
  from assms have "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p) 
                    lt (monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q)" (is "?u  ?v")
    by (rule is_regular_spairD3)
  hence "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p - monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q) =
          ord_term_lin.max ?u ?v" by (rule lt_minus_distinct_eq_max)
  moreover from p  0 1 have "?u = (?l - ?t1)  fst (spp_of p)" by (simp add: lt_monom_mult fst_spp_of)
  moreover from q  0 2 have "?v = (?l - ?t2)  fst (spp_of q)" by (simp add: lt_monom_mult fst_spp_of)
  ultimately show ?thesis
    by (simp add: spair_spp_def spair_def Let_def spp_of_def rep_list_minus rep_list_monom_mult)
qed

lemma sig_trd_spp_body_alt_Some:
  assumes "find_sig_reducer (map spp_of bs) v (punit.lt p) 0 = Some i"
  shows "sig_trd_spp_body (map spp_of bs) v (p, r) =
          (punit.lower (p - local.punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
                  (punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i))) (punit.lt p), r)"
      (is ?thesis1)
    and "sig_trd_spp_body (map spp_of bs) v (p, r) =
          (p - local.punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
                  (punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)), r)"
      (is ?thesis2)
proof -
  have "?thesis1  ?thesis2"
  proof (cases "p = 0")
    case True
    show ?thesis by (simp add: assms, simp add: True)
  next
    case False
    from assms have "i < length bs" by (rule find_sig_reducer_SomeD)
    hence eq1: "snd (map spp_of bs ! i) = rep_list (bs ! i)" by (simp add: snd_spp_of)
    from assms have "rep_list (bs ! i)  0" and "punit.lt (rep_list (bs ! i)) adds punit.lt p"
      by (rule find_sig_reducer_SomeD)+
    hence nz: "rep_list (bs ! i)  0" and adds: "punit.lt (rep_list (bs ! i)) adds punit.lt p"
      by (simp_all add: snd_spp_of)
    from nz have "punit.lc (rep_list (bs ! i))  0" by (rule punit.lc_not_0)
    moreover from False have "punit.lc p  0" by (rule punit.lc_not_0)
    ultimately have eq2: "punit.lt (punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
                        (punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i))) = punit.lt p"
      (is "punit.lt ?p = _") using nz adds by (simp add: lp_monom_mult adds_minus)
    have ?thesis1 by (simp add: assms Let_def eq1 punit.lower_minus punit.tail_monom_mult[symmetric],
                      simp add: punit.tail_def eq2)
    moreover have ?thesis2
    proof (simp add: ?thesis1 punit.lower_id_iff disj_commute[of "p = ?p"] del: sig_trd_spp_body.simps)
      show "punit.lt (p - ?p)  punit.lt p  p = ?p"
      proof (rule disjCI)
        assume "p  ?p"
        hence "p - ?p  0" by simp
        moreover note eq2
        moreover from punit.lc (rep_list (bs ! i))  0 have "punit.lc ?p = punit.lc p" by simp
        ultimately show "punit.lt (p - ?p)  punit.lt p" by (rule punit.lt_minus_lessI)
      qed
    qed
    ultimately show ?thesis ..
  qed
  thus ?thesis1 and ?thesis2 by blast+
qed

lemma sig_trd_aux_alt_spp:
  assumes "fst args  keys (rep_list (snd args))"
  shows "rep_list (sig_trd_aux bs args) =
              sig_trd_spp_aux (map spp_of bs) (lt (snd args))
                (rep_list (snd args) - punit.higher (rep_list (snd args)) (fst args),
                 punit.higher (rep_list (snd args)) (fst args))"
proof -
  from assms have "sig_trd_aux_dom bs args" by (rule sig_trd_aux_domI)
  thus ?thesis using assms
  proof (induct args rule: sig_trd_aux.pinduct)
    case (1 t p)
    define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
                              None  p
                            | Some i  p -
                                monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                                 (t - punit.lt (rep_list (bs ! i))) (bs ! i))"
    define p'' where "p'' = punit.lower (rep_list p') t"
    from 1(3) have t_in: "t  keys (rep_list p)" by simp
    hence "t  keys (rep_list p - punit.higher (rep_list p) t)" (is "_  keys ?p")
      by (simp add: punit.keys_minus_higher)
    hence "?p  0" by auto
    hence eq1: "sig_trd_spp_aux bs0 v0 (?p, r0) = sig_trd_spp_aux bs0 v0 (sig_trd_spp_body bs0 v0 (?p, r0))"
      for bs0 v0 r0 by (simp add: sig_trd_spp_aux_simps del: sig_trd_spp_body.simps)
    from t_in have lt_p: "punit.lt ?p = t" and lc_p: "punit.lc ?p = lookup (rep_list p) t"
      and tail_p: "punit.tail ?p = punit.lower (rep_list p) t"
      by (rule punit.lt_minus_higher, rule punit.lc_minus_higher, rule punit.tail_minus_higher)
    have "lt p' = lt p  punit.higher (rep_list p') t = punit.higher (rep_list p) t 
          (i. find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i  lookup (rep_list p') t = 0)"
      (is "?A  ?B  ?C")
    proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
      case None
      thus ?thesis by (simp add: p'_def)
    next
      case (Some i)
      hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                            (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
      from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
      hence eq: "t - punit.lt (rep_list (bs ! i)) + punit.lt (rep_list (bs ! i)) = t" by (rule adds_minus)
      from t_in Some have *: "sig_red_single (≺t) (≼) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
        unfolding p' by (rule find_sig_reducer_SomeD_red_single)
      hence **: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
        by (rule sig_red_single_red_single)
      from * have ?A by (rule sig_red_single_regular_lt)
      moreover from punit.red_single_higher[OF **] have ?B by (simp add: eq)
      moreover have ?C
      proof (intro allI impI)
        from punit.red_single_lookup[OF **] show "lookup (rep_list p') t = 0" by (simp add: eq)
      qed
      ultimately show ?thesis by (intro conjI)
    qed
    hence lt_p': "lt p' = lt p" and higher_p': "punit.higher (rep_list p') t = punit.higher (rep_list p) t"
      and lookup_p': "i. find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i  lookup (rep_list p') t = 0"
      by blast+
    show ?case
    proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric], intro conjI impI)
      assume "p'' = 0"
      hence p'_decomp: "punit.higher (rep_list p) t + monomial (lookup (rep_list p') t) t = rep_list p'"
        using punit.higher_lower_decomp[of "rep_list p'" t] by (simp add: p''_def higher_p')
      show "rep_list p' = sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)"
      proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
        case None
        hence p': "p' = p" by (simp add: p'_def)
        from p'' = 0 have eq2: "punit.tail ?p = 0" by (simp add: tail_p p''_def p')
        from p'_decomp show ?thesis by (simp add: p' eq1 lt_p lc_p None eq2 sig_trd_spp_aux_simps)
      next
        case (Some i)
        hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                            (t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
        from p'' = 0 have eq2: "punit.lower (rep_list p - punit.higher (rep_list p) t -
                punit.monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                  (t - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)))
                t = 0"
          by (simp add: p''_def p' rep_list_minus rep_list_monom_mult punit.lower_minus punit.lower_higher_zeroI)
        from Some have "lookup (rep_list p') t = 0" by (rule lookup_p')
        with p'_decomp have eq3: "rep_list p' = punit.higher (rep_list p) t" by simp
        show ?thesis by (simp add: sig_trd_spp_body_alt_Some(1) eq1 eq2 lt_p lc_p Some del: sig_trd_spp_body.simps,
                         simp add: sig_trd_spp_aux_simps eq3)
      qed
    next
      assume "p''  0"
      hence "punit.lt p''  t" unfolding p''_def by (rule punit.lt_lower_less)
      have higher_p'_2: "punit.higher (rep_list p') (punit.lt p'') =
                          punit.higher (rep_list p) t + monomial (lookup (rep_list p') t) t"
      proof (simp add: higher_p'[symmetric], rule poly_mapping_eqI)
        fix s
        show "lookup (punit.higher (rep_list p') (punit.lt p'')) s =
              lookup (punit.higher (rep_list p') t + monomial (lookup (rep_list p') t) t) s"
        proof (rule ordered_powerprod_lin.linorder_cases)
          assume "t  s"
          moreover from punit.lt p''  t this have "punit.lt p''  s"
            by (rule ordered_powerprod_lin.less_trans)
          ultimately show ?thesis by (simp add: lookup_add punit.lookup_higher_when lookup_single)
        next
          assume "t = s"
          with punit.lt p''  t show ?thesis by (simp add: lookup_add punit.lookup_higher_when)
        next
          assume "s  t"
          show ?thesis
          proof (cases "punit.lt p''  s")
            case True
            hence "lookup (punit.higher (rep_list p') (punit.lt p'')) s = lookup (rep_list p') s"
              by (simp add: punit.lookup_higher_when)
            also from s  t have "... = lookup p'' s" by (simp add: p''_def punit.lookup_lower_when)
            also from True have "... = 0" using punit.lt_le_iff by auto
            finally show ?thesis using s  t
              by (simp add: lookup_add lookup_single punit.lookup_higher_when)
          next
            case False
            with s  t show ?thesis by (simp add: lookup_add punit.lookup_higher_when lookup_single)
          qed
        qed
      qed
      have "rep_list (sig_trd_aux bs (punit.lt p'', p')) =
              sig_trd_spp_aux (map spp_of bs) (lt (snd (punit.lt p'', p')))
               (rep_list (snd (punit.lt p'', p')) -
                punit.higher (rep_list (snd (punit.lt p'', p'))) (fst (punit.lt p'', p')),
                punit.higher (rep_list (snd (punit.lt p'', p'))) (fst (punit.lt p'', p')))"
        using p'_def p''_def p''  0
      proof (rule 1(2))
        from p''  0 have "punit.lt p''  keys p''" by (rule punit.lt_in_keys)
        also have "...  keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
        finally show "fst (punit.lt p'', p')  keys (rep_list (snd (punit.lt p'', p')))" by simp
      qed
      also have "... = sig_trd_spp_aux (map spp_of bs) (lt p)
                         (rep_list p' - punit.higher (rep_list p') (punit.lt p''),
                          punit.higher (rep_list p') (punit.lt p''))"
        by (simp only: lt_p' fst_conv snd_conv)
      also have "... = sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)"
      proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
        case None
        hence p': "p' = p" by (simp add: p'_def)
        have "rep_list p - (punit.higher (rep_list p) t + monomial (lookup (rep_list p) t) t) =
              punit.lower (rep_list p) t"
          using punit.higher_lower_decomp[of "rep_list p" t] by (simp add: diff_eq_eq ac_simps)
        with higher_p'_2 show ?thesis by (simp add: eq1 lt_p lc_p tail_p p' None)
      next
        case (Some i)
        hence p': "rep_list p - punit.monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
                          (t - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)) = rep_list p'"
          by (simp add: p'_def rep_list_minus rep_list_monom_mult)
        from Some have "lookup (rep_list p') t = 0" by (rule lookup_p')
        with higher_p'_2 show ?thesis
          by (simp add: sig_trd_spp_body_alt_Some(2) eq1 lt_p lc_p tail_p Some
              diff_right_commute[of "rep_list p" "punit.higher (rep_list p) t"] p' del: sig_trd_spp_body.simps)
      qed
      finally show "rep_list (sig_trd_aux bs (punit.lt p'', p')) =
                     sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)" .
    qed
  qed
qed

lemma sig_trd_alt_spp: "spp_of (sig_trd bs p) = sig_trd_spp (map spp_of bs) (spp_of p)"
  unfolding sig_trd_def
proof (split if_split, intro conjI impI)
  assume "rep_list p = 0"
  thus "spp_of p = sig_trd_spp (map spp_of bs) (spp_of p)" by (simp add: spp_of_def sig_trd_spp_aux_simps)
next
  let ?args = "(punit.lt (rep_list p), p)"
  assume "rep_list p  0"
  hence a: "fst ?args  keys (rep_list (snd ?args))" by (simp add: punit.lt_in_keys)
  hence "(sig_red (≺t) (≼) (set bs))** (snd ?args) (sig_trd_aux bs ?args)"
    by (rule sig_trd_aux_red_rtrancl)
  hence eq1: "lt (sig_trd_aux bs ?args) = lt (snd ?args)" by (rule sig_red_regular_rtrancl_lt)
  have eq2: "punit.higher (rep_list p) (punit.lt (rep_list p)) = 0"
    by (auto simp: punit.higher_eq_zero_iff punit.lt_max simp flip: not_in_keys_iff_lookup_eq_zero
            dest: punit.lt_max_keys)
  show "spp_of (sig_trd_aux bs (punit.lt (rep_list p), p)) = sig_trd_spp (map spp_of bs) (spp_of p)"
    by (simp add: spp_of_def eq1 eq2 sig_trd_aux_alt_spp[OF a])
qed

lemma is_regular_spair_alt_spp: "is_regular_spair p q  is_regular_spair_spp (spp_of p) (spp_of q)"
  by (auto simp: is_regular_spair_spp_def fst_spp_of snd_spp_of intro: is_regular_spairI
          dest: is_regular_spairD1 is_regular_spairD2 is_regular_spairD3)

lemma sig_of_spair_alt_spp: "sig_of_pair p = sig_of_pair_spp (app_pair spp_of p)"
proof (rule sum_prodE)
  fix a b
  assume p: "p = Inl (a, b)"
  show ?thesis by (simp add: p spair_sigs_def spair_sigs_spp_def spp_of_def)
qed simp

lemma pair_ord_alt_spp: "pair_ord x y  pair_ord_spp (app_pair spp_of x) (app_pair spp_of y)"
  by (simp add: pair_ord_spp_def pair_ord_def sig_of_spair_alt_spp)

lemma new_spairs_alt_spp:
  "map (app_pair spp_of) (new_spairs bs p) = new_spairs_spp (map spp_of bs) (spp_of p)"
proof (induct bs)
  case Nil
  show ?case by simp
next
  case (Cons b bs)
  have "map (app_pair spp_of) (insort_wrt pair_ord (Inl (p, b)) (new_spairs bs p)) =
        insort_wrt pair_ord_spp (app_pair spp_of (Inl (p, b))) (map (app_pair spp_of) (new_spairs bs p))"
    by (rule map_insort_wrt, rule pair_ord_alt_spp[symmetric])
  thus ?case by (simp add: is_regular_spair_alt_spp Cons)
qed

lemma add_spairs_alt_spp:
  assumes "x. x  set bs  Inl (spp_of p, spp_of x)  app_pair spp_of ` set ps"
  shows "map (app_pair spp_of) (add_spairs ps bs p) =
          add_spairs_spp (map (app_pair spp_of) ps) (map spp_of bs) (spp_of p)"
proof -
  have "map (app_pair spp_of) (merge_wrt pair_ord (new_spairs bs p) ps) =
        merge_wrt pair_ord_spp (map (app_pair spp_of) (new_spairs bs p)) (map (app_pair spp_of) ps)"
  proof (rule map_merge_wrt, rule ccontr)
    assume "app_pair spp_of ` set (new_spairs bs p)  app_pair spp_of ` set ps  {}"
    then obtain q' where "q'  app_pair spp_of ` set (new_spairs bs p)"
      and q'_in: "q'  app_pair spp_of ` set ps" by blast
    from this(1) obtain q where "q  set (new_spairs bs p)" and q': "q' = app_pair spp_of q" ..
    from this(1) obtain x where x_in: "x  set bs" and q: "q = Inl (p, x)"
      by (rule in_new_spairsE)
    have q': "q' = Inl (spp_of p, spp_of x)" by (simp add: q q')
    have "q'  app_pair spp_of ` set ps" unfolding q' using x_in by (rule assms)
    thus False using q'_in ..
  qed (simp only: pair_ord_alt_spp)
  thus ?thesis by (simp add: add_spairs_def add_spairs_spp_def new_spairs_alt_spp)
qed

lemma rb_aux_invD_app_args:
  assumes "rb_aux_inv (fst (app_args vec_of ((bs, ss, ps), z)))"
  shows "list_all spp_inv bs" and "list_all spp_inv_pair ps"
proof -
  from assms(1) have inv: "rb_aux_inv (map vec_of bs, ss, map (app_pair vec_of) ps)" by simp
  hence "rb_aux_inv1 (map vec_of bs)" by (rule rb_aux_inv_D1)
  hence "0  rep_list ` set (map vec_of bs)" by (rule rb_aux_inv1_D2)
  hence "0  vec_of ` set bs" using rep_list_zero by fastforce
  hence 1: "b  set bs  spp_inv b" for b by (auto simp: spp_inv_alt)
  thus "list_all spp_inv bs" by (simp add: list_all_def)

  have 2: "x  set bs" if "vec_of x  set (map vec_of bs)" for x
  proof -
    from that have "vec_of x  vec_of ` set bs" by simp
    then obtain y where "y  set bs" and eq: "vec_of x = vec_of y" ..
    from this(1) have "spp_inv y" by (rule 1)
    moreover have "vec_of y = vec_of x" by (simp only: eq)
    ultimately have "y = x" by (rule vec_of_inj)
    with y  set bs show ?thesis by simp
  qed

  show "list_all spp_inv_pair ps" unfolding list_all_def
  proof (rule ballI)
    fix p
    assume "p  set ps"
    show "spp_inv_pair p"
    proof (rule sum_prodE)
      fix a b
      assume p: "p = Inl (a, b)"
      from p  set ps have "Inl (a, b)  set ps" by (simp only: p)
      hence "app_pair vec_of (Inl (a, b))  app_pair vec_of ` set ps" by (rule imageI)
      hence "Inl (vec_of a, vec_of b)  set (map (app_pair vec_of) ps)" by simp
      with inv have "vec_of a  set (map vec_of bs)" and "vec_of b  set (map vec_of bs)"
        by (rule rb_aux_inv_D3)+
      have "spp_inv a" by (rule 1, rule 2, fact)
      moreover have "spp_inv b" by (rule 1, rule 2, fact)
      ultimately show ?thesis by (simp add: p)
    qed simp
  qed
qed

lemma app_args_spp_of_vec_of:
  assumes "rb_aux_inv (fst (app_args vec_of args))"
  shows "app_args spp_of (app_args vec_of args) = args"
proof -
  obtain bs ss ps z where args: "args = ((bs, ss, ps), z)" using prod.exhaust by metis
  from assms have "list_all spp_inv bs" and *: "list_all spp_inv_pair ps" unfolding args
    by (rule rb_aux_invD_app_args)+
  from this(1) have "map (spp_of  vec_of) bs = bs" by (rule map_spp_of_vec_of)
  moreover from * have "map (app_pair spp_of  app_pair vec_of) ps = ps"
    by (rule map_app_pair_spp_of_vec_of)
  ultimately show ?thesis by (simp add: args)
qed

lemma poly_of_pair_alt_spp:
  assumes "distinct fs" and "rb_aux_inv (bs, ss, p # ps)"
  shows "spp_of (poly_of_pair p) = spp_of_pair (app_pair spp_of p)"
proof -
  show ?thesis
  proof (rule sum_prodE)
    fix a b
    assume p: "p = Inl (a, b)"
    hence "Inl (a, b)  set (p # ps)" by simp
    with assms(2) have "is_regular_spair a b" by (rule rb_aux_inv_D3)
    thus ?thesis by (simp add: p spair_alt_spp)
  next
    fix j
    assume p: "p = Inr j"
    hence "Inr j  set (p # ps)" by simp
    with assms(2) have "j < length fs" by (rule rb_aux_inv_D4)
    thus ?thesis by (simp add: p spp_of_def lt_monomial rep_list_monomial[OF assms(1)] term_simps)
  qed
qed

lemma rb_aux_alt_spp:
  assumes "rb_aux_inv (fst args)"
  shows "app_args spp_of (rb_aux args) = rb_spp_aux (app_args spp_of args)"
proof -
  from assms have "rb_aux_dom args" by (rule rb_aux_domI)
  thus ?thesis using assms
  proof (induct args rule: rb_aux.pinduct)
    case (1 bs ss z)
    show ?case by (simp add: rb_aux.psimps(1)[OF 1(1)] rb_spp_aux_Nil)
  next
    case (2 bs ss p ps z)
    let ?q = "sig_trd bs (poly_of_pair p)"

    from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
    hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
    hence "0  rep_list ` set bs" by (rule rb_aux_inv1_D2)
    hence "0  set bs" by (force simp: rep_list_zero)
    hence eq1: "sig_crit_spp (map spp_of bs) ss' (app_pair spp_of p)  sig_crit bs ss' p" for ss'
      by (simp add: sig_crit_alt_spp)
    from fs_distinct * have eq2: "sig_trd_spp (map spp_of bs) (spp_of_pair (app_pair spp_of p)) = spp_of ?q"
      by (simp only: sig_trd_alt_spp poly_of_pair_alt_spp)

    show ?case
    proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
      note refl
      moreover assume a: "sig_crit bs (new_syz_sigs ss bs p) p"
      moreover from * this have "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))"
        unfolding fst_conv by (rule rb_aux_inv_preserved_1)
      ultimately have "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
                        rb_spp_aux (app_args spp_of ((bs, new_syz_sigs ss bs p, ps), z))"
        by (rule 2(2))
      also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
        by (simp add: rb_spp_aux_Cons eq1 a new_syz_sigs_alt_spp[symmetric])
      finally show "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
            rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
      thus "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
            rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
    next
      assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list ?q = 0"
      from * b have "rb_aux_inv (fst ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z))"
        unfolding fst_conv by (rule rb_aux_inv_preserved_2)
      with refl a refl b have "app_args spp_of (rb_aux ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z)) =
                               rb_spp_aux (app_args spp_of ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z))"
        by (rule 2(3))
      also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
        by (simp add: rb_spp_aux_Cons eq1 a Let_def eq2 snd_spp_of b fst_spp_of new_syz_sigs_alt_spp[symmetric])
      finally show "app_args spp_of (rb_aux ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z)) =
                    rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
    next
      assume a: "¬ sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list ?q  0"

      have "Inl (spp_of ?q, spp_of x)  app_pair spp_of ` set ps" for x
      proof
        assume "Inl (spp_of ?q, spp_of x)  app_pair spp_of ` set ps"
        then obtain y where "y  set ps" and eq0: "Inl (spp_of ?q, spp_of x) = app_pair spp_of y" ..
        obtain a b where y: "y = Inl (a, b)" and "spp_of ?q = spp_of a"
        proof (rule sum_prodE)
          fix a b
          assume "y = Inl (a, b)"
          moreover from eq0 have "spp_of ?q = spp_of a" by (simp add: y = Inl (a, b))
          ultimately show ?thesis ..
        next
          fix j
          assume "y = Inr j"
          with eq0 show ?thesis by simp
        qed
        from this(2) have "lt ?q = lt a" by (simp add: spp_of_def)
        from y  set ps have "y  set (p # ps)" by simp
        with * have "a  set bs" unfolding y by (rule rb_aux_inv_D3(1))
        hence "lt ?q  lt ` set bs" unfolding lt ?q = lt a by (rule imageI)
        moreover from * a b have "lt ?q  lt ` set bs" by (rule rb_aux_inv_preserved_3)
        ultimately show False by simp
      qed
      hence eq3: "add_spairs_spp (map (app_pair spp_of) ps) (map spp_of bs) (spp_of ?q) =
                  map (app_pair spp_of) (add_spairs ps bs ?q)" by (simp add: add_spairs_alt_spp)

      from * a b have "rb_aux_inv (fst ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z))"
        unfolding fst_conv by (rule rb_aux_inv_preserved_3)
      with refl a refl b
      have "app_args spp_of (rb_aux ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z)) =
            rb_spp_aux (app_args spp_of ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z))"
        by (rule 2(4))
      also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
        by (simp add: rb_spp_aux_Cons eq1 a Let_def eq2 fst_spp_of snd_spp_of b eq3 new_syz_sigs_alt_spp[symmetric])
      finally show "app_args spp_of (rb_aux ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z)) =
                     rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
    qed
  qed
qed

corollary rb_spp_aux_alt:
  "rb_aux_inv (fst (app_args vec_of args)) 
    rb_spp_aux args = app_args spp_of (rb_aux (app_args vec_of args))"
  by (simp only: rb_aux_alt_spp app_args_spp_of_vec_of)

corollary rb_spp_aux:
  "hom_grading dgrad 
    punit.is_Groebner_basis (set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))))"
          (is "_  ?thesis1")
  "ideal (set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))) = ideal (set fs)"
          (is "?thesis2")
  "set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))  punit_dgrad_max_set dgrad"
          (is "?thesis3")
  "0  set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
          (is "?thesis4")
  "hom_grading dgrad  is_pot_ord  is_regular_sequence fs 
    snd (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)) = z"
          (is "_  _  _  ?thesis5")
  "rword_strict = rw_rat_strict  p  set (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))) 
    q  set (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))  p  q 
    punit.lt (snd p) adds punit.lt (snd q)  punit.lt (snd p)  fst q t punit.lt (snd q)  fst p"
proof -
  let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
  have eq0: "app_pair vec_of  Inr = Inr" by (intro ext, simp)
  have eq1: "fst (fst (app_args spp_of a)) = map spp_of (fst (fst a))" for a::"(_ × ('t list) × _) × _"
  proof -
    obtain bs ss ps z where "a = ((bs, ss, ps), z)" using prod.exhaust by metis
    thus ?thesis by simp
  qed
  have eq2: "snd  spp_of = rep_list" by (intro ext, simp add: snd_spp_of)
  have "rb_aux_inv (fst (app_args vec_of ?args))" by (simp add: eq0 rb_aux_inv_init)
  hence eq3: "rb_spp_aux ?args = app_args spp_of (rb_aux (app_args vec_of ?args))"
    by (rule rb_spp_aux_alt)

  {
    assume "hom_grading dgrad"
    with rb_aux_is_Groebner_basis show ?thesis1 by (simp add: eq0 eq1 eq2 eq3 del: set_map)
  }

  from ideal_rb_aux show ?thesis2 by (simp add: eq0 eq1 eq2 eq3 del: set_map)

  from dgrad_max_set_closed_rb_aux show ?thesis3 by (simp add: eq0 eq1 eq2 eq3 del: set_map)

  from rb_aux_nonzero show ?thesis4 by (simp add: eq0 eq1 eq2 eq3 del: set_map)

  {
    assume "is_pot_ord" and "hom_grading dgrad" and "is_regular_sequence fs"
    hence "snd (rb_aux ?args) = z" by (rule rb_aux_no_zero_red)
    thus ?thesis5 by (simp add: snd_app_args eq0 eq3)
  }

  {
    from rb_aux_nonzero have "0  rep_list ` set (fst (fst (rb_aux ?args)))"
      (is "0  rep_list ` ?G") by simp
    assume "rword_strict = rw_rat_strict"
    hence "is_min_sig_GB dgrad ?G" by (rule rb_aux_is_min_sig_GB)
    hence rl: "g. g  ?G  ¬ is_sig_red (≼t) (=) (?G - {g}) g" by (simp add: is_min_sig_GB_def)
    assume "p  set (fst (fst (rb_spp_aux ?args)))"
    also have "... = spp_of ` ?G" by (simp add: eq0 eq1 eq3)
    finally obtain p' where "p'  ?G" and p: "p = spp_of p'" ..
    assume "q  set (fst (fst (rb_spp_aux ?args)))"
    also have "... = spp_of ` ?G" by (simp add: eq0 eq1 eq3)
    finally obtain q' where "q'  ?G" and q: "q = spp_of q'" ..
    from this(1) have 1: "¬ is_sig_red (≼t) (=) (?G - {q'}) q'" by (rule rl)
    assume "p  q" and "punit.lt (snd p) adds punit.lt (snd q)"
    hence "p'  q'" and adds: "punit.lt (rep_list p') adds punit.lt (rep_list q')"
      by (auto simp: p q snd_spp_of)
    show "punit.lt (snd p)  fst q t punit.lt (snd q)  fst p"
    proof (rule ccontr)
      assume "¬ punit.lt (snd p)  fst q t punit.lt (snd q)  fst p"
      hence le: "punit.lt (rep_list q')  lt p' t punit.lt (rep_list p')  lt q'"
        by (simp add: p q spp_of_def)
      from p'  q' p'  ?G have "p'  ?G - {q'}" by simp
      moreover from p'  ?G 0  rep_list ` ?G have "rep_list p'  0" by fastforce
      moreover from q'  ?G 0  rep_list ` ?G have "rep_list q'  0" by fastforce
      moreover note adds
      moreover have "ord_term_lin.is_le_rel (≼t)" by simp
      ultimately have "is_sig_red (≼t) (=) (?G - {q'}) q'" using le by (rule is_sig_red_top_addsI)
      with 1 show False ..
    qed
  }
qed

end

end

end

end

end

definition gb_sig_z ::
    "(('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  ('a 0 'b) list  (('t × ('a 0 'b::field)) list × nat)"
  where "gb_sig_z rword_strict fs0 =
              (let fs = rev (remdups (rev (removeAll 0 fs0)));
                   res = rb_spp_aux fs rword_strict (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0) in
                  (fst (fst res), snd res))"

text ‹The second return value of @{const gb_sig_z} is the total number of zero-reductions.›

definition gb_sig :: "(('t × ('a 0 'b))  ('t × ('a 0 'b))  bool)  ('a 0 'b) list  ('a 0 'b::field) list"
  where "gb_sig rword_strict fs0 = map snd (fst (gb_sig_z rword_strict fs0))"

theorem
  assumes "fs. is_strict_rewrite_ord fs rword_strict"
  shows gb_sig_isGB: "punit.is_Groebner_basis (set (gb_sig rword_strict fs))" (is ?thesis1)
    and gb_sig_ideal: "ideal (set (gb_sig rword_strict fs)) = ideal (set fs)" (is ?thesis2)
    and dgrad_p_set_closed_gb_sig:
        "dickson_grading d  set fs  punit.dgrad_p_set d m  set (gb_sig rword_strict fs)  punit.dgrad_p_set d m"
          (is "_  _  ?thesis3")
    and gb_sig_nonzero: "0  set (gb_sig rword_strict fs)" (is ?thesis4)
    and gb_sig_no_zero_red: "is_pot_ord  is_regular_sequence fs  snd (gb_sig_z rword_strict fs) = 0"
proof -
  from ex_hgrad obtain d0::"'a  nat" where "dickson_grading d0  hom_grading d0" ..
  hence dg: "dickson_grading d0" and hg: "hom_grading d0" by simp_all
  define fs1 where "fs1 = rev (remdups (rev (removeAll 0 fs)))"
  note assms dg
  moreover have "distinct fs1" and "0  set fs1" by (simp_all add: fs1_def)
  ultimately have "ideal (set (gb_sig rword_strict fs)) = ideal (set fs1)" and ?thesis4
    unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)+
  thus ?thesis2 and ?thesis4 by (simp_all add: fs1_def ideal.span_Diff_zero)

  from assms dg distinct fs1 0  set fs1 hg show ?thesis1
    unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)

  {
    assume dg: "dickson_grading d" and *: "set fs  punit.dgrad_p_set d m"
    show ?thesis3
    proof (cases "set fs  {0}")
      case True
      hence "removeAll 0 fs = []"
        by (metis (no_types, lifting) Diff_iff ex_in_conv set_empty2 set_removeAll subset_singleton_iff)
      thus ?thesis by (simp add: gb_sig_def gb_sig_z_def Let_def rb_spp_aux_Nil)
    next
      case False
      have "set fs1  set fs" by (fastforce simp: fs1_def)
      hence "Keys (set fs1)  Keys (set fs)" by (rule Keys_mono)
      hence "d ` Keys (set fs1)  d ` Keys (set fs)" by (rule image_mono)
      hence "insert (d 0) (d ` Keys (set fs1))  insert (d 0) (d ` Keys (set fs))" by (rule insert_mono)
      moreover have "insert (d 0) (d ` Keys (set fs1))  {}" by simp
      moreover have "finite (insert (d 0) (d ` Keys (set fs)))"
        by (simp add: finite_Keys)
      ultimately have le: "Max (insert (d 0) (d ` Keys (set fs1))) 
                            Max (insert (d 0) (d ` Keys (set fs)))" by (rule Max_mono)
      from assms dg have "set (gb_sig rword_strict fs)  punit_dgrad_max_set (TYPE('b)) fs1 d"
        using distinct fs1 0  set fs1
        unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)
      also have "punit_dgrad_max_set (TYPE('b)) fs1 d  punit_dgrad_max_set (TYPE('b)) fs d"
        by (rule punit.dgrad_p_set_subset, simp add: dgrad_max_def le)
      also from dg * False have "...  punit.dgrad_p_set d m"
        by (rule punit_dgrad_max_set_subset_dgrad_p_set)
      finally show ?thesis .
    qed
  }

  {
    assume "is_regular_sequence fs"
    note assms dg distinct fs1 0  set fs1 hg
    moreover assume "is_pot_ord"
    moreover from is_regular_sequence fs have "is_regular_sequence fs1" unfolding fs1_def
      by (intro is_regular_sequence_remdups is_regular_sequence_removeAll_zero)
    ultimately show "snd (gb_sig_z rword_strict fs) = 0"
      unfolding gb_sig_def gb_sig_z_def snd_conv fs1_def Let_def by (rule rb_spp_aux)
  }
qed

theorem gb_sig_z_is_min_sig_GB:
  assumes "p  set (fst (gb_sig_z rw_rat_strict fs))" and "q  set (fst (gb_sig_z rw_rat_strict fs))"
    and "p  q" and "punit.lt (snd p) adds punit.lt (snd q)"
  shows "punit.lt (snd p)  fst q t punit.lt (snd q)  fst p"
proof -
  define fs1 where "fs1 = rev (remdups (rev (removeAll 0 fs)))"
  from ex_hgrad obtain d0::"'a  nat" where "dickson_grading d0  hom_grading d0" ..
  hence "dickson_grading d0" ..
  note rw_rat_strict_is_strict_rewrite_ord this
  moreover have "distinct fs1" and "0  set fs1" by (simp_all add: fs1_def)
  moreover note refl assms
  ultimately show ?thesis unfolding gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)
qed

text ‹Summarizing, these are the four main results proved in this theory:
   @{thm gb_sig_isGB},
   @{thm gb_sig_ideal},
   @{thm gb_sig_no_zero_red}, and
   @{thm gb_sig_z_is_min_sig_GB}.›

end (* qpm_inf_term *)

end (* theory *)