Theory Cring_Multivariable_Poly

theory Cring_Multivariable_Poly
imports  "HOL-Algebra.Indexed_Polynomials" Padic_Ints.Cring_Poly
begin

(**************************************************************************************************)
(**************************************************************************************************)
section‹Multivariable Polynomials Over a Commutative Ring›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  This theory extends the content of \texttt{HOL-Algebra.Indexed\_Polynomials}, mainly in the
  context of a commutative base ring. In particular, the ring of polynomials over an arbitrary
  variable set is explicitly witnessed as a ring itself, which is commutative if the base is
  commutative, and a domain if the base ring is a domain. A universal property for polynomial
  evaluation is proved, which allows us to embed polynomial rings in a ring of functions over the
  base ring, as well as construe multivariable polynomials as univariate polynomials over a small
  base polynomial ring.
›

type_synonym 'a monomial = "'a multiset"
type_synonym ('b,'a) mvar_poly = "'a multiset  'b"
type_synonym ('a,'b) ring_hom = "'a  'b"
type_synonym 'a u_poly = "nat  'a"
(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Lemmas about multisets›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  Since multisets function as monomials in this formalization, we'll need some simple lemmas
  about multisets in order to define degree functions and certain lemmas about factorizations.
  Those are provided in this section.
›

lemma count_size:
  assumes "size m  K"
  shows "count m i  K"
  using assms
  by (metis count_le_replicate_mset_subset_eq dual_order.trans order_refl size_mset_mono size_replicate_mset)

text‹The following defines the set of monomials with nonzero coefficients for a given polynomial:›

definition monomials_of :: "('a,'c) ring_scheme  ('a, 'b) mvar_poly  ('b monomial) set" where
"monomials_of R P = {m. P m  𝟬R}"

context ring
begin

lemma monomials_of_index_free:
  assumes "index_free P i"
  assumes "m  monomials_of R P"
  shows "count m i = 0"
  using assms
  unfolding monomials_of_def index_free_def
  by (meson count_inI mem_Collect_eq)

lemma index_freeI:
  assumes "m. m  monomials_of R P  count m i = 0"
  shows "index_free P i"
  unfolding index_free_def
proof safe
  fix m
  assume "i ∈# m"
  then have "count m i  0"
    by simp
  then have "m  monomials_of R P"
    using assms
    by meson
  then show "P m = 𝟬"
    unfolding monomials_of_def
    by blast
qed

text‹\texttt{monomials\_of} R is subadditive›

lemma monomials_of_add:
"monomials_of R (P  Q)  (monomials_of R P)  (monomials_of R Q)"
proof
  fix x
  assume "x  monomials_of R (P  Q) "
  then have "P x  Q x 𝟬"
    by (simp add: indexed_padd_def monomials_of_def)
  then have "P x 𝟬  Q x  𝟬"
    by auto
  then show "x  (monomials_of R P)  (monomials_of R Q)"
    by (simp add: monomials_of_def)
qed

lemma monomials_of_add_finite:
  assumes "finite (monomials_of R P)"
  assumes "finite (monomials_of R Q)"
  shows "finite (monomials_of R (P  Q))"
  by (meson assms(1) assms(2) finite_Un finite_subset monomials_of_add)

lemma monomials_ofE:
  assumes "m  monomials_of R p"
  shows "p m  𝟬"
  using assms
  unfolding monomials_of_def
  by simp

lemma complement_of_monomials_of:
  assumes "m  monomials_of R P"
  shows "P m = 𝟬"
  using assms
  unfolding monomials_of_def
  by blast

text‹Multiplication by an indexed variable corresponds to adding that index to each monomial:›

lemma monomials_of_p_mult:
"monomials_of R (P  i) = {m.  n  (monomials_of R P). m = add_mset i n}"
proof
  show "monomials_of R (P  i)  {m. nmonomials_of R P. m = add_mset i n}"
    proof
      fix m
      assume A: "m  monomials_of R (P  i)"
      show "m  {m. nmonomials_of R P. m = add_mset i n}"
      proof-
        have "(P  i) m  𝟬"
          by (simp add: A monomials_ofE)
        then have "P (m - {# i #}) 𝟬"
          unfolding indexed_pmult_def
          by presburger
        then have 0: "(m - {# i #})  monomials_of R P"
          by (meson complement_of_monomials_of)
        have 1: " m = add_mset i  (m - {# i #})"
          by (metis (P  i) m  𝟬 add_mset_remove_trivial_If indexed_pmult_def)
        then show ?thesis using 0 1
          by blast
      qed
    qed
  show "{m. nmonomials_of R P. m = add_mset i n}  monomials_of R (P  i)"
    unfolding monomials_of_def indexed_pmult_def
    by auto
qed

lemma monomials_of_p_mult':
"monomials_of R (p  i) = add_mset i ` (monomials_of R p)"
  using  monomials_of_p_mult
  by (simp add: monomials_of_p_mult image_def)

lemma monomials_of_p_mult_finite:
  assumes "finite (monomials_of R P)"
  shows "finite (monomials_of R (P  i))"
  using assms monomials_of_p_mult'[of P i]
  by simp

text‹Monomials of a constant either consist of the empty multiset, or nothing:›

lemma monomials_of_const:
"(monomials_of R (indexed_const k)) = (if (k = 𝟬) then {} else {{#}})"
  unfolding monomials_of_def indexed_const_def
  by simp

lemma monomials_of_const_finite:
"finite (monomials_of R (indexed_const k))"
  by (simp add: monomials_of_const)

text‹A polynomial always has finitely many monomials:›
lemma monomials_finite:
  assumes "P  indexed_pset K I"
  shows "finite (monomials_of R P)"
  using assms
  apply(induction P)
  using monomials_of_const_finite apply blast
    using monomials_of_add_finite apply blast
      by (simp add: monomials_of_p_mult_finite)
end

(**************************************************************************************************)
(**************************************************************************************************)
    subsection‹Turning monomials into polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

text‹Constructor for turning a monomial into a polynomial›

definition mset_to_IP :: "('a, 'b) ring_scheme  'c monomial  ('a,'c) mvar_poly" where
"mset_to_IP R m n= (if (n = m) then 𝟭Relse 𝟬R)"

definition var_to_IP :: "('a, 'b) ring_scheme  'c  ('a,'c) mvar_poly" ("pvar") where
"var_to_IP R i = mset_to_IP R {#i#}"

context ring
begin

lemma mset_to_IP_simp[simp]:
"mset_to_IP R m m = 𝟭"
  by (simp add: mset_to_IP_def)

lemma mset_to_IP_simp'[simp]:
  assumes "n m"
  shows "mset_to_IP R m n = 𝟬"
  by (simp add: assms mset_to_IP_def)

lemma(in cring) monomials_of_mset_to_IP:
  assumes "𝟭 𝟬"
  shows "monomials_of R (mset_to_IP R m) = {m}"
  unfolding monomials_of_def mset_to_IP_def
proof
  show "{ma. (if ma = m then 𝟭 else 𝟬)  𝟬}  {m}"
    by auto
  show "{m}  {ma. (if ma = m then 𝟭 else 𝟬)  𝟬}"
    using assms by auto
qed

end

text‹The set of monomials of a fixed P› which include a given variable:›

definition monomials_with :: "('a, 'b) ring_scheme  'c  ('a,'c) mvar_poly  ('c monomial) set" where
"monomials_with R i P = {m. m  monomials_of R P  i ∈# m}"

context ring
begin

lemma monomials_withE:
  assumes "m  monomials_with R i P"
  shows "i ∈# m"
        "m  monomials_of R P"
  using assms unfolding monomials_with_def
  apply blast
     using assms unfolding monomials_with_def
     by blast

lemma monomials_withI:
  assumes "i ∈# m"
  assumes "m  monomials_of R P"
  shows "m  monomials_with R i P"
  using assms
  unfolding monomials_with_def
  by blast

end

text‹Restricting a polynomial to be zero outside of a given monomial set:›

definition restrict_poly_to_monom_set ::
  "('a, 'b) ring_scheme  ('a,'c) mvar_poly  ('c monomial) set ('a,'c) mvar_poly"  where
"restrict_poly_to_monom_set R P m_set m = (if m  m_set then P m else 𝟬R)"

context ring
begin

lemma restrict_poly_to_monom_set_coeff:
  assumes "carrier_coeff P"
  shows "carrier_coeff (restrict_poly_to_monom_set R P Ms)"
  by (metis assms carrier_coeff_def restrict_poly_to_monom_set_def zero_closed)

lemma restrict_poly_to_monom_set_coeff':
  assumes "P  indexed_pset K I"
  assumes "I  {}"
  assumes "m. P m  S"
  assumes "𝟬  S"
  shows "m.  (restrict_poly_to_monom_set R P m_set m)  S"
  using assms
  unfolding restrict_poly_to_monom_set_def
  by simp

lemma restrict_poly_to_monom_set_monoms:
  assumes "P  indexed_pset I K"
  assumes "m_set  monomials_of R P"
  shows "monomials_of R (restrict_poly_to_monom_set R P m_set) = m_set"
  using assms
  unfolding monomials_of_def restrict_poly_to_monom_set_def
  by (simp add: subset_iff)
end

(**************************************************************************************************)
(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Degree Functions›
(**************************************************************************************************)
(**************************************************************************************************)
(**************************************************************************************************)

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsubsection‹Total Degree Function›
  (**************************************************************************************************)
  (**************************************************************************************************)


lemma multi_set_size_count:
  fixes m :: "'c monomial"
  shows "size m  count m i"
  by (metis count_le_replicate_mset_subset_eq order_refl size_mset_mono size_replicate_mset)

text‹Total degree function›

definition total_degree :: "('a, 'b) ring_scheme  ('a, 'c) mvar_poly  nat" where
"total_degree R P = (if (monomials_of R P = {}) then 0 else Max (size ` (monomials_of R P)))"

context ring
begin

lemma total_degree_ineq:
  assumes "m  monomials_of R P"
  assumes "finite (monomials_of R P)"
  shows "total_degree R P  size m"
  unfolding total_degree_def using assms
  by force

lemma total_degree_in_monomials_of:
  assumes "monomials_of R P  {}"
  assumes "finite (monomials_of R P)"
  obtains m where "m  monomials_of R P  size m = total_degree R P"
    using assms Max_in[of "(size ` monomials_of R P)"] unfolding total_degree_def
    by (metis (mono_tags, lifting) empty_is_image finite_imageI image_iff)

lemma total_degreeI:
  assumes "finite (monomials_of R P)"
  assumes "m. m  monomials_of R P  size m = n"
  assumes "m. m  monomials_of R P  size m  n"
  shows "n = total_degree R P"
proof(cases "monomials_of R P = {}")
  case True
  then show ?thesis
    using assms by blast
next
  case False
  obtain m where m_def: "m  monomials_of R P  size m = n"
    using assms by blast
  have 0: "n  (size ` (monomials_of R P))"
    using m_def
    by blast
  have "k. k  (size ` (monomials_of R P))  k  n"
    using assms
    by blast
  then have 1: "m.  m  (monomials_of R P)  size m n"
    by blast
  obtain m' where m'_def: "m'  monomials_of R P  size m' = total_degree R P"
    using assms total_degree_in_monomials_of
    by blast
  then have 2: "size m'  n"
    using 1
    by blast
  have 3: "n size m'"
    using assms m'_def total_degree_ineq by auto
  show ?thesis using 2 3
    using dual_order.antisym m'_def by blast
qed
end

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsubsection‹Degree in One Variable›
  (**************************************************************************************************)
  (**************************************************************************************************)

definition degree_in_var ::
  "('a, 'b) ring_scheme  ('a, 'c) mvar_poly  'c  nat" where
"degree_in_var R P i =  (if (monomials_of R P = {}) then 0 else Max ((λm. count m i) ` (monomials_of R P)))"

context ring
begin

lemma degree_in_var_ineq:
  assumes "m  monomials_of R P"
  assumes "finite (monomials_of R P)"
  shows "degree_in_var R P i  count m i"
  unfolding degree_in_var_def using assms
  by force

lemma degree_in_var_in_monomials_of:
  assumes "monomials_of R P  {}"
  assumes "finite (monomials_of R P)"
  obtains m where "m  monomials_of R P  count m i = degree_in_var R P i"
  using assms Max_in[of "((λm. count m i) ` (monomials_of R P))"] unfolding degree_in_var_def
  by (metis (no_types, lifting) empty_is_image finite_imageI image_iff)

lemma degree_in_varI:
  assumes "finite (monomials_of R P)"
  assumes "m. m  monomials_of R P  count m i = n"
  assumes "c. c  monomials_of R P  count c i  n"
  shows "n = degree_in_var R P i"
proof-
  obtain l where l_def: "l  monomials_of R P  count l i = degree_in_var R P i"
    by (metis assms(1) assms(2) degree_in_var_in_monomials_of equals0D)
  have "degree_in_var R P i  n"
    using assms(3) l_def
    by force
  then show ?thesis
    using assms(1) assms(2) dual_order.antisym degree_in_var_ineq
    by fastforce
qed

text‹Total degree bounds degree in a single variable:›

lemma total_degree_degree_in_var:
  assumes "finite (monomials_of R P)"
  shows "total_degree R P  degree_in_var R P i"
proof(cases " (monomials_of R P) = {}")
  case True
  then show ?thesis
    unfolding total_degree_def degree_in_var_def
    by simp
next
  case False
  then   obtain m1 where m1_def: "m1  monomials_of R P  count m1 i = degree_in_var R P i"
    by (meson assms degree_in_var_in_monomials_of)
  have "size m1 count m1 i"
    by (simp add: multi_set_size_count)
  then show ?thesis
    by (simp add: assms local.ring_axioms m1_def order.trans ring.total_degree_ineq)
qed
end

text‹The set of monomials of maximal degree in variable i›:›

definition max_degree_monoms_in_var ::
  "('a, 'b) ring_scheme  ('a, 'c) mvar_poly  'c  ('c monomial) set" where
"max_degree_monoms_in_var R P i = {m. m  monomials_of R P  count m i = degree_in_var R P i}"

context ring
begin

lemma max_degree_monoms_in_var_memI:
  assumes "m  monomials_of R P"
  assumes "count m i = degree_in_var R P i"
  shows "m  max_degree_monoms_in_var R P i"
  using assms unfolding max_degree_monoms_in_var_def
  by blast

lemma max_degree_monoms_in_var_memE:
  assumes "m  max_degree_monoms_in_var R P i"
  shows   "m  monomials_of R P"
          "count m i = degree_in_var R P i"
  using assms unfolding max_degree_monoms_in_var_def
  apply blast
    using assms unfolding max_degree_monoms_in_var_def
    by blast
end

text‹The set of monomials of P› of fixed degree in variable i›:›

definition fixed_degree_in_var ::
  "('a, 'b) ring_scheme  ('a, 'c) mvar_poly  'c  nat  ('c monomial) set" where
"fixed_degree_in_var R P i n = {m. m  monomials_of R P  count m i = n}"

context ring
begin

lemma fixed_degree_in_var_subset:
"fixed_degree_in_var R P i n  monomials_of R P"
  unfolding fixed_degree_in_var_def
  by blast

lemma fixed_degree_in_var_max_degree_monoms_in_var:
"max_degree_monoms_in_var R P i = fixed_degree_in_var R P i (degree_in_var R P i)"
  unfolding max_degree_monoms_in_var_def fixed_degree_in_var_def
  by auto

lemma fixed_degree_in_varI:
  assumes "m  monomials_of R P"
  assumes "count m i = n"
  shows "m  fixed_degree_in_var R P i n"
  unfolding fixed_degree_in_var_def
  using assms
  by blast

lemma fixed_degree_in_varE:
  assumes "m  fixed_degree_in_var R P i n"
  shows "m  monomials_of R P"
          "count m i = n"
  apply (meson assms fixed_degree_in_var_subset in_mono)
    using assms fixed_degree_in_var_def by force

definition restrict_to_var_deg ::
 "('a,'c) mvar_poly  'c  nat  'c monomial  'a"  where
"restrict_to_var_deg P i n = restrict_poly_to_monom_set R P (fixed_degree_in_var R P i n)"

lemma restrict_to_var_deg_var_deg:
  assumes "finite (monomials_of R P)"
  assumes "Q = restrict_to_var_deg P i n"
  assumes "monomials_of R Q  {}"
  shows "n = degree_in_var R Q i"
  apply(rule degree_in_varI)
    apply (metis assms(1) assms(2) fixed_degree_in_varE(1) monomials_ofE restrict_poly_to_monom_set_def
      restrict_to_var_deg_def rev_finite_subset subsetI)
      apply (metis (full_types) assms(2) assms(3) equals0I fixed_degree_in_varE(2)
      monomials_ofE restrict_poly_to_monom_set_def restrict_to_var_deg_def)
        by (metis assms(2) eq_iff fixed_degree_in_varE(2) monomials_ofE restrict_poly_to_monom_set_def restrict_to_var_deg_def)

lemma index_free_degree_in_var[simp]:
  assumes "index_free P i"
  shows "degree_in_var R P i = 0"
proof(cases "monomials_of R P = {}")
  case True
  then show ?thesis
    using assms
    unfolding degree_in_var_def
    by simp
next
  case False
  then have 0: "degree_in_var R P i =  Max ((λm. count m i) ` (monomials_of R P))"
    unfolding degree_in_var_def
    by simp
  have"((λm. count m i) ` (monomials_of R P)) = {0}"
  proof
    show "(λm. count m i) ` monomials_of R P  {0}"
      using False assms monomials_of_index_free[of P i]
      by auto
    show "{0}  (λm. count m i) ` monomials_of R P"
      using False  (λm. count m i) ` monomials_of R P  {0}
      by auto
  qed
  then show ?thesis using 0
    by simp
qed

lemma degree_in_var_index_free:
  assumes "degree_in_var R P i = 0"
  assumes "finite (monomials_of R P)"
  shows "index_free P i"
  apply(rule index_freeI)
  by (metis assms(1) assms(2) degree_in_var_ineq le_zero_eq)

end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Constructing the Multiplication Operation on the Ring of Indexed Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

  (**********************************************************************)
  (**********************************************************************)
  subsubsection‹The Set of Factors of a Fixed Monomial›
  (**********************************************************************)
  (**********************************************************************)

  text‹The following function maps a monomial to the set of monomials which divide it:›

definition mset_factors :: "'c monomial  ('c monomial) set" where
"mset_factors m = {n. n ⊆# m}"

context ring
begin

lemma monom_divides_factors:
"n  (mset_factors m) n ⊆# m"
  unfolding mset_factors_def by auto

lemma mset_factors_mono:
  assumes "n ⊆# m"
  shows "mset_factors n  mset_factors m"
  unfolding mset_factors_def
  by (simp add: Collect_mono_iff assms subset_mset.order_trans)

lemma mset_factors_size_bound:
  assumes "n  mset_factors m"
  shows "size n  size m"
  using assms
  unfolding mset_factors_def
  by (simp add: size_mset_mono)

lemma sets_to_inds_finite:
  assumes "finite I"
  shows "finite S  finite (PiE S (λ_. I))"
  using assms
  by (simp add: finite_PiE)

end

    (**********************************************************************************************)
    (**********************************************************************************************)
    subsubsection‹Finiteness of the Factor Set of a Monomial›
    (**********************************************************************************************)
    (**********************************************************************************************)

text‹
  This section shows that any monomial m has only finitely many factors. This is done by mapping
  the set of factors injectively into a finite extensional function set. Explicitly, a monomial is
  just mapped to its count function, restricted to the set of indices where the count is nonzero.
›

definition mset_factors_to_fun ::
  "('a,'b) ring_scheme  'c monomial  'c monomial  ('c  nat)" where
"mset_factors_to_fun R m n = (if (n  mset_factors m) then
                                           (restrict  (count n) (set_mset m)) else undefined)"
context ring
begin

lemma mset_factors_to_fun_prop:
  assumes "size m = n"
  shows "mset_factors_to_fun R m  (mset_factors m)  (PiE (set_mset m) (λ_. {0.. n}))"
proof
  fix x
  assume A: "x  (mset_factors m)"
  show "mset_factors_to_fun R m x  (set_mset m) E {0..n} "
  proof
    show "xa. xa ∈# m  mset_factors_to_fun R m x xa  {0..n}"
      proof-
        fix y
        assume Ay: "y ∈# m"
        show P0: "mset_factors_to_fun R m x y  {0..n}"
        proof-
          have "mset_factors_to_fun R m x = restrict (count x) (set_mset m)"
            using A  unfolding mset_factors_to_fun_def
            by simp
          then have "mset_factors_to_fun R m x y  = count x y"
            using Ay
            by simp
          then show ?thesis
            using A INTEG.R.mset_factors_size_bound assms count_size by fastforce
        qed
      qed
    fix z
    assume Ay: "z ∉#m"
    show "mset_factors_to_fun R m x z = undefined"
      using A Ay unfolding mset_factors_to_fun_def
      by simp
  qed
qed

lemma mset_factors_to_fun_inj:
  shows "inj_on (mset_factors_to_fun R m) (mset_factors m) "
proof
  fix x y
  assume A: "x  mset_factors m" "y  mset_factors m"
  show "mset_factors_to_fun R m x = mset_factors_to_fun R m y  x = y"
  proof-
    assume A0:  "mset_factors_to_fun R m x = mset_factors_to_fun R m y"
    show " x = y"
    proof-
      have "i. count x i = count y i"
      proof- fix i
        show "count x i = count y i"
        proof(cases "i ∈# m")
          case True
          then show ?thesis using A0 A unfolding mset_factors_to_fun_def
            by (metis restrict_def)
        next
          case False
          then show ?thesis
             using A0 A unfolding mset_factors_to_fun_def
             by (metis monom_divides_factors count_inI mset_subset_eqD)
        qed
      qed
      then show ?thesis
        using multiset_eqI by blast
    qed
  qed
qed

lemma finite_target:
"finite (PiE (set_mset m) (λ_. {0..(n::nat)}))"
proof-
  have 0: "finite (set_mset m)"
    by simp
  have 1: "finite ({0..n})"
    by simp
  then show ?thesis using 0
    by (simp add: finite_PiE)
qed

text‹A multiset has only finitely many factors:›

lemma mset_factors_finite[simp]:
"finite (mset_factors m)"
proof-
  have 0: "inj_on (mset_factors_to_fun R m) (mset_factors m) "
    by (simp add: mset_factors_to_fun_inj)
  have 1: "(mset_factors_to_fun R m)  (mset_factors m)  (PiE (set_mset m) (λ_. {0 .. (size m)}))"
    by (metis mset_factors_to_fun_prop)
  have 2: "finite (PiE (set_mset m) (λ_. {0 .. (size m)}))"
    by (simp add: finite_target)
  have 3: "((mset_factors_to_fun R m) ` (mset_factors m))  (PiE (set_mset m) (λ_. {0 .. (size m)}))"
    using 1 2
    by blast
  then have "finite ((mset_factors_to_fun R m) ` (mset_factors m))"
    using 2 finite_subset by auto
  then show ?thesis using 0 1 2
    finite_imageD[of "mset_factors_to_fun R m" "mset_factors m" ]
    by blast
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Definition of Indexed Polynomial Multiplication.›
(**************************************************************************************************)
(**************************************************************************************************)

context ring
begin

text‹Monomial division:›

lemma monom_divide:
  assumes "n  mset_factors m"
  shows "(THE k. n + k = m ) = m - n "
  apply(rule the_equality)
  using assms unfolding mset_factors_def
  apply simp
    using assms unfolding mset_factors_def
    by auto

text‹A monomial is a factor of itself:›

lemma m_factor[simp]:
"m  mset_factors m"
  using local.ring_axioms ring.monom_divides_factors by blast

end

text‹The definition of polynomial multiplication:›

definition P_ring_mult :: "('a, 'b) ring_scheme  ('a,'c) mvar_poly  ('a,'c) mvar_poly  'c monomial  'a"
  where
"P_ring_mult R P Q m = (finsum R (λx. (P x) R(Q (m - x))) (mset_factors m))"

abbreviation(in ring) P_ring_mult_in_ring (infixl "p" 65)where
"P_ring_mult_in_ring   P_ring_mult R"

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Distributivity Laws for Polynomial Multiplication›
(**************************************************************************************************)
(**************************************************************************************************)

context ring
begin

lemma P_ring_rdistr:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  shows "a p (b  c) = (a p b) (a p c)"
proof
  fix m
  show "(a p (b  c)) m = (a p b  (a p c)) m"
  proof-
    have RHS: "(a p b  (a p c)) m =
                  (xmset_factors m. a x  b (m - x))  (xmset_factors m. a x  c (m - x))"
      unfolding indexed_padd_def P_ring_mult_def by auto
    have LHS: "(a p (b  c)) m=
                  ( xmset_factors m. a x  b (m - x)  a x  c (m - x))"
      unfolding indexed_padd_def  P_ring_mult_def
      by (meson assms(1) assms(2) assms(3) local.ring_axioms r_distr ring.carrier_coeff_def)
    have RHS': "(a p b  (a p c)) m =
                  (xmset_factors m. a x  b (m - x)  a x  c (m - x))"
    proof-
      have 0: "(λx. a x  b (m - x))  mset_factors m  carrier R"
      proof
        fix x assume "x  mset_factors m"
        then show "a x  b (m - x)  carrier R"
          using assms carrier_coeffE
          by blast
      qed
      have 1: "(λx. a x  c (m - x))  mset_factors m  carrier R"
      proof
        fix x assume "x  mset_factors m"
        then show "a x  c (m - x)  carrier R"
          using assms carrier_coeffE
          by blast
      qed
      then show ?thesis
        using 0 1 RHS assms finsum_addf[of "λx. a x  b (m - x)" "mset_factors m"
                                                    "λx. a x  c (m - x)"]
        by metis
    qed
    show ?thesis using LHS RHS' by auto
  qed
qed

lemma P_ring_ldistr:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  shows "  (b  c) p  a = (b p a) (c p a)"
proof
  fix m
  show "((b  c) p  a) m = ((b p a) (c p a)) m"
  proof-
    have RHS: "((b p a) (c p a)) m =
                  (xmset_factors m. b x  a (m - x))  (xmset_factors m. c x  a (m - x))"
      unfolding indexed_padd_def P_ring_mult_def by auto
    have LHS: "((b  c) p  a)  m=
                  ( xmset_factors m. b x  a (m - x)  c x  a (m - x))"
      unfolding indexed_padd_def  P_ring_mult_def
      by (meson assms(1) assms(2) assms(3) l_distr local.ring_axioms ring.carrier_coeff_def)
    have RHS': "((b p a) (c p a)) m =
                  (xmset_factors m. b x  a (m - x)  c x  a (m - x))"
    proof-
      have 0: "(λx. b x  a (m - x))  mset_factors m  carrier R"
      proof
        fix x assume "x  mset_factors m"
        then show "b x  a (m - x)  carrier R"
          using assms carrier_coeffE
          by blast
      qed
      have 1: "(λx. c x  a (m - x))  mset_factors m  carrier R"
      proof
        fix x assume "x  mset_factors m"
        then show "c x  a (m - x)  carrier R"
          using assms carrier_coeffE
          by blast
      qed
      then show ?thesis
        using 0 1 RHS assms finsum_addf[of "λx. b x  a (m - x)" "mset_factors m"
                                                    "λx. c x  a (m - x)"]
        by metis
    qed
    show ?thesis using LHS RHS' by auto
  qed
qed
end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Multiplication Commutes with \texttt{indexed\_pmult}›
(**************************************************************************************************)
(**************************************************************************************************)

context ring
begin

text‹This lemma shows how we can write the factors of a monomial $m$ times a variable $i$ in terms
      of the factors of m:›

lemma mset_factors_add_mset:
"mset_factors (add_mset i m) = mset_factors m  add_mset i ` (mset_factors m)"
proof
  show "mset_factors (add_mset i m)  mset_factors m  add_mset i ` mset_factors m"
  proof fix x
    assume A: "x  mset_factors (add_mset i m)"
    show "x  mset_factors m  add_mset i ` mset_factors m"
    proof(cases "i ∈# x")
      case True
      then have "x - {#i#} ⊆# m"
        using A INTEG.R.monom_divides_factors subset_eq_diff_conv by force
      then show ?thesis
        by (metis INTEG.R.monom_divides_factors True UnI2 add_mset_remove_trivial image_iff multi_member_split)
    next
      case False
      have 0: "x ⊆# add_mset i m"
        using A INTEG.R.monom_divides_factors by blast
      hence "x ⊆#  m"
        using mset_subset_eqI[of x m] 0 False
        by (metis count_add_mset count_greater_zero_iff count_inI less_Suc_eq_le
            subseteq_mset_def union_single_eq_member)
      thus ?thesis
        using INTEG.R.monom_divides_factors by blast
    qed
  qed
  show "mset_factors m  add_mset i ` mset_factors m  mset_factors (add_mset i m)"
  proof fix x assume A: "x  mset_factors m  add_mset i ` mset_factors m"
    have "x ⊆# add_mset i m"
    proof(cases "x  mset_factors m")
      case True
      then have "x ⊆# m"
        by (simp add: INTEG.R.monom_divides_factors)
      hence "(a. count x a  count (add_mset i m) a)"
        using mset_subset_eq_count[of x m] count_add_mset[of i m]
              nat_le_linear not_less_eq_eq by fastforce
      thus ?thesis
        using mset_subset_eqI by blast
    next
      case False
      then obtain n where n_def: "n  mset_factors m  x = add_mset i n"
        using A by blast
      then have "x ⊆# add_mset i m"
        by (simp add: INTEG.R.monom_divides_factors)
      then show ?thesis
        by simp
    qed
    thus "x  mset_factors (add_mset i m)"
      by (simp add: INTEG.R.monom_divides_factors)
  qed
qed

end
(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Associativity of Polynomial Multiplication.›
(**************************************************************************************************)
(**************************************************************************************************)
context ring
begin

lemma finsum_eq:
  assumes "f  S  carrier R"
  assumes "g  S  carrier R"
  assumes "(λ x  S. f x) = (λ x  S. g x)"
  shows " finsum R f S = finsum R g S"
        by (metis assms(1) assms(3) finsum_cong' restrict_apply')

lemma finsum_eq_induct:
  assumes "f  S  carrier R"
  assumes "g  T  carrier R"
  assumes "finite S"
  assumes "finite T"
  assumes "bij_betw h S T"
  assumes "s. s  S  f s = g (h s)"
  shows "finite U  U  S  finsum R f U = finsum R g (h ` U)"
  apply(induct rule: finite_induct)
  apply (simp; fail)
proof-
  fix x
  fix F :: "'c set"
  assume F_fin: "finite F"
  assume x_not_in: "x  F"
  assume IH: "(F  S  finsum R f F = finsum R g (h ` F)) "
  show "insert x F  S  finsum R f (insert x F) = finsum R g (h ` insert x F)"
  proof-
    assume A: "insert x F  S"
    then have F_sub: "F S"
      by simp
    have x_in: "x  S"
      using A by blast
    have fx_in: "f x   carrier R"
      using x_in assms(1)
      by auto
    have ghx_fx_eq:
      "f x = g (h x)"
      using x_in assms
      by blast
    have I: "finsum R f F = finsum R g (h ` F)"
      using F_sub IH by blast
    have f_fact: "f  F  carrier R "
      using assms(1) F_sub
      by blast
    have "finsum R f (insert x F) = (f x) add_monoid R(finsum R f F) "
      using F_fin  x_not_in comm_monoid.finprod_insert[of "(add_monoid R)" F x f ]
      unfolding finsum_def
      using f_fact fx_in local.add.comm_monoid_axioms
      by auto
    have "finsum R g (h ` insert x F) = finsum R g (insert (h x) (h ` F))"
      by simp
    then have "finsum R g (h ` insert x F) = (g (h x)) add_monoid Rfinsum R g (h ` F)"
    proof-
       have 0: "finite (h ` F)"
         by (simp add: F_fin)
       have 1: "h x  h ` F"
       proof-
         have 10: "bij_betw h F (h` F)"
           using assms(5) F_sub bij_betw_subset
           by blast
         show ?thesis
         proof
           assume "h x  h ` F "
           then  obtain s where s_def: "s  F  h s = h x"
             using 10
             by auto
           have "s  S"
             using F_sub s_def by blast
           then have "s = x"
             using x_in assms(5)
               s_def
             unfolding bij_betw_def inj_on_def
             by blast
           then show False using x_not_in
             using s_def by blast
         qed
       qed
       have 2: "g  h ` F  carrier (add_monoid R)"
       proof
         fix y
         assume Ay: "y  h ` F"
         then obtain s where s_def: "y = h s  s  F"
           by blast
         then have s_S: "s  S"
           using F_sub by blast
         have "h ` F  T"
           using F_sub assms(5)
           unfolding bij_betw_def
           by blast
         then have "g y  carrier R"
           using Ay assms(2)
           by blast
         then show "g y  carrier (add_monoid R)"
           by simp
       qed
       have "g (h x)  carrier (add_monoid R)"
         using fx_in ghx_fx_eq
         by auto
       then show ?thesis
         using 0 1 2 comm_monoid.finprod_insert[of "(add_monoid R)" "(h ` F)"  "h x" g ]
         unfolding finsum_def
         by (simp add: local.add.comm_monoid_axioms)
    qed
    then have "finsum R g (h ` insert x F) = f x  add_monoid R(finsum R f F)"
      using I ghx_fx_eq by auto
    then show ?thesis
      by (simp add: finsum R f (insert x F) = f x add_monoid Rfinsum R f F)
  qed
qed

lemma finsum_bij_eq:
  assumes "f  S  carrier R"
  assumes "g  T  carrier R"
  assumes "finite S"
  assumes "bij_betw h S T"
  assumes "s. s  S  f s = g (h s)"
  shows "finsum R f S = finsum R g T"
proof-
  have 0: "finite T"
    using assms bij_betw_finite
    by blast
  have 1: "(s. s  S  f s = g (h s)) "
    using assms
    by blast
  have "(s. s  S  f s = g (h s))  finite S  S  S  finsum R f S = finsum R g (h ` S)"
    using 0 assms finsum_eq_induct[of f S g T h S ]
  by blast
  then have "finsum R f S = finsum R g (h ` S)"
    using assms(5) assms(3)
    by blast
  then show ?thesis
    using assms(5) assms(4) bij_betw_imp_surj_on
    by blast
qed

lemma monom_comp:
  assumes "x ⊆# m"
  assumes "y ⊆# m - x"
  shows "x ⊆# m - y"
using assms
  by (metis add_diff_cancel_left' subset_eq_diff_conv
      subset_mset.le_diff_conv2 subset_mset.order_refl subset_mset.order_trans)

lemma monom_comp':
  assumes "x ⊆# m"
  assumes "y = m - x"
  shows "x = m - y"
  using assms
  by (metis add_diff_cancel_right' subset_mset.add_diff_inverse)

text ‹
  This lemma turns iterated sums into sums over a product set. The first lemma is only a technical
  phrasing of \texttt{double\_finsum'} to facilitate an inductive proof, and likely can and should
  be simplified.
›

lemma double_finsum_induct:
  assumes "finite S"
  assumes "s. s  S  finite (F s)"
  assumes "P = (λS. {(s, t). s  S   t  (F s)})"
  assumes "s y. s  S  y  (F s)  g s y  carrier R"
  shows "finite T  T  S  finsum R (λs. finsum R (g s) (F s)) T =
         finsum R (λ c. g (fst c) (snd c)) (P  T)"
  apply(induct rule: finite_induct)
  apply (simp add: assms(3); fail)
proof-
  fix x
  fix T :: "'c set"
  assume AT: "finite T"
  assume x_notin: "x  T"
  assume IH: "(T  S  (sT. finsum R (g s) (F s)) = (cP T. g (fst c) (snd c)))"
  assume A: "insert x T S"
  show " (sinsert x T. finsum R (g s) (F s)) = (cP (insert x T). g (fst c) (snd c))"
  proof-
    have 0: "(λs. finsum R (g s) (F s))  T  carrier R"
    proof
      fix v
      assume A0: "v  T"
      then have A1: "v  S"
        using A by blast
      then show "finsum R (g v) (F v)  carrier R"
      proof-
        have 00: "finite (F v)"
          using assms A
          using v  T by blast
        have "g v  F v  carrier R"
        proof
          fix l assume "l  F v"
          then show "g v l  carrier R"
            using A1 assms(4)[of v l]
            by blast
        qed
        then show ?thesis
          using finsum_closed by blast
      qed
    qed
    have 1: "finsum R (g x) (F x)  carrier R"
    proof-
      have "x  S"
        using A by blast
      then show ?thesis using assms(4) finsum_closed[of "(g x)" "F x"]
        by blast
    qed
    have 2: " (sinsert x T. finsum R (g s) (F s)) = finsum R (g x) (F x)  (s T. finsum R (g s) (F s))"
      using 0 1 finsum_insert[of T x "(λs. finsum R (g s) (F s))"] AT x_notin
      by blast
    have 3: "P (insert x T) = P {x}  P T"
    proof
      show "P (insert x T)  P {x}  P T"
      proof
        fix y
        assume "y  P (insert x T)"
        then show "y  P {x}  P T"
          using assms
          by blast
      qed
      show "P {x}  P T  P (insert x T)"
      proof
        fix y
        assume Ay: "y  P {x}  P T"
        show "y   P (insert x T)"
        proof(cases "y  P {x}")
          case True
          then have "y  ({(s, t). s  {x}  t  F s})"
            using assms(3) by auto
          then have "fst y = x  snd y  F x"
            using Product_Type.Collect_case_prodD by auto
          then show ?thesis using assms(3)
            using fst_def by auto
        next
          case False
          then have "y  P T"
            using Ay
            by blast
          then have "y  ({(s, t). s  T  t  F s})"
            using assms(3) by blast
          then obtain s t where st_def: "y = (s, t) s  T  t  F s"
            by blast
          then have "y = (s, t) s  (insert x T)  t  F s"
            by blast
          then show ?thesis using assms(3)
            by simp
        qed
      qed
    qed
    have 4: "P {x}  P T = {}"
    proof
      show "P {x}  P T  {}"
      proof
        fix y
        assume B: "y  P {x}  P T"
        then have "fst y = x"
        proof-
          have "y  {(s, t). s  {x}  t  F s}"
            using assms(3) B by auto
          then obtain s t where "y = (s, t)  s  {x}  t  F s"
            by blast
          then show ?thesis
            by auto
        qed
        have "fst y  T"
        proof-
          have "y  {(s, t). s  T  t  F s}"
            using assms(3) B by auto
          then obtain s t where "y = (s, t)  s  T  t  F s"
            by blast
          then show ?thesis
            by simp
        qed
        then have False
          using x_notin
          by (simp add: fst y = x)
        then show "y  {}"
          by simp
      qed
      show "{}  P {x}  P T"
        by simp
    qed
    have 5: "(λc. g (fst c) (snd c))  P {x}  carrier R"
    proof
      fix y
      assume X0: "y  P {x}"
      obtain s t where st_def: "y = (s, t)  (s, t)  P {x}"
        by (metis X0 old.prod.exhaust)
      then have st: "s = x  t  F x"
        using assms(3) by blast
      then have "g (fst y) (snd y) = g x t t  F x"
        by (simp add: st_def)
      then show "g (fst y) (snd y)  carrier R"
        using assms(4)[of x t] A
        by simp
    qed
    have 6: "(λc. g (fst c) (snd c))  P T  carrier R"
    proof
      fix y
      assume X0: "y  P T"
      obtain s t where st_def: "y = (s, t)  (s, t)  P T"
        by (metis X0 old.prod.exhaust)
      then have st: "s  T  t  F s"
        using assms(3) by blast
      then have "g (fst y) (snd y) = g s t t  F s"
        by (simp add: st_def)
      then show "g (fst y) (snd y)  carrier R"
        using assms(4)[of s t] A st
        by auto
    qed
    have 07: "x. x  S   finite (P {x})"
    proof-
      fix x
      assume "x  S"
      have "bij_betw snd (P {x}) (F x)"
        unfolding bij_betw_def
      proof
        show "inj_on snd (P {x})"
        proof
          fix a b
          assume Aa: "a  P {x}"
          assume Ab: "b  P {x}"
          have 0: "fst a = x"
          proof-
            have" a  {(s, t). s  {x}  t  F s} "
              using Aa  assms(3)
              by blast
            then obtain s t where st_def: "a = (s, t)  s  {x}  t  F s"
              by blast
            then show ?thesis
              by auto
          qed
          have 1: "fst b = x"
          proof-
            have" b  {(s, t). s  {x}  t  F s} "
              using Ab  assms(3)
              by blast
            then obtain s t where st_def: "b = (s, t)  s  {x}  t  F s"
              by blast
            then show ?thesis
              by auto
          qed
          show "snd a = snd b  a = b"
            using 0 1
            by (simp add: prod_eqI)
        qed
        show "snd ` P {x} = F x"
        proof
          show "snd ` P {x}  F x"
          proof
            fix y
            assume 0: "y  snd ` P {x}"
            then obtain q where q_def: " q  P {x}  y = snd q"
              by blast
            then obtain s t where st: "q = (s, t)  s  {x}  t  F s"
              using assms(3) by blast
            then have  1: "s = x"
              by blast
            have 2: "snd q = t"
              by (simp add: st)
            then show "y  F x"
              using q_def st by blast
          qed
          show "F x  snd ` P {x}"
          proof
            fix y
            assume "y  F x"
            then have C: "(x, y)  P {x}"
              using assms(3)
              by simp
            then have "y = snd (x, y)"
              by auto
            then show "y  snd ` P {x}"
              using C
              by blast
          qed
        qed
      qed
      then show "finite (P {x})"
        using assms(2)[of x] bij_betw_finite
        x  S
        by blast
    qed
    have 7: "finite (P {x})"
      using 07 A
      by blast
    have 8: "finite (P T)"
    proof-
      have "D. finite D  D  S  finite (P D)"
      proof-
        fix D
        show "finite D  D  S   finite (P D)"
          apply(induct rule: finite_induct)
        proof-
          have "P {} = {}" using assms(3)
            by blast
          then show "finite (P {})"
            by auto
          show "x F. finite F  x  F  (F  S  finite (P F))  insert x F  S  finite (P (insert x F))"
          proof-
            fix x
            fix Q :: "'c set "
            assume fin: "finite Q"
            assume notin: "x  Q"
            assume fin_pf: "(Q  S  finite (P Q))"
            assume I: "insert x Q  S "
            show "finite (P (insert x Q))"
            proof-
              have x_in: "x  S"
                using I by blast
              have 0: "(P (insert x Q))  (P Q)  P {x}"
              proof
                fix y
                assume "y  P (insert x Q)"
                obtain s t where st: "y = (s, t)  s  (insert x Q)  t  F s"
                  using assms(3) x_in  y  P (insert x Q)
                  by blast
                show "y  (P Q)  P {x}"
                proof(cases "s  Q")
                  case True
                  then have "y  P Q"
                    using st assms(3)
                    by simp
                  then show ?thesis using st assms(3)
                    by blast
                next
                  case False
                  then have "s = x"
                    using st by blast
                  then have  "s {x}  t  F s"
                    using st by blast
                  then have "y  P {x}"
                    using st assms(3)
                    by auto
                  then show ?thesis by auto
                qed
              qed
              have 1: "finite  (P Q)"
                using I fin_pf by blast
              have "finite (P {x})"
                using "07" x_in by blast
              then show ?thesis using 0 1
                using finite_subset by auto
            qed
          qed
        qed
      qed
      then show ?thesis
        using A AT by blast
    qed
    have 9: "(qP (insert x T). g (fst q) (snd q)) =
          (qP T. g (fst q) (snd q))  (qP {x}. g (fst q) (snd q))"
      using 8 7 6 5 4 3 finsum_Un_disjoint[of "P {x}" "P T" "λq. g (fst q) (snd q)"]
      by (simp add: "7" finite (P {x}); finite (P T); P {x}  P T = {}; (λp. g (fst p) (snd p))
            P {x}  carrier R; (λp. g (fst p) (snd p))  P T  carrier R 
           (pP {x}  P T. g (fst p) (snd p)) = (pP {x}. g (fst p) (snd p)) 
             (pP T. g (fst p) (snd p)) add.m_comm)
    have 10: "(pP (insert x T). g (fst p) (snd p))
        =  (sT. finsum R (g s) (F s)) (pP {x}. g (fst p) (snd p))"
      using IH 9 A
      by auto
    have 11: "(pP {x}. g (fst p) (snd p)) = finsum R (g x) (F x)"
    proof-
      obtain h :: "('c × 'd)  'd" where h_def: "h = snd"
        by simp
      have 110: "bij_betw h (P {x}) (F x)"
        unfolding bij_betw_def
      proof
        show "inj_on h (P {x})"
          unfolding inj_on_def
        proof
          fix q
          assume Ap: "q   P {x}"
          show " yP {x}. h q = h y  q = y"
          proof
            fix y
            assume q_def: "y  P {x}"
            then have C0: "fst y = x"
               using assms(3)
               by (simp add: case_prod_beta)
            have C1: "fst q = x"
              using assms(3) Ap
               by (simp add: case_prod_beta)
             show  "h q = h y  q = y"
               using C0 C1 h_def
               by (simp add: prod_eq_iff)
          qed
        qed
        show "h ` P {x} = F x"
        proof
          show "h ` P {x}  F x"
          proof
            fix y
            assume "y  h ` P {x}"
            then obtain d where d_def: "y = h d  d  P {x}"
              by blast
            then obtain s t where st_def: "d = (s, t)"
              by (meson surj_pair)
            then have "s = x  t  F x"
              using assms(3) d_def
              by blast
            then  show "y  F x"
              using assms(3)
              by (simp add: d_def h_def st_def)
          qed
          show "F x  h ` P {x}"
          proof
            fix y
            assume E0:  "y  F x"
            have E1: "y = h (x , y)"
              by (simp add: h_def)
            have "(x, y)  P {x}"
              using E0 assms(3) by blast
            then show "y  h ` P {x} "
              using E1 assms(3) by blast
          qed
        qed
      qed
      have 111: "g x  F x  carrier R "
      proof
        fix y
        assume "y  F x"
        then show "g x y  carrier R"
          using assms A
          by blast
      qed
      have 112: "(s. s  P {x}  g (fst s) (snd s) = g x (h s))"
      proof-
        fix s
        assume "s  P{x}"
        then have "s  {(s, t). s = x  t  F s}"
          using assms(3) by blast
        then obtain t where "s = (x, t)  t  F x"
          using assms(3)
          by blast
        then show "g (fst s) (snd s) = g x (h s)"
          by (simp add: h_def)
      qed



      show ?thesis using 5 7  110 111 112 finsum_bij_eq[of "λp. g (fst p) (snd p)" "P {x}" "g x" "F x" h ]
        by auto
    qed
    have 12: "(pP (insert x T). g (fst p) (snd p))
        =  (sT. finsum R (g s) (F s)) finsum R (g x) (F x)"
      using 10 11 by auto
    then show ?thesis using 2
      by (simp add: "2" "0" "1" add.m_comm)
  qed
qed

lemma double_finsum:
  assumes "finite S"
  assumes "s. s  S  finite (F s)"
  assumes "P = {(s, t). s  S   t  (F s)}"
  assumes "s y. s  S  y  (F s)  g s y  carrier R"
  shows "finsum R (λs. finsum R (g s) (F s)) S =
         finsum R (λ p. g (fst p) (snd p)) P"
proof-
  obtain P' where P'_def: "P' =  (λS. {(s, t). s  S   t  (F s)})"
    by simp
  have "finsum R (λs. finsum R (g s) (F s)) S =
         finsum R (λ p. g (fst p) (snd p)) (P' S)"
    using double_finsum_induct[of S F P' g S]
          assms P'_def
    by blast
  then show ?thesis
    using P'_def assms
    by blast
qed

end

text‹
  The product index set for the double sums in the coefficients of the
  $((a \otimes_p b) \otimes_p c)$. It is the set of pairs $(x,y)$ of monomials, such that
  $x$ is a factor of monomial $m$, and $y$ is a factor of monomial $x$.
›

definition right_cuts :: "'c monomial  ('c monomial × 'c monomial) set" where
"right_cuts m = {(x, y). x ⊆# m  y ⊆# x}"

context ring
begin

lemma right_cuts_alt_def:
"right_cuts m = {(x, y). x  mset_factors m  y  mset_factors x}"
  unfolding mset_factors_def right_cuts_def
  by simp

lemma right_cuts_finite:
"finite (right_cuts m)"
proof-
  have "finite (mset_factors m × mset_factors m)"
    using mset_factors_finite
    by blast
  have "right_cuts m  (mset_factors m × mset_factors m)"
  proof
    fix p
    assume p_def: "p  right_cuts m"
    obtain x y where xy: "p = (x , y)  x ⊆# m  y ⊆#x"
      using p_def unfolding right_cuts_def
      by blast
    then have "x  mset_factors m  y  mset_factors m"
      using monom_divides_factors
      by auto
    then show "p  (mset_factors m × mset_factors m)"
      by (simp add: xy)
  qed
  then show ?thesis
    using finite (mset_factors m × mset_factors m) finite_subset
    by blast
qed

lemma assoc_aux0:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  assumes "g = (λx y.  a x  (b y  c (m - x - y)))"
  shows "s y. s  mset_factors m  y  mset_factors (m - x)
             g s y  carrier R"
  using assms carrier_coeffE by blast

lemma assoc_aux1:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  assumes "g = (λx y.  (a y   b (x - y))  c (m - x))"
  shows "s y. s  mset_factors m  y  mset_factors x  g s y  carrier R"
  using assms carrier_coeffE by blast
end

text‹
  The product index set for the double sums in the coefficients of the
  $(a \otimes_p (b \otimes_p c))$. It is the set of pairs $(x,y)$ such that $x$ is a factor
  of $m$ and $y$ is a factor of $m/x$.
›

definition left_cuts :: "'c monomial  ('c monomial × 'c monomial) set" where
"left_cuts m = {(x, y). x ⊆#m  y ⊆# (m - x)}"

context ring
begin

lemma left_cuts_alt_def:
"left_cuts m = {(x, y). x  mset_factors m  y  mset_factors (m - x)}"
  unfolding mset_factors_def left_cuts_def
  by simp

text‹This lemma witnesses the bijection between left and right cuts for the proof of associativity:›

lemma left_right_cuts_bij:
"bij_betw (λ (x,y). (x + y, x))  (left_cuts m) (right_cuts m)"
  unfolding bij_betw_def right_cuts_def left_cuts_def
proof
  show "inj_on (λ(x, y). (x + y, x)) {(x, y). x ⊆# m  y ⊆# m - x}"
    unfolding inj_on_def
    by auto
  show "(λ(x, y). (x + y, x)) ` {(x, y). x ⊆# m  y ⊆# m - x} = {(x, y). x ⊆# m  y ⊆# x}"
  proof
    show "(λ(x, y). (x + y, x)) ` {(x, y). x ⊆# m  y ⊆# m - x}  {(x, y). x ⊆# m  y ⊆# x}"
    proof
      fix p
      assume "p   (λ(x, y). (x + y, x)) ` {(x, y). x ⊆# m  y ⊆# m - x}"
      then obtain a b where ab: "a ⊆# m  b ⊆# m - a  p = (λ(x, y). (x + y, x)) (a, b)"
        by blast
      then have "p = (a + b, a)"
        by simp
      then show "p  {(x, y). x ⊆# m  y ⊆# x}"
        using ab
        by (metis (no_types, lifting) case_prodI mem_Collect_eq mset_subset_eq_add_left
            subset_mset.le_diff_conv2 union_commute)
    qed
    show "{(x, y). x ⊆# m  y ⊆# x}  (λ(x, y). (x + y, x)) ` {(x, y). x ⊆# m  y ⊆# m - x}"
    proof
      fix p
      assume p_def: "p  {(x, y). x ⊆# m  y ⊆# x}"
      then obtain a b where ab: "p = (a, b)  a ⊆# m  b ⊆# a"
        by blast
      then have "p = (λ(x, y). (x + y, x)) (b, a - b)"
        using ab
        by simp
      then show " p  (λ(x, y). (x + y, x)) ` {(x, y). x ⊆# m  y ⊆# m - x}"
        by (metis (mono_tags, lifting) ab case_prodI image_eqI mem_Collect_eq
            subset_mset.diff_add subset_mset.dual_order.trans subset_mset.le_diff_conv2)
    qed
  qed
qed

lemma left_cuts_sum:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  shows "(a p (b p c)) m = (p  left_cuts m. a (fst p)  (b (snd p)  c (m - (fst p) - (snd p))))"
proof-
  have U: "(a p (b p c)) m = (xmset_factors m. (xamset_factors (m - x). a x  (b xa  c (m - x - xa))))"
    unfolding P_ring_mult_def
  proof-
    obtain f where f_def:
          "f = (λx . a x  (xamset_factors (m - x). b xa  c (m - x - xa)))"
      by simp
    obtain g where g_def:
      "g =  (λx . xamset_factors (m - x). a x  (b xa  c (m - x - xa)))"
      by simp
    have 0: "restrict f (mset_factors m) = restrict g (mset_factors m)"
    proof
      fix x
      show "restrict f (mset_factors m) x = restrict g (mset_factors m) x"
      proof(cases "x  mset_factors m")
        case True
         have T0: "restrict f (mset_factors m) x = a x  (xamset_factors (m - x). b xa  c (m - x - xa))"
           using f_def True by auto
         have T1: "restrict g (mset_factors m) x = (xamset_factors (m - x). a x  (b xa  c (m - x - xa)))"
           using True g_def by auto
         show "restrict f (mset_factors m) x = restrict g (mset_factors m) x"
           using assms finsum_rdistr[of "mset_factors (m - x)" "a x" "λ xa. b xa  c (m - x - xa)"]
           by (metis (mono_tags, lifting) mset_factors_finite Pi_I T0 T1 carrier_coeffE m_closed)
      next
        case False
        then have "restrict f (mset_factors m)  x = undefined" using f_def
          by (simp add: restrict_def)
        have  "restrict g (mset_factors m) x = undefined" using g_def False
          using restrict_def by auto
        then show ?thesis using f_def g_def
          using restrict f (mset_factors m) x = undefined
          by auto

      qed
    qed
    have 1: "f  mset_factors m  carrier R"
      using f_def assms
      by (simp add: carrier_coeffE)
    have 2: "g  mset_factors m  carrier R"
      using g_def assms
      by (simp add: carrier_coeffE)
    show "(xmset_factors m. a x  (xamset_factors (m - x). b xa  c (m - x - xa))) =
    (xmset_factors m. xamset_factors (m - x). a x  (b xa  c (m - x - xa)))"
      using f_def g_def finsum_eq[of "f"
                "mset_factors m" "g"] 0 1 2
      by blast
  qed
  have 0: "(s. s  mset_factors m  finite (mset_factors (m - s)))"
    by simp
  have 1: "finite (mset_factors m)"
    by simp
  have 2: "(s y. s  mset_factors m  y  mset_factors (m - s)  a s  (b y  c (m - s - y))  carrier R)"
    using assms assoc_aux0[of a b c ]
    by blast
    have "(xmset_factors m. (xamset_factors (m - x). a x  (b xa  c (m - x - xa)))) =
          (p  left_cuts m. a (fst p)  (b (snd p)  c (m - (fst p) - (snd p)))) "
    using assms left_cuts_alt_def[of m] 0 1 2
          double_finsum[of "mset_factors m" "λx. mset_factors (m - x)" "left_cuts m" "(λx y.  a x  (b y  c (m - x - y)))"]
    by blast
  then show ?thesis using U
    by auto
qed

lemma right_cuts_sum:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  shows "(a p b p c) m = (p  right_cuts m. a (snd  p)  (b ((fst p) -(snd p))  c (m - (fst p))))"
proof-
  have 0: "finite (mset_factors m)"
    by simp
  have 1: "(s. s  mset_factors m  finite (mset_factors s))"
    by auto
  have 2: "right_cuts m = {(s, t). s  mset_factors m  t  mset_factors s}"
    unfolding right_cuts_def
    by (simp add: monom_divides_factors)
  have 3: "(s y. s  mset_factors m  y  mset_factors s  a y  b (s - y)  c (m - s)  carrier R)"
    using assoc_aux1 assms(1) assms(2) assms(3)
    by blast
  have 4: "(smset_factors m. (ymset_factors s. a y  b (s - y)  c (m - s))) =
    (pright_cuts m. a (snd p)  b (fst p - snd p)  c (m - fst p))"
    using 0 1 2 3
      double_finsum[of "mset_factors m" _ "right_cuts m"
        "(λx y.  (a y   b (x - y))  c (m - x))"]
    by auto
  have 5: "(xmset_factors m. (xamset_factors x. a xa  b (x - xa))  c (m - x)) =
           (xmset_factors m. xamset_factors x. a xa  b (x - xa)  c (m - x))"
  proof-
    obtain f where f_def: "f =( λx. (xamset_factors x. a xa  b (x - xa))  c (m - x))"
      by simp
    obtain g where g_def: "g = (λx.  xamset_factors x. a xa  b (x - xa)  c (m - x))"
      by simp
    have 50: "s. s  (mset_factors m)  f s = g s"
    proof-
      fix x
      assume As:  "x  mset_factors m"
      show "f x = g x"
      proof-
        have f_eq: "f x = (xamset_factors x. a xa  b (x - xa))  c (m - x)"
          using f_def
          by auto
        have g_eq: "g x =  (xamset_factors x. a xa  b (x - xa)  c (m - x))"
          using g_def
          by auto
        have f_eq': "f x =  (xamset_factors x. (a xa  b (x - xa))  c (m - x))"
          using f_eq finsum_ldistr[of "mset_factors x" "c (m - x)" "λxa. (a xa  b (x - xa))" ] assms
          by (simp add: finite (mset_factors x); c (m - x)  carrier R; (λxa. a xa  b (x - xa))  mset_factors x  carrier R  (imset_factors x. a i  b (x - i))  c (m - x) = (imset_factors x. a i  b (x - i)  c (m - x)) Pi_I carrier_coeffE)
        then show "f x = g x"
          by (simp add: g_eq)
      qed
    qed
    have 51: "f  mset_factors m  carrier R"
    proof
      fix x
      assume "x  mset_factors m"
      have "(xamset_factors x. a xa  b (x - xa))  c (m - x ) carrier R"
        using assms carrier_coeffE finsum_closed[of "λxa.  a xa  b (x - xa)" "mset_factors x"]
        by blast
      then show "f x  carrier R" using assms f_def by auto
    qed
    have 52: "g  mset_factors m  carrier R"
    proof
      fix x
      assume "x  mset_factors m"
      show "g x  carrier R"
        using assms finsum_closed[of "λxa. a xa  b (x - xa)  c (m - x)" "mset_factors x"] g_def
        by (metis (no_types, lifting) "50" "51" Pi_iff x  mset_factors m)
    qed

    show ?thesis
      using 50 51 52 finsum_eq[of f "(mset_factors m) " g]
      by (metis (mono_tags, lifting) f_def finsum_cong' g_def)
  qed
  then have  6: "(xmset_factors m. (xamset_factors x. a xa  b (x - xa))  c (m - x)) =
    (pright_cuts m. a (snd p)  b (fst p - snd p)  c (m - fst p))"
    by (simp add: "4")
  have 7: "(pright_cuts m. a (snd p)  b (fst p - snd p)  c (m - fst p))
      = (pright_cuts m. a (snd p)  (b (fst p - snd p)  c (m - fst p)))"
    using assms
    by (meson local.ring_axioms m_assoc ring.carrier_coeff_def)
  then show ?thesis
    using assms 5 6
    unfolding P_ring_mult_def
    by simp
qed

text‹The Associativity of Polynomial Multiplication:›

lemma P_ring_mult_assoc:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  assumes "carrier_coeff c"
  shows "a p (b p c) = (a p b) p c"
proof
  fix m
  show "(a p (b p c)) m = (a p b p c) m"
  proof-
    obtain f where f_def: "f = (λp.  a (snd  p)  (b ((fst p) -(snd p))  c (m - (fst p))))"
      by simp
    obtain g where g_def: "g = (λp. a (fst p)  (b (snd p)  c (m - (fst p) - (snd p))))"
      by simp
    have f_dom: "f  right_cuts m  carrier R"
      using assms f_def unfolding right_cuts_def
      by (simp add: carrier_coeffE)
    have g_dom: "g  left_cuts m  carrier R"
      using assms g_def unfolding left_cuts_def
      by (simp add: carrier_coeffE)
    have 0: "finite (right_cuts m)"
      by (simp add: right_cuts_finite)
    have 1: "bij_betw (λ (x,y). (x + y, x))  (left_cuts m) (right_cuts m)"
      by (simp add: left_right_cuts_bij)
    have 2: "(s. s  left_cuts m  g s = f (case s of (x, y)  (x + y, x)))"
    proof-
      fix s
      assume "s  left_cuts m"
      then obtain x y where xy: "s = (x, y)  x ⊆# m  y ⊆# m - x"
        using  left_cuts_def
        by blast
      then have g_eq: "g s = a x  (b y  c (m - x - y))"
        using g_def fst_conv
        by auto
      have f_eq: "f (case s of (x, y)  (x + y, x)) = f (x + y, x)"
        by (simp add: xy)
      then have f_eq': "f (case s of (x, y)  (x + y, x)) = a x  (b ((x + y) - x)  c (m - (x + y)))"
        using f_def
        by simp
      then show "g s = f (case s of (x, y)  (x + y, x))"
        by (simp add: g_eq)
    qed
    have 3: "finsum R g (left_cuts m) = finsum R f (right_cuts m)"
    using 0 1 2 finsum_bij_eq[of g "left_cuts m" f "right_cuts m" "λ (x,y). (x + y, x)" ]
      using bij_betw_finite f_dom g_dom by blast
    then show ?thesis using assms right_cuts_sum left_cuts_sum
      by (metis (mono_tags, lifting) f_def f_dom finsum_cong' g_def g_dom)
  qed
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Commutativity of Polynomial Multiplication›
(**************************************************************************************************)
(**************************************************************************************************)

context ring
begin

lemma mset_factors_bij:
"bij_betw (λx. m - x) (mset_factors m) (mset_factors m)"
  apply(rule bij_betwI')
  apply (metis monom_comp' monom_divides_factors)
   apply (simp add: monom_divides_factors)
    by (meson monom_comp' monom_divides_factors diff_subset_eq_self)

lemma(in cring) P_ring_mult_comm:
  assumes "carrier_coeff a"
  assumes "carrier_coeff b"
  shows "a p b  = b p a"
proof
  fix m
  show "(a p b) m = (b p a) m"
  unfolding P_ring_mult_def
  apply (rule finsum_bij_eq[of "λ x. a x  b (m - x)" "mset_factors m"
                          "λx. b x  a (m - x)" "mset_factors m" "λx. m - x"])
  proof
    show "x. x  mset_factors m  a x  b (m - x)  carrier R"
    proof-
      fix x
      assume "x  mset_factors m"
      show "a x  b (m - x)  carrier R"
        using assms carrier_coeffE
        by blast
    qed
    show "(λx. b x  a (m - x))  mset_factors m  carrier R"
    proof
      fix x
      assume "x  mset_factors m"
      show "b x  a (m - x)  carrier R"
        using assms  carrier_coeffE
        by blast
    qed
    show "finite (mset_factors m)"
      by simp
    show "bij_betw ((-) m) (mset_factors m) (mset_factors m)"
      by (simp add: mset_factors_bij)
    show "s. s  mset_factors m  a s  b (m - s) = b (m - s)  a (m - (m - s))"
    proof-
      fix s
      assume "s  mset_factors m"
      then show "a s  b (m - s) = b (m - s)  a (m - (m - s))"
        using assms carrier_coeffE
        by (metis local.ring_axioms m_comm ring.monom_comp' ring.monom_divides_factors)
    qed
  qed
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Closure properties for multiplication›
(**************************************************************************************************)
(**************************************************************************************************)

text‹Building monomials from polynomials:›

lemma indexed_const_P_mult_eq:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "(indexed_const a) p  (indexed_const b) = indexed_const (a  b)"
proof-
  have 0: "monomials_of R (indexed_const a) = (if (a = 𝟬) then {} else {{#}})"
    unfolding monomials_of_def indexed_const_def
    by auto
  have 1: "monomials_of R (indexed_const b) = (if (b = 𝟬) then {} else {{#}})"
    unfolding monomials_of_def indexed_const_def
    by auto
  show ?thesis
    unfolding P_ring_mult_def
  proof
    fix m:: "'c monomial"
    show "(xmset_factors m. indexed_const a x  indexed_const b (m - x)) = indexed_const (a  b) m "
    proof(cases "m = {#}")
      case True
      then have T0: "mset_factors m = {{#}}"
        unfolding mset_factors_def
        by simp
      have "(xmset_factors m. indexed_const a x  indexed_const b (m - x)) =
                   indexed_const a {#}  indexed_const b ({#} - {#}) "
      proof-
        have 0: "(λx. indexed_const a x  indexed_const b (m - x))  {}  carrier R"
          by blast
        have 1: "indexed_const a {#}  indexed_const b (m - {#})  carrier R"
          by (simp add: assms(1) assms(2) indexed_const_def)
        have 2: "(x{{#}}. indexed_const a x  indexed_const b (m - x)) =
                 indexed_const a {#}  indexed_const b (m - {#})  (x{}. indexed_const a x  indexed_const b (m - x))"
        using True T0 0 1  finsum_insert[of "{}" "{#}" "λx. indexed_const a x  indexed_const b (m - x)"  ]
            by (simp add: indexed_const_def)
        then show ?thesis
          using True T0 0 1  finsum_insert[of "{}" "{#}" "λx. indexed_const a x  indexed_const b (m - x)"  ]
              finsum_empty[of "λx. indexed_const a x  indexed_const b (m - x)"]
          by (simp add: finite {}; {#}  {}; (λx. indexed_const a x  indexed_const b (m - x))  {}  carrier R; indexed_const a {#}  indexed_const b (m - {#})  carrier R  (x{{#}}. indexed_const a x  indexed_const b (m - x)) = indexed_const a {#}  indexed_const b (m - {#})  (x{}. indexed_const a x  indexed_const b (m - x)) indexed_const_def)
      qed
      then show ?thesis
        by (simp add: True indexed_const_def)
    next
      case False
      then show ?thesis using 0 1
        by (simp add: add.finprod_one_eqI assms(1) assms(2)
            indexed_const_def )
    qed
  qed
qed

lemma indexed_const_P_multE:
  assumes "P  indexed_pset I (carrier R)"
  assumes "c  carrier R"
  shows "(P p (indexed_const c)) m =  (P m)  c"
  unfolding P_ring_mult_def
proof-
  have 3: "m  mset_factors m - {m}"
    by simp
  have 4: "finite ((mset_factors m) - {m})"
    by simp
  have 5: "P m  indexed_const c (m - m)  carrier R "
  by (metis assms(1) assms(2) cancel_comm_monoid_add_class.diff_cancel
      carrier_coeffE indexed_const_def indexed_pset_in_carrier local.ring_axioms
      ring.ring_simprules(5) subsetI)
  have 0: "(λx. P x  indexed_const c (m - x))  mset_factors m - {m}  carrier R"
  proof
    fix x
    assume "x  mset_factors m - {m}"
    then show "P x  indexed_const c (m - x)  carrier R "
      using assms
      by (meson carrier_coeffE indexed_const_in_carrier indexed_pset_in_carrier m_closed subsetI)
  qed
  have 1: "P m  indexed_const c ( m - m) =  (P m  c)"
    by (simp add: indexed_const_def)
  have 2: "x. x  mset_factors m   P x  indexed_const c (m - x)  = (if x = m then (P m  c) else 𝟬)"
  proof-
    fix x
    assume "x  mset_factors m "
    then have "indexed_const c (m - x) = (if x = m then c else 𝟬)"
      unfolding indexed_const_def
      by (metis cancel_comm_monoid_add_class.diff_cancel
          diff_zero monom_comp' monom_divides_factors)
    then show  "P x  indexed_const c (m - x)  = (if x = m then (P m  c) else 𝟬)"
      by (metis assms(1) carrier_coeffE indexed_pset_in_carrier
            local.semiring_axioms semiring.r_null set_eq_subset)
  qed
  then have "(x(mset_factors m) - {m}. P x  indexed_const c (m - x)) = 𝟬"
    using assms
    by (metis (no_types, lifting) DiffD1 DiffD2 add.finprod_one_eqI singletonI)
  then show "(x(mset_factors m). P x  indexed_const c (m - x)) = P m  c"
    using assms 0 1 3 4 5 finsum_insert[of "(mset_factors m) - {m}"
                                    m "λx. P x  indexed_const c (m - x) "]
    by (metis (no_types, lifting) m_factor add.l_cancel_one
              insert_Diff_single insert_absorb zero_closed)
qed

lemma indexed_const_var_mult:
  assumes "P  indexed_pset I (carrier R)"
  assumes "c  carrier R"
  assumes "i  I"
  shows "P  i p indexed_const c = (P p (indexed_const c))  i "
proof
  fix m
  show "(P  i p indexed_const c) m = (P p indexed_const c  i) m"
  proof(cases "i ∈# m")
    case True
    then have T0: "(P  i p indexed_const c) m  = (P  i) m  c"
      using assms indexed_const_P_multE[of "P  i" I c m]
      by (simp add: indexed_pset.indexed_pmult)
    then show ?thesis using assms indexed_const_P_multE[of P I c m]
      unfolding  indexed_pmult_def
      using True indexed_const_P_multE by fastforce
  next
    case False
        then have T0: "(P  i p indexed_const c) m  = (P  i) m  c"
      using assms indexed_const_P_multE[of "P  i" I c m]
      by (simp add: indexed_pset.indexed_pmult)
    then show ?thesis using assms indexed_const_P_multE[of P I c m]
      unfolding  indexed_pmult_def
      using False by auto
  qed
qed

lemma indexed_const_P_mult_closed:
  assumes "a  indexed_pset I (carrier R)"
  assumes "c  carrier R"
  shows "a p (indexed_const c)  indexed_pset I (carrier R)"
  apply(rule indexed_pset.induct[of a I "(carrier R)" ])
     apply (simp add: assms(1); fail)
proof-
  show "k. (k  carrier R)  ((indexed_const k) p (indexed_const c))  ((carrier R) [𝒳I])"
    using assms
    by (metis indexed_const_P_mult_eq indexed_pset.simps m_closed)
  show "P Q. P  ((carrier R) [𝒳I]) 
           P p indexed_const c  ((carrier R) [𝒳I]) 
           Q  ((carrier R) [𝒳I])  Q p indexed_const c  ((carrier R) [𝒳I])  P  Q p indexed_const c  ((carrier R) [𝒳I])"
    using P_ring_ldistr
    by (metis assms(2) carrier_coeff_def indexed_const_in_carrier indexed_pset.indexed_padd indexed_pset_in_carrier subsetI)
  show "P i. P  ((carrier R) [𝒳I])  P p indexed_const c  ((carrier R) [𝒳I])  i  I  P  i p indexed_const c  ((carrier R) [𝒳I])"
  proof-
    fix P i
    assume A0: "P  ((carrier R) [𝒳I])"
    assume A1: " (P p indexed_const c)  ((carrier R) [𝒳I])"
    assume A2: "i  I"
    show "P  i p indexed_const c  (carrier R [𝒳I])"
      using assms  A0 A1 A2 indexed_const_var_mult local.ring_axioms ring.indexed_pset.simps
      by (metis (no_types, opaque_lifting))
  qed
qed

lemma monom_add_mset:
"mset_to_IP R (add_mset i m) = mset_to_IP R m  i"
  unfolding indexed_pmult_def
  by (metis (no_types, opaque_lifting) add_mset_diff_bothsides diff_empty mset_to_IP_def multi_member_split union_single_eq_member)

text‹Monomials are closed under multiplication:›

lemma monom_mult:
"mset_to_IP R m p mset_to_IP R n = mset_to_IP R (m + n)"
proof-
  have "(λma. xmset_factors ma. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬))
       = (λma. (if ma = n + m then 𝟭 else 𝟬))"
  proof
    fix ma
    show "(xmset_factors ma. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)) = (if ma = n + m then 𝟭 else 𝟬)"

    proof-
      have 0: "x. x  mset_factors ma 
            (λ x. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)) x =
                  (if (x = m) then (if ma - x = n then 𝟭 else 𝟬) else 𝟬) "
        by simp
      then have 1: "(x((mset_factors ma) - {m}). (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)) =
            𝟬"
        by (metis (no_types, lifting) add.finprod_one_eqI mem_Collect_eq set_diff_eq singletonI)
      have 2: "finite (mset_factors ma - {m}) "
        by simp
      have 3: "m  mset_factors ma - {m}"
        by simp
      have 4: "(λx. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬))  mset_factors ma - {m}  carrier R"
      proof
        fix x
        assume " x  mset_factors ma - {m}"
        show "(if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)  carrier R "
          by auto
      qed
      show ?thesis
      proof(cases "m  (mset_factors ma)")
        case True
        have T0: " (xinsert m (mset_factors ma - {m}). (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)) =
      (if m = m then 𝟭 else 𝟬)  (if ma - m = n then 𝟭 else 𝟬) 
      (xmset_factors ma - {m}. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬))"
          using True 1 2 3 4 finsum_insert[of "((mset_factors ma) - {m})" m
                                   "λx. (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬)"]
          using m_closed one_closed zero_closed by presburger
        have T1: "(mset_factors ma) = insert m (mset_factors ma - {m})"
          using True
          by blast
        then have  "(x(mset_factors ma). (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬))
                      = (if m = m then 𝟭 else 𝟬)  (if ma - m = n then 𝟭 else 𝟬)"
          using T0 T1 1 add.l_cancel_one insert_Diff_single insert_absorb2 l_one mk_disjoint_insert one_closed zero_closed
          by presburger
        then have "(x(mset_factors ma). (if x = m then 𝟭 else 𝟬)  (if ma - x = n then 𝟭 else 𝟬))
                      = (if ma - m = n then 𝟭 else 𝟬)"
          by simp
        then show ?thesis
          by (metis (no_types, lifting) monom_divides_factors True add.commute add_diff_cancel_right' subset_mset.add_diff_inverse)
      next
        case False
        then show ?thesis
          by (metis (no_types, lifting) "0" monom_divides_factors add.finprod_one_eqI mset_subset_eq_add_right)
      qed
    qed
  qed
  then show ?thesis
  unfolding mset_to_IP_def P_ring_mult_def
  by (simp add: union_commute)
qed

lemma poly_index_mult:
  assumes "a  indexed_pset I (carrier R)"
  assumes "i  I"
  shows "a  i = a p mset_to_IP R {#i#}"
proof
  fix m
  show "(a  i) m = (a p mset_to_IP R {#i#}) m"
  proof(cases "i ∈# m")
    case True
    have T0: "(a  i) m = a (m - {#i#})"
      by (simp add: True indexed_pmult_def)
    have T1: "(a p mset_to_IP R {#i#}) m = a (m - {#i#})"
    proof-
      have T10: "(a p mset_to_IP R {#i#}) m
                = (xmset_factors m. a x  mset_to_IP R {#i#} (m - x)) "
        unfolding P_ring_mult_def
        by auto
      have T11: "x. x  mset_factors m 
                  mset_to_IP R {#i#} (m - x) = (if x = (m - {#i#}) then 𝟭 else 𝟬)"
        unfolding mset_to_IP_def
        by (metis Multiset.diff_cancel Multiset.diff_right_commute True diff_single_eq_union
            diff_single_trivial diff_zero monom_comp' monom_divides_factors multi_drop_mem_not_eq)
      have T12: "x. x  mset_factors m 
              a x  mset_to_IP R {#i#} (m - x) = (if  x = (m - {#i#}) then a (m - {#i#}) else 𝟬)"
      proof-
        fix x
        assume "x  mset_factors m"
        then show " a x  mset_to_IP R {#i#} (m - x) = (if x = m - {#i#} then a (m - {#i#}) else 𝟬)"
          apply(cases "x = m - {#i#}")
           apply (metis T0 T11 x  mset_factors m assms(1) carrier_coeffE empty_subsetI
              ideal_is_subalgebra indexed_pset_in_carrier oneideal r_one subalgebra_in_carrier)
          by (metis T11 x  mset_factors m assms(1) carrier_coeffE carrier_is_subalgebra
              empty_subsetI  indexed_pset_in_carrier  r_null  subalgebra_in_carrier)
      qed
      have  T13: "(x(mset_factors m) - {m - {#i#}}. a x  mset_to_IP R {#i#} (m - x)) = 𝟬"
        using T12
        by (metis (no_types, lifting) DiffE add.finprod_one_eqI singletonI)
      have T14: "finite (mset_factors m - {m - {#i#}})"
        using mset_factors_finite by blast
      have T15: "m - {#i#}  mset_factors m - {m - {#i#}}"
        by simp
      have T16: "  (λx. a x  mset_to_IP R {#i#} (m - x))  mset_factors m - {m - {#i#}}  carrier R"
      proof
        fix x
        assume "x  mset_factors m - {m - {#i#}}"
        then show " a x  mset_to_IP R {#i#} (m - x)  carrier R"
          using assms T12 by auto
      qed
      have "m - (m - {#i#}) = {#i#}"
        by (metis monom_comp' True single_subset_iff)
      then have T17: " a (m - {#i#})  mset_to_IP R {#i#} (m - (m - {#i#}))  carrier R"
        unfolding mset_to_IP_def apply auto using assms
        by (meson carrier_coeffE carrier_is_subalgebra empty_subsetI indexed_pset_in_carrier m_closed one_closed subalgebra_in_carrier)
      have T18:"(xinsert (m - {#i#}) (mset_factors m - {m - {#i#}}). a x  mset_to_IP R {#i#} (m - x)) =
      a (m - {#i#})  mset_to_IP R {#i#} (m - (m - {#i#}))  (xmset_factors m - {m - {#i#}}. a x  mset_to_IP R {#i#} (m - x))"
        using T12 T13 T14 T15 T16 T17 unfolding P_ring_mult_def
        using finsum_insert[of "mset_factors m - {m - {#i#}}" "m - {#i#}"
            "λx. a x  mset_to_IP R {#i#} (m - x)"]
        by blast
      have T19: "a (m - {#i#})  mset_to_IP R {#i#} (m - (m - {#i#})) = a (m - {#i#}) "
      proof-
        have " (m - (m - {#i#})) = {#i#}"
          using True
          by (metis monom_comp' single_subset_iff)
        then have "mset_to_IP R {#i#} (m - (m - {#i#})) = 𝟭"
          by (metis mset_to_IP_def)
        have "a (m - {#i#})  carrier R"
          using assms
          by (meson carrier_coeffE carrier_is_subalgebra exp_base_closed
              indexed_pset_in_carrier one_closed subalgebra_in_carrier)
        then show ?thesis using assms
          using mset_to_IP R {#i#} (m - (m - {#i#})) = 𝟭 by auto
      qed
      have T20: "(xinsert (m - {#i#}) ((mset_factors m) - {m - {#i#}}). a x  mset_to_IP R {#i#} (m - x)) =
       a (m - {#i#})  (xmset_factors m - {m - {#i#}}. a x  mset_to_IP R {#i#} (m - x))"
        using T18 T19 by auto
      have T21: "insert (m - {#i#}) ((mset_factors m) - {m - {#i#}}) = mset_factors m"
      proof-
        have "(m - {#i#})  mset_factors m"
          by (simp add: monom_divides_factors)
        then show ?thesis
          by blast
      qed
      then show ?thesis using T13 T20 True unfolding P_ring_mult_def
        using T17 T19 by auto
    qed
    then show ?thesis
      by (simp add: T0)
  next
    case False
    have F0: "(a  i) m = 𝟬 "
      by (simp add: False indexed_pmult_def)
    have F1: "  (a p mset_to_IP R {#i#}) m = 𝟬"
      unfolding P_ring_mult_def
    proof-
      have "x. x  mset_factors m   a x  mset_to_IP R {#i#} (m - x) = 𝟬"
      proof-
        fix x
        assume A: "x  mset_factors m"
        have B: "m - x  {#i#}" using False A
          by (metis diff_subset_eq_self single_subset_iff)
        then show "a x  mset_to_IP R {#i#} (m - x) = 𝟬"
          using assms False unfolding mset_to_IP_def
          using  carrier_coeffE indexed_pset_in_carrier by fastforce
      qed
      then show "(xmset_factors m. a x  mset_to_IP R {#i#} (m - x)) = 𝟬"
        by (simp add: add.finprod_one_eqI)
    qed
    then show ?thesis
      by (simp add: F0)
  qed
qed

lemma mset_to_IP_mult_closed:
  assumes "a  indexed_pset I (carrier R)"
  shows "set_mset m  I  a p mset_to_IP R m  indexed_pset I (carrier R)"
proof(induction m)
  case empty
  then have "mset_to_IP R {#} = indexed_const 𝟭"
    unfolding mset_to_IP_def indexed_const_def by auto
  then show ?case
    by (simp add: mset_to_IP R {#} = indexed_const 𝟭 assms indexed_const_P_mult_closed)
next
  case (add x m)
  fix x
  fix m :: "'c monomial"
  assume A: "set_mset (add_mset x m)  I"
  assume B: "set_mset m  I  a p mset_to_IP R m  (carrier R [𝒳I])"
  show " a p mset_to_IP R (add_mset x m)  (carrier R [𝒳I])"
  proof-
    have 0: "x  I"
      using A by auto
    have 1: "set_mset m  I"
      using A by auto
    have 2: " a p mset_to_IP R m  (carrier R [𝒳I])"
      using "1" B by blast
    have 3: " a p mset_to_IP R (add_mset x m) = (a p mset_to_IP R m)  x"
    proof-
      have 30: " a p mset_to_IP R (add_mset x m) = a p (mset_to_IP R m  x)"
        using monom_add_mset[of x m]
        by auto
      have 31: "(a p mset_to_IP R m)  x = a p (mset_to_IP R m  x)"
      proof
        fix y
        show "(a p mset_to_IP R m  x) y = (a p (mset_to_IP R m  x)) y "
        proof-
          have 310: "(a p mset_to_IP R m  x) y = (a p mset_to_IP R m p mset_to_IP R {#x#}) y"
            using poly_index_mult "0" "2"
            by fastforce
          have 311: " (mset_to_IP R m  x) = mset_to_IP R m p mset_to_IP R {#x#}"
            using  "0" "2"
            by (metis add_mset_add_single monom_add_mset monom_mult)
          have 312: "carrier_coeff a "
            using assms indexed_pset_in_carrier by blast
          have 313: "carrier_coeff (mset_to_IP R m)"
            by (simp add: carrier_coeff_def mset_to_IP_def)
          have 314: "carrier_coeff (mset_to_IP R {#x#})"
            by (metis carrier_coeff_def mset_to_IP_def one_closed zero_closed)
          show ?thesis
            using 310 311 312 313 314
                  P_ring_mult_assoc[of a "mset_to_IP R m" "mset_to_IP R {#x#}"]
            by simp
        qed
      qed
      then show ?thesis
        by (simp add: monom_add_mset)
    qed
    then show ?thesis
      by (simp add: "0" "1" B indexed_pset.indexed_pmult)
  qed
qed

text‹
  A predicate which identifies when the variables used in a given polynomial $P$ are only
   drawn from a fixed variable set $I$.
›
abbreviation monoms_in where
"monoms_in I P  (m  monomials_of R P. set_mset m  I)"

lemma monoms_of_const:
"monomials_of R (indexed_const k) = (if k = 𝟬 then {} else {{#}})"
  unfolding indexed_const_def monomials_of_def
  by auto

lemma const_monoms_in:
"monoms_in I (indexed_const k)"
  unfolding indexed_const_def monomials_of_def
  using  count_empty count_eq_zero_iff monomials_ofE subsetI
  by simp

lemma mset_to_IP_indices:
  shows  "P  indexed_pset I (carrier R)  monoms_in I P"
  apply(erule indexed_pset.induct[of])
   apply (simp add: const_monoms_in; fail)
proof-
  show "P Q. P  (carrier R [𝒳I]) 
           mmonomials_of R P. set_mset m  I 
           Q  (carrier R [𝒳I])  mmonomials_of R Q. set_mset m  I  mmonomials_of R (P  Q). set_mset m  I"
  proof
    fix P Q
    fix m
    assume A: "P  (carrier R [𝒳I])" "mmonomials_of R P. set_mset m  I"
                "Q  (carrier R [𝒳I])" "mmonomials_of R Q. set_mset m  I"
                  " m  monomials_of R (P  Q)"
    show "set_mset m  I"
    proof-
      have "m  monomials_of R P  m  monomials_of R Q"
        using A using monomials_of_add[of P Q]
        by blast
      then show ?thesis using A
        by blast
    qed
  qed
  show "P i. P  (carrier R [𝒳I])  mmonomials_of R P. set_mset m  I  i  I  mmonomials_of R (P  i). set_mset m  I"
  proof
    fix P
    fix i
    fix m
    assume A: "P  (carrier R [𝒳I])" "mmonomials_of R P. set_mset m  I" "i  I" "m  monomials_of R (P  i)"
    obtain n where "n  monomials_of R P  m = add_mset i n"
      using A
      by (metis image_iff monomials_of_p_mult')
    then show "set_mset m  I"
      using A
      by auto
  qed
qed

lemma mset_to_IP_indices':
  assumes "P  indexed_pset I (carrier R)"
  assumes "m  monomials_of R P"
  shows "set_mset m  I"
  using assms(1) assms(2) mset_to_IP_indices by blast

lemma one_mset_to_IP:
 "mset_to_IP R {#} = indexed_const 𝟭"
  unfolding mset_to_IP_def indexed_const_def
  by blast

lemma mset_to_IP_closed:
  shows "set_mset m  I mset_to_IP R m  indexed_pset I (carrier R) "
  apply(induction m)
  apply (simp add: indexed_pset.indexed_const one_mset_to_IP)
  by (metis (no_types, lifting) add_mset_commute indexed_pset.simps
      monom_add_mset mset_add subset_iff union_single_eq_member)

lemma mset_to_IP_closed':
  assumes "P  indexed_pset I (carrier R)"
  assumes "m  monomials_of R P"
  shows "mset_to_IP R m  indexed_pset I (carrier R)"
  by (meson assms(1) assms(2) mset_to_IP_closed mset_to_IP_indices')

text‹This lemma expresses closure under multiplcation for indexed polynomials.›

lemma  P_ring_mult_closed:
  assumes "carrier_coeff P"
  assumes "carrier_coeff Q"
  shows  "carrier_coeff (P_ring_mult R P Q)"
  unfolding carrier_coeff_def
proof
  fix m
  have "(λx. P x  Q (m - x))  mset_factors m  carrier R"
  proof
    fix x
    assume "x  mset_factors m"
    then show "P x  Q (m - x)  carrier R"
      using assms  carrier_coeffE
      by blast
  qed
  then show "(P p Q) m  carrier R"
    using assms finsum_closed[of "(λx. (P x)  (Q (m - x)))" "mset_factors m"]
    unfolding carrier_coeff_def  P_ring_mult_def
    by blast
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Multivariable Polynomial Induction›
(**************************************************************************************************)
(**************************************************************************************************)

lemma mpoly_induct:
  assumes "Q. Q  indexed_pset I (carrier R)  card (monomials_of R Q) = 0  P Q"
  assumes "n. (Q. Q  indexed_pset I (carrier R)  card (monomials_of R Q)   n  P Q)
              (Q. Q  indexed_pset I (carrier R)  card (monomials_of R Q)  (Suc n)  P Q)"
  assumes "Q  indexed_pset I (carrier R)"
  shows "P Q"
proof-
  have "m. Q. Q  indexed_pset I (carrier R)  card(monomials_of R Q)  m  P Q"
  proof-
    fix m
    show "Q. Q  indexed_pset I (carrier R)  card(monomials_of R Q)  m  P Q"
      apply(induction m)
      using assms(1)
       apply blast
      using assms
      by blast
  qed
  then show ?thesis
    using assms(3) by blast
qed

lemma monomials_of_card_zero:
  assumes "Q  indexed_pset I (carrier R)  card (monomials_of R Q) = 0"
  shows "Q = indexed_const 𝟬"
proof-
  have 0: "carrier_coeff  Q"
    using assms  indexed_pset_in_carrier
    by blast
  have "m. Q m = 𝟬"
    unfolding monomials_of_def
    using assms monomials_finite
    by (metis card_0_eq complement_of_monomials_of empty_iff)
  then show  ?thesis
  using 0 assms unfolding indexed_const_def
  by auto
qed

text‹Polynomial induction on the number of monomials with nonzero coefficient:›

lemma mpoly_induct':
  assumes "P (indexed_const 𝟬)"
  assumes "n. (Q. Q  indexed_pset I (carrier R)  card (monomials_of R Q)   n  P Q)
              (Q. Q  indexed_pset I (carrier R)  card (monomials_of R Q) = (Suc n)  P Q)"
  assumes "Q  indexed_pset I (carrier R)"
  shows "P Q"
  apply(rule mpoly_induct)
  using assms(1) monomials_of_card_zero apply blast
proof-
  show "n Q. (Q. Q  (carrier R [𝒳I])  card (monomials_of R Q)  n  P Q)  Q  (carrier R [𝒳I])  card (monomials_of R Q)  Suc n  P Q"
  proof-
    fix n
    fix Q
    assume A0: "(Q. Q  (carrier R [𝒳I])  card (monomials_of R Q)  n  P Q)"
    assume A1: "Q  (carrier R [𝒳I])  card (monomials_of R Q)  Suc n"
    show "P Q"
      apply(cases  "card (monomials_of R Q) = Suc n")
      using assms A0 A1
       apply blast
      using assms A0 A1
      using le_SucE by blast
  qed
  show "Q  (carrier R [𝒳I])"
    using assms by auto
qed

lemma monomial_poly_split:
  assumes "P  indexed_pset I (carrier R)"
  assumes "m  monomials_of R P"
  shows "(restrict_poly_to_monom_set R P ((monomials_of R P) - {m}))  ((mset_to_IP R m) p (indexed_const (P m))) = P"
proof fix x
  show "(restrict_poly_to_monom_set R P (monomials_of R P - {m})  (mset_to_IP R m p indexed_const (P m))) x = P x"
  proof(cases "x = m")
    case True
    have T0: "(restrict_poly_to_monom_set R P (monomials_of R P - {m})) x = 𝟬"
      unfolding restrict_poly_to_monom_set_def
      by (simp add: True)
    have T1: "(mset_to_IP R m p indexed_const (P m)) x = P x"
      using assms True indexed_const_P_multE[of "mset_to_IP R m" I "P m" m]
              mset_to_IP_simp
      by (metis Idl_subset_ideal' carrier_coeffE genideal_one indexed_pset_in_carrier
          l_one mset_to_IP_closed' one_closed)
    then show ?thesis
      using T0 True
      unfolding indexed_padd_def
      using assms(1) carrier_coeffE indexed_pset_in_carrier l_zero
      by fastforce
  next
    case False
    then have F0: "(restrict_poly_to_monom_set R P (monomials_of R P - {m})) x = P x"
    proof(cases "x  monomials_of R P")
      case True
      then show ?thesis
        by (simp add: False restrict_poly_to_monom_set_def)
    next
      case False
      then show ?thesis
        by (simp add: complement_of_monomials_of restrict_poly_to_monom_set_def)
    qed
    have F1: "mset_to_IP R m  (carrier R [𝒳I])"
      using assms(1) assms(2) mset_to_IP_closed' by blast
    have F2: "P m  carrier R"
      using assms(1) carrier_coeffE indexed_pset_in_carrier by blast
    have F3: "(mset_to_IP R m p indexed_const (P m)) x = 𝟬"
      using False F1 F2 assms indexed_const_P_multE[of "mset_to_IP R m" I "P m" x]
      by (simp add: F1 m  monomials_of R P)
    then show ?thesis using F0 unfolding indexed_padd_def
      using assms(1) carrier_coeffE indexed_pset_in_carrier
      by fastforce
  qed
qed

lemma restrict_not_in_monoms:
  assumes "a  monomials_of R P"
  shows "restrict_poly_to_monom_set R P A = restrict_poly_to_monom_set R P (insert a A)"
proof
  fix x
  show " restrict_poly_to_monom_set R P A x = restrict_poly_to_monom_set R P (insert a A) x "
    unfolding restrict_poly_to_monom_set_def using assms unfolding monomials_of_def
  by simp
qed

lemma restriction_closed':
  assumes "P  indexed_pset I (carrier R)"
  assumes "finite ms"
  shows "(restrict_poly_to_monom_set R P ms)  indexed_pset I (carrier R)"
  apply(rule finite.induct[of ms])
  apply (simp add: assms(2); fail)
proof-
  show "restrict_poly_to_monom_set R P {}  (carrier R [𝒳I])"
  proof-
    have "restrict_poly_to_monom_set R P {} = indexed_const 𝟬"
      unfolding restrict_poly_to_monom_set_def indexed_const_def by auto
    then show ?thesis
      by (simp add: indexed_pset.indexed_const)
  qed
  show "A a. finite A  restrict_poly_to_monom_set R P A  (carrier R [𝒳I])  restrict_poly_to_monom_set R P (insert a A)  (carrier R [𝒳I])"
  proof-
    fix A
    fix a
    assume A: "restrict_poly_to_monom_set R P A  (carrier R [𝒳I])"
    show "restrict_poly_to_monom_set R P (insert a A)  (carrier R [𝒳I])"
    proof(cases "a  monomials_of R P")
      case True
      then have T0: "mset_to_IP R a  (carrier R [𝒳I])"
        using assms(1) mset_to_IP_closed' by blast
      then have T1: "(mset_to_IP R a p indexed_const (P a))   (carrier R [𝒳I])"
        by (meson assms(1) carrier_coeffE indexed_pset_in_carrier
      local.ring_axioms ring.indexed_const_P_mult_closed subset_refl)

      show ?thesis
      proof(cases "a  A")
        case True
        then show ?thesis
          by (simp add: A insert_absorb)
      next
        case False
        have "restrict_poly_to_monom_set R P (insert a A) = restrict_poly_to_monom_set R P A   (mset_to_IP R a p indexed_const (P a))"
        proof
          fix x
          show "restrict_poly_to_monom_set R P (insert a A) x = (restrict_poly_to_monom_set R P A  (mset_to_IP R a p indexed_const (P a))) x"
          proof(cases "x = a")
            case True
            then have FT0: "(if x  insert a A then P x else 𝟬) = P x"
              by simp
            have FT1: "(if x  A then P x else 𝟬) = 𝟬"
              by (simp add: False True)
            have FT2: "P a  carrier R"
              using assms(1) carrier_coeffE indexed_pset_in_carrier by blast
            have FT3: "(mset_to_IP R a p indexed_const (P a)) x = P x"
              using T0 FT2 True indexed_const_P_multE[of "mset_to_IP R a" I "P a" x] mset_to_IP_simp[of a]
              by simp
            then show ?thesis
              using False
              unfolding restrict_poly_to_monom_set_def indexed_padd_def
              using FT0 FT1 FT3  assms
            by (simp add: FT2 True)
          next
            case F: False
            then have FT0: "(if x  insert a A then P x else 𝟬) = (if x  A then P x else 𝟬)"
              by simp
            have FT2: "P a  carrier R"
              using assms(1) carrier_coeffE indexed_pset_in_carrier by blast
            have FT3: "(mset_to_IP R a p indexed_const (P a)) x = 𝟬"
              using T0 FT2 True indexed_const_P_multE[of "mset_to_IP R a" I "P a" x] mset_to_IP_simp[of a]
              by (simp add: F mset_to_IP_def)
            show ?thesis
              unfolding restrict_poly_to_monom_set_def indexed_padd_def
            proof-

              show "(if x  insert a A then P x else 𝟬) = (if x  A then P x else 𝟬)  (mset_to_IP R a p indexed_const (P a)) x"
                apply(cases "x  A")
                using FT0 FT2 FT3 F False assms  carrier_coeffE indexed_pset_in_carrier local.ring_axioms
                apply fastforce
                    using FT0 FT2 FT3 F False assms
                    by simp
                qed
          qed
        qed
        then show ?thesis
          using A T1 indexed_pset.simps by blast
      qed
    next
      case False
      then show ?thesis
        using A restrict_not_in_monoms by fastforce
    qed
  qed
qed

lemma restriction_restrict:
"restrict_poly_to_monom_set R P ms = restrict_poly_to_monom_set R P (ms  monomials_of R P)"
proof
  fix x
  show "restrict_poly_to_monom_set R P ms x = restrict_poly_to_monom_set R P (ms  monomials_of R P) x"
    unfolding restrict_poly_to_monom_set_def monomials_of_def
  by simp
qed

lemma restriction_closed:
  assumes "P  indexed_pset I (carrier R)"
  assumes "Q = restrict_poly_to_monom_set R P ms"
  shows "Q  indexed_pset I (carrier R)"
proof-
  have "Q = restrict_poly_to_monom_set R P (ms  monomials_of R P)"
    using assms restriction_restrict
    by blast
  then show ?thesis
    using assms restriction_closed'[of P I "(ms  monomials_of R P)"]
    using monomials_finite
    by blast
qed

lemma monomial_split_card:
  assumes "P  indexed_pset I (carrier R)"
  assumes "m  monomials_of R P"
  shows "card (monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m})))=
        card (monomials_of R P) -1"
proof-
  have 0: "(monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m}))) =
        (monomials_of R P) - {m}"
    using assms(1)
    by (meson Diff_subset restrict_poly_to_monom_set_monoms)
  then have 1:  "card (monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m}))) =
              card ((monomials_of R P) - {m})"
    by auto
  have " card ((monomials_of R P) - {m}) =  card (monomials_of R P) - 1"
    using assms(2)
  using assms(1) monomials_finite by fastforce
    then show ?thesis
    by (simp add: "1")
qed

lemma P_ring_mult_closed':
  assumes "a  indexed_pset I (carrier R)"
  assumes "b  indexed_pset I (carrier R)"
  shows "a p b  indexed_pset I (carrier R)"
  apply(rule mpoly_induct'[of "λb. a p b  indexed_pset I (carrier R)" I b])
  using assms(1) indexed_const_P_mult_closed apply blast
proof-
  show "b  (carrier R [𝒳I])"
    using assms by auto
  show "n Q. (Q. Q   (carrier R [𝒳I])  card (monomials_of R Q)  n  a p Q   (carrier R [𝒳I])) 
           Q   (carrier R [𝒳I])  card (monomials_of R Q) = Suc n  a p Q   (carrier R [𝒳I])"
  proof-
    fix n
    fix Q
    assume A0: "(Q. Q   (carrier R [𝒳I])  card (monomials_of R Q)  n  a p Q   (carrier R [𝒳I]))"
    assume A1: "Q   (carrier R [𝒳I])  card (monomials_of R Q) = Suc n"
    obtain m where m_def: "m  monomials_of R Q"
      using A1
      by fastforce
    obtain P where P_def: "P = (restrict_poly_to_monom_set R Q ((monomials_of R Q) - {m}))"
      by simp
    have "P  (carrier R [𝒳I])"
      using P_def A1 restriction_closed
      by blast
    have 0: "a p P   (carrier R [𝒳I])"
      using A0 P_def
      by (metis A1 One_nat_def P  (carrier R [𝒳I]) diff_Suc_Suc m_def
          minus_nat.diff_0 monomial_split_card nat_le_linear)
    have 1: "a p (mset_to_IP R m)  (carrier R [𝒳I])"
      using assms mset_to_IP_mult_closed[of a I m] A1 m_def mset_to_IP_indices
      by blast
    have 2: "a p (mset_to_IP R m p indexed_const (Q m))  (carrier R [𝒳I])"
      using P_ring_mult_assoc[of a "mset_to_IP R m" " indexed_const (Q m)"]
        1
      by (metis A1 assms(1) carrier_coeffE indexed_pset.indexed_const
          indexed_pset_in_carrier local.ring_axioms m_def mset_to_IP_closed'
          ring.indexed_const_P_mult_closed set_eq_subset)
    have 3: "(P  (mset_to_IP R m p indexed_const (Q m))) = Q"
      using P_def monomial_poly_split[of Q I m]
      using A1 m_def by blast
    have 4: "(a p P )  (a p (mset_to_IP R m p indexed_const (Q m))) = a p Q"
      using assms 2 3 P_ring_rdistr[of a P "(mset_to_IP R m p indexed_const (Q m))"]
      by (metis A1 Idl_subset_ideal' P_ring_mult_closed P  (carrier R [𝒳I])
          carrier_coeffE indexed_pset.indexed_const indexed_pset_in_carrier m_def
          mset_to_IP_closed' onepideal principalideal.generate)
    then show " a p Q   (carrier R [𝒳I])"
      by (metis "0" "2" indexed_pset.indexed_padd)
  qed
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Subtraction of Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

definition P_ring_uminus :: "('a,'b) ring_scheme  ('a,'c) mvar_poly  ('a,'c) mvar_poly" where
"P_ring_uminus R P = (λm. R(P m))"

context ring

begin


lemma P_ring_uminus_eq:
  assumes "a  indexed_pset I (carrier R)"
  shows "P_ring_uminus R a  = a p (indexed_const ( 𝟭))"
proof
    fix x
    have "(a p indexed_const ( 𝟭)) x = a x   𝟭"
      using indexed_const_P_multE[of a I " 𝟭" x] assms
      by blast
    then show "P_ring_uminus R a x = (a p indexed_const ( 𝟭)) x"
      unfolding P_ring_uminus_def
      using assms
      by (metis Idl_subset_ideal' carrier_coeffE indexed_pset_in_carrier
        one_closed onepideal principalideal.generate r_minus r_one)
qed

lemma P_ring_uminus_closed:
  assumes "a  indexed_pset I (carrier R)"
  shows "P_ring_uminus R a  indexed_pset I (carrier R)"
  using assms P_ring_uminus_eq
  by (metis add.l_inv_ex indexed_const_P_mult_closed minus_equality one_closed)

lemma P_ring_uminus_add:
  assumes "a  indexed_pset I (carrier R)"
  shows "P_ring_uminus R a  a = indexed_const 𝟬"
proof
  fix x
  show "(P_ring_uminus R a  a) x = indexed_const 𝟬 x"
    using assms
    unfolding P_ring_uminus_def indexed_const_def indexed_padd_def
    by (meson carrier_coeffE indexed_pset_in_carrier
        local.ring_axioms ring.ring_simprules(9) set_eq_subset)
qed

text‹multiplication by 1›

lemma one_mult_left:
  assumes "a  indexed_pset I (carrier R)"
  shows "(indexed_const 𝟭) p a = a"
proof
  fix m
  show "(indexed_const 𝟭 p a) m = a m "
    unfolding indexed_const_def P_ring_mult_def
  proof-
    have 0: "(x((mset_factors m) - {{#}}). (if x = {#} then 𝟭 else 𝟬)  a (m - x)) = 𝟬"
    proof-
      have "x. x((mset_factors m) - {{#}})   (if x = {#} then 𝟭 else 𝟬)  a (m - x) = 𝟬  a (m - x)"
        by simp
      then have "x. x((mset_factors m) - {{#}})   (if x = {#} then 𝟭 else 𝟬)  a (m - x) = 𝟬"
        using assms
        by (metis carrier_coeffE  indexed_pset_in_carrier local.ring_axioms ring.ring_simprules(24) set_eq_subset)
      then show ?thesis
        by (meson add.finprod_one_eqI)
    qed
    have 1: "(λx. (if x = {#} then 𝟭 else 𝟬)  a (m - x)) {#} = a m"
      by (metis (full_types) assms carrier_coeffE diff_zero
          indexed_pset_in_carrier local.ring_axioms ring.ring_simprules(12) set_eq_subset)
    have 2: "(xinsert {#} (mset_factors m - {{#}}). (if x = {#} then 𝟭 else 𝟬)  a (m - x)) =
              (λx. (if x = {#} then 𝟭 else 𝟬)  a (m - x)) {#} 
             (x((mset_factors m) - {{#}}). (if x = {#} then 𝟭 else 𝟬)  a (m - x))"
    proof-
      have 20:"finite (mset_factors m - {{#}})"
        by simp
      have 21: " {#}  mset_factors m - {{#}}"
        by blast
      have 22: " (λx. (if x = {#} then 𝟭 else 𝟬)  a (m - x))  mset_factors m - {{#}}  carrier R"
      proof
        fix x
        assume A: "x  mset_factors m - {{#}}"
        show "(if x = {#} then 𝟭 else 𝟬)  a (m - x)  carrier R"
          apply(cases "x = {#}")
          using A
          apply blast
              using assms carrier_coeffE indexed_pset_in_carrier  local.ring_axioms
              one_closed ring.ring_simprules(5) set_eq_subset zero_closed
              by (metis (mono_tags, opaque_lifting))
      qed
      have 23: "(if {#} = {#} then 𝟭 else 𝟬)  a (m - {#})  carrier R"
        by (metis "1" assms carrier_coeffE indexed_pset_in_carrier set_eq_subset)
      show ?thesis
        using 20 21 22 23 finsum_insert[of  "((mset_factors m) - {{#}})" "{#}"
            " (λx. (if x = {#} then 𝟭 else 𝟬)  a (m - x))"]
        by blast
    qed
    have 3: "insert {#} (mset_factors m - {{#}}) = mset_factors m"
    proof-
      have "{#}  mset_factors m"
        using monom_divides_factors subset_mset.zero_le
        by blast
      then show ?thesis
        by blast
    qed
    show "(xmset_factors m. (if x = {#} then 𝟭 else 𝟬)  a (m - x)) = a m"
      using 0 1 2 3 assms
      by (metis (no_types, lifting) Idl_subset_ideal' carrier_coeffE
          genideal_one indexed_pset_in_carrier one_closed r_zero)
  qed
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
                     subsection‹The Carrier of the Ring of Indexed Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)


abbreviation(input) Pring_set where
"Pring_set R I  ring.indexed_pset R I (carrier R) "

context ring

begin

lemma Pring_set_zero:
  assumes "f  Pring_set R I"
  assumes "¬ set_mset m   I"
  shows " f m = 𝟬R⇙"
proof-
  have "¬ set_mset m   I  f m = 𝟬R⇙"
  apply(induction m)
  apply simp
  by (meson assms complement_of_monomials_of mset_to_IP_indices')
  then show ?thesis
    using assms(2) by blast
qed

lemma(in ring) Pring_cfs_closed:
  assumes "P  Pring_set R I"
  shows "P m  carrier R"
  using assms carrier_coeffE indexed_pset_in_carrier
  by blast

lemma indexed_pset_mono_aux:
  assumes "P  indexed_pset I S"
  shows "S  T  P  indexed_pset I T"
  using assms
  apply(induction P)
  using indexed_pset.indexed_const apply blast
  using indexed_pset.indexed_padd apply blast
  by (simp add: indexed_pset.indexed_pmult)

lemma indexed_pset_mono:
  assumes "S  T"
  shows "indexed_pset I S  indexed_pset I T"
  using assms indexed_pset_mono_aux
  by blast

end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Scalar Multiplication›
(**************************************************************************************************)
(**************************************************************************************************)

definition poly_scalar_mult :: "('a, 'b) ring_scheme  'a  ('a,'c) mvar_poly  ('a,'c) mvar_poly" where
"poly_scalar_mult R c P = (λ m. c RP m)"

lemma(in cring) poly_scalar_mult_eq:
  assumes "c  carrier R"
  shows "P  Pring_set R (I :: 'c set)  poly_scalar_mult R c P = indexed_const c p P"
proof(erule indexed_pset.induct)
  show "k. k  carrier R  poly_scalar_mult R c (indexed_const k) = indexed_const c p indexed_const k"
  proof-
    fix k
    assume A0: "k  carrier R"
    show "poly_scalar_mult R c (indexed_const k)  = (indexed_const c p indexed_const k) "
      unfolding poly_scalar_mult_def
    proof
      show "m. c  indexed_const k m = (indexed_const c p indexed_const k) m"
          using indexed_const_P_mult_eq
          by (metis A0  assms indexed_const_P_mult_eq indexed_const_def local.semiring_axioms semiring.r_null)
    qed
  qed
  show "P Q. P  Pring_set R I 
           poly_scalar_mult R c P = indexed_const c p P 
           Q  Pring_set R I  poly_scalar_mult R c Q = indexed_const c p Q  poly_scalar_mult R c (P  Q) = indexed_const c p (P  Q)"
  proof
    fix P Q :: "'c monomial  'a"
    fix x :: "'c monomial"
    assume A0: "P  Pring_set R I"
    assume A1: "poly_scalar_mult R c P = indexed_const c p P"
    assume A2: "Q  Pring_set R I"
    assume A3: "poly_scalar_mult R c Q = indexed_const c p Q"
    show "poly_scalar_mult R c (P  Q) x = (indexed_const c p (P  Q)) x"
    proof-
      have P0: "poly_scalar_mult R c (P  Q) = (poly_scalar_mult R c P)  (poly_scalar_mult R c Q)"
        unfolding poly_scalar_mult_def
      proof
        fix m
        show "c  (P  Q) m = ((λm. c  P m)  (λm. c  Q m)) m"
        proof-
          have LHS: "c  (P  Q) m = c  ( P m  Q m) "
            by (simp add: indexed_padd_def)
          then have LHS1: "c  (P  Q) m = (c   P m)  (c Q m) "
          proof-
            have 0: "carrier_coeff P"
              using A0 indexed_pset_in_carrier
              by blast
            have 1: "P m  carrier R"
              using 0 carrier_coeffE by blast
            have 2: "carrier_coeff Q"
              using A2 indexed_pset_in_carrier
              by blast
            have 3: "Q m  carrier R" using 2
              using carrier_coeff_def
              by blast
            show ?thesis using 1 3 assms
              by (simp add: LHS r_distr)
          qed
          then show ?thesis
            by (simp add: indexed_padd_def)
        qed
      qed
      have P1: "poly_scalar_mult R c (P  Q) = (indexed_const c p P)  (indexed_const c p Q)"
        using P0 A1 A3
        by simp
      have P2: "indexed_const c  Pring_set R I"
        using assms indexed_pset.indexed_const by blast
      show ?thesis
        using P2 A0 A2 P1
        by (metis P_ring_rdistr indexed_pset_in_carrier set_eq_subset)
    qed
  qed
  show "P i. P  Pring_set R I 
           poly_scalar_mult R c P = indexed_const c p P  i  I  poly_scalar_mult R c (P  i) = indexed_const c p (P  i)"
  proof-
    fix P
    fix i
    assume A0:  "P  Pring_set R I"
    assume A1: "poly_scalar_mult R c P = indexed_const c p P"
    assume A2: "i  I"
    show "poly_scalar_mult R c (P  i) = indexed_const c p (P  i)"
    proof
      fix x
      show "poly_scalar_mult R c (P  i) x = (indexed_const c p (P  i)) x"
      proof-
        have RHS: "(indexed_const c p (P  i)) = (indexed_const c p P)  i"
        proof-
          have B0: "P  i = P p (mset_to_IP R {#i#})"
            by (meson A0 A2 poly_index_mult)
          have B1: " (indexed_const c p P)  i = (indexed_const c p P) p (mset_to_IP R {#i#})"
            by (meson A0 A2 P_ring_mult_closed' assms indexed_pset.simps poly_index_mult)
          have B2: "(indexed_const c p (P  i)) = indexed_const c p P p (mset_to_IP R {#i#})"
            by (metis A0 A2 B1 assms cring.P_ring_mult_comm indexed_const_var_mult
                indexed_pmult_in_carrier indexed_pset.indexed_const indexed_pset_in_carrier
                is_cring set_eq_subset)
          show ?thesis using A0 A1 A2 B2 B1 assms
          by (simp add: A0)
        then have RHS': "(indexed_const c p (P  i)) = (poly_scalar_mult R c P)  i"
          using A0 A1 A2 assms
          by simp
        qed
        show ?thesis apply(cases "i ∈# x")
          unfolding poly_scalar_mult_def
           apply (metis A1 RHS poly_scalar_mult_def indexed_pmult_def )
          by (metis RHS assms indexed_pmult_def r_null)
      qed
    qed
  qed
qed

lemma(in cring) poly_scalar_mult_const:
  assumes "c  carrier R"
  assumes "k  carrier R"
  shows "poly_scalar_mult R k (indexed_const c) = indexed_const (k  c)"
  using assms poly_scalar_mult_eq
  by (simp add: poly_scalar_mult_eq indexed_const_P_mult_eq indexed_pset.indexed_const)

lemma(in cring) poly_scalar_mult_closed:
  assumes "c  carrier R"
  assumes "P  Pring_set R I"
  shows "poly_scalar_mult R c P  Pring_set R I"
  using assms poly_scalar_mult_eq
  by (metis P_ring_mult_closed' indexed_pset.indexed_const)

lemma(in cring) poly_scalar_mult_zero:
  assumes "P  Pring_set R I"
  shows "poly_scalar_mult R 𝟬 P = indexed_const 𝟬"
proof
  fix x
  show "poly_scalar_mult R 𝟬 P x = indexed_const 𝟬 x"
    unfolding poly_scalar_mult_def
    using assms
    by (metis Pring_cfs_closed indexed_zero_def l_null)
qed

lemma(in cring) poly_scalar_mult_one:
  assumes "P  Pring_set R I"
  shows "poly_scalar_mult R 𝟭 P = P"
proof
  fix x show "poly_scalar_mult R 𝟭 P x = P x"
    using assms
    by (metis one_closed one_mult_left poly_scalar_mult_eq)
qed

lemma(in cring) times_poly_scalar_mult:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  assumes "k  carrier R"
  shows "P p (poly_scalar_mult R k Q) = poly_scalar_mult R k (P p Q)"
proof-
  have "P p (poly_scalar_mult R k Q) = P p (indexed_const k) p Q"
  by (metis assms(1) assms(2) assms(3)  indexed_pset_mono_aux
      local.ring_axioms poly_scalar_mult_eq ring.P_ring_mult_assoc ring.indexed_pset.intros(1)
      ring.indexed_pset_in_carrier subset_refl)
  then have  "P p (poly_scalar_mult R k Q) =  (indexed_const k) p P p Q"
  by (metis P_ring_mult_comm assms(1) assms(3) local.ring_axioms ring.indexed_pset.indexed_const ring.indexed_pset_in_carrier subset_refl)
  then show ?thesis
    by (metis (no_types, opaque_lifting) P_ring_mult_closed' Pring_cfs_closed assms(1) assms(2) assms(3)
        carrier_coeff_def  indexed_pset.indexed_const  local.ring_axioms poly_scalar_mult_eq ring.P_ring_mult_assoc)
qed

lemma(in cring) poly_scalar_mult_times:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  assumes "k  carrier R"
  shows " poly_scalar_mult R k (Q p P) = (poly_scalar_mult R k Q) p P"
  using assms times_poly_scalar_mult
  by (metis (no_types, opaque_lifting) P_ring_mult_comm  cring.P_ring_mult_comm
      cring.poly_scalar_mult_closed is_cring local.ring_axioms ring.indexed_pset_in_carrier subset_refl)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Defining the Ring of Indexed Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)
definition Pring :: "('b, 'e) ring_scheme   'a set   ('b, ('b,'a) mvar_poly) module" where

"Pring R I    carrier = Pring_set R I,
                Group.monoid.mult = P_ring_mult R ,
                one = ring.indexed_const R 𝟭R,
                zero = ring.indexed_const R  𝟬R,
                add = ring.indexed_padd R,
                smult = poly_scalar_mult R"

context ring

begin

lemma Pring_car:
"carrier (Pring R I) = Pring_set R I"
  unfolding Pring_def
  by auto

text‹Definitions of the operations and constants:›

lemma Pring_mult:
"a Pring R Ib = a p b"
  unfolding Pring_def
  by simp

lemma Pring_add:
"a Pring R Ib = a  b"
  unfolding Pring_def
  by simp

lemma Pring_zero:
"𝟬Pring R I= indexed_const 𝟬"
  unfolding Pring_def by simp

lemma Pring_one:
"𝟭Pring R I= indexed_const 𝟭"
  unfolding Pring_def by simp

lemma Pring_smult:
"(⊙Pring R I) = (poly_scalar_mult R)"
  unfolding Pring_def by simp

lemma Pring_carrier_coeff:
  assumes "a  carrier (Pring R I)"
  shows "carrier_coeff a"
  using assms indexed_pset_in_carrier[of "(carrier R)" a I] Pring_car
  by blast

lemma Pring_carrier_coeff'[simp]:
  assumes "a  carrier (Pring R I)"
  shows "a m  carrier R"
  using Pring_car[of I] assms carrier_coeffE indexed_pset_in_carrier
  by blast

lemma Pring_add_closed:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  shows "a Pring R Ib  carrier (Pring R I)"
  using assms  Pring_def[of R I]
  by (simp add: Pring_def[of R I] indexed_pset.indexed_padd)

lemma Pring_mult_closed:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  shows "a Pring R Ib  carrier (Pring R I)"
  using assms  P_ring_mult_closed'[of a I b] Pring_car[of I] Pring_mult[of I a b]
  by (simp add: a Pring R Ib = a p b carrier (Pring R I) = Pring_set R I)

lemma Pring_one_closed:
"𝟭Pring R I carrier (Pring R I)"
proof-
  have "indexed_const 𝟭  carrier (Pring R I)"
    using Pring_car indexed_pset.simps by blast
  then show ?thesis
    unfolding Pring_def by auto
qed

lemma Pring_zero_closed:
"𝟬Pring R I carrier (Pring R I)"
proof-
  have "indexed_const 𝟬  carrier (Pring R I)"
    using Pring_car indexed_pset.simps by blast
  then show ?thesis
    unfolding Pring_def by auto
qed

lemma Pring_var_closed:
  assumes "i  I"
  shows "var_to_IP R i  carrier (Pring R I)"
unfolding var_to_IP_def
  by (metis Pring_car Pring_one_closed assms indexed_pset.indexed_pmult
      local.ring_axioms monom_add_mset one_mset_to_IP ring.Pring_one)

text‹Properties of addition:›

lemma Pring_add_assoc:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  assumes "c  carrier (Pring R I)"
  shows "a Pring R I(b Pring R Ic) = (a Pring R Ib) Pring R Ic"
proof-
  have "a  (b  c) = (a   b)   c"
  proof
    fix x
    have "carrier_coeff a" "carrier_coeff b" "carrier_coeff c"
      using assms Pring_car[of I] Pring_carrier_coeff   apply blast
        using assms Pring_car[of I] Pring_carrier_coeff   apply blast
          using assms Pring_car[of I] Pring_carrier_coeff   by blast
    then have "a x  carrier R" "b x  carrier R" "c x  carrier R"
      using carrier_coeffE apply blast
        using carrier_coeff b carrier_coeffE apply blast
          using carrier_coeff c carrier_coeffE by blast
    show "(a  (b  c)) x = (a  b  c) x"
      unfolding indexed_padd_def
      using assms
      by (simp add: a x  carrier R b x  carrier R c x  carrier R add.m_assoc)
  qed
  then show ?thesis
    using assms
  by (simp add: Pring_add)
qed

lemma Pring_add_comm:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  shows "a Pring R Ib = b Pring R Ia"
proof-
  have "a  b = b  a"
  proof fix x
    show "(a  b) x = (b  a) x"
      using assms
      by (metis abelian_monoid.a_comm abelian_monoid_axioms
          indexed_padd_def local.ring_axioms ring.Pring_carrier_coeff')
  qed
  then show ?thesis
    by (simp add: Pring_add)
qed

lemma Pring_add_zero:
  assumes "a  carrier (Pring R I)"
  shows "a Pring R I𝟬Pring R I= a"
        "𝟬Pring R IPring R Ia  = a"
  using assms Pring_zero Pring_add
   apply (metis Pring_carrier_coeff indexed_padd_zero(1))
  using assms Pring_zero Pring_add
   by (metis Pring_carrier_coeff indexed_padd_zero(2))

text‹Properties of multiplication›

lemma Pring_mult_assoc:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  assumes "c  carrier (Pring R I)"
  shows "a Pring R I(b Pring R Ic) = (a Pring R Ib) Pring R Ic"
  using assms P_ring_mult_assoc[of a b c]
  by (metis Pring_carrier_coeff Pring_mult)

lemma Pring_mult_comm:
  assumes "cring R"
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  shows "a Pring R Ib = b Pring R Ia"
  using assms Pring_carrier_coeff[of a I] Pring_carrier_coeff[of b I]
        Pring_mult[of I b a]  Pring_mult[of I a b] cring.P_ring_mult_comm[of R a b]
  by metis

lemma Pring_mult_one:
  assumes "a  carrier (Pring R I)"
  shows "a Pring R I𝟭Pring R I= a"
proof
  fix x
  show "(a Pring R I𝟭Pring R I) x = a x "
  proof-
    have 0: "(a Pring R I𝟭Pring R I) x = (a p indexed_const 𝟭) x"
        using assms Pring_mult[of I a "𝟭Pring R I⇙" ] Pring_one[of I]
        by metis
    have 1: "𝟭  carrier R"
      by simp
    have 2: "(a Pring R I𝟭Pring R I) x =  a x  𝟭"
      using 0 1 indexed_const_P_multE[of a I 𝟭 x]
        assms Pring_car[of I]
      by metis
    then show ?thesis
      using assms by auto
  qed
qed

lemma Pring_mult_one':
  assumes "a  carrier (Pring R I)"
  shows "𝟭Pring R IPring R Ia = a"
  using one_mult_left[of a I]
        assms Pring_one Pring_mult
  by (simp add: Pring_mult Pring_one Pring_car)

text‹Distributive laws›

lemma Pring_mult_rdistr:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  assumes "c  carrier (Pring R I)"
  shows "a Pring R I(b Pring R Ic) = (a Pring R Ib) Pring R I(a Pring R Ic)"
proof-
  have "a p (b  c) = a p b  (a p c)"
  using P_ring_rdistr[of a b c]
        assms Pring_carrier_coeff
  by metis
  then have "a p (b Pring R Ic) = a p b Pring R I(a p c)"
    using       Pring_add[of I b c] Pring_add[of I "a p b" "a p c"]
    by auto
  then have "a Pring R I(b Pring R Ic) = (a p b) Pring R I(a p c)"
    using  Pring_mult[of I a "(b Pring R Ic)"]
    by auto
  then have "a Pring R I(b Pring R Ic) = (a Pring R Ib) Pring R I(a p c)"
    using Pring_mult[of I a b] by metis
  then show ?thesis
    using Pring_mult[of I a c] by metis
qed

lemma Pring_mult_ldistr:
  assumes "a  carrier (Pring R I)"
  assumes "b  carrier (Pring R I)"
  assumes "c  carrier (Pring R I)"
  shows "(b Pring R Ic) Pring R Ia  = (b Pring R Ia) Pring R I(c Pring R Ia)"
proof-
  have "(b  c) p a = b p a  (c p a)"
  using P_ring_ldistr[of a b c]
        assms Pring_carrier_coeff
  by metis
  then have " (b Pring R Ic)  p a = b p a Pring R I(c p a)"
    using       Pring_add[of I b c] Pring_add[of I "b p a" "c p a"]
    by auto
  then have " (b Pring R Ic) Pring R Ia = (b p a) Pring R I(c p a)"
    using  Pring_mult[of I "(b Pring R Ic)" a]
    by auto
  then have "(b Pring R Ic) Pring R Ia = (b Pring R Ia) Pring R I(c p a)"
    using Pring_mult[of I b a] by metis
  then show ?thesis
    using Pring_mult[of I c a] by metis
qed

text‹Properties of subtraction:›

lemma Pring_uminus:
  assumes "a  carrier (Pring R I)"
  shows "P_ring_uminus R a  carrier (Pring R I)"
  using P_ring_uminus_closed[of a I] Pring_car[of I] assms
  by metis

lemma Pring_subtract:
  assumes "a  carrier (Pring R I)"
  shows "P_ring_uminus R a Pring R Ia = 𝟬Pring R I⇙"
        "a Pring R IP_ring_uminus R a = 𝟬Pring R I⇙"
  using assms Pring_add[of I "P_ring_uminus R a " a] Pring_zero[of I]
  apply  (simp add: Pring_car local.ring_axioms ring.P_ring_uminus_add)
    using assms Pring_add[of I "P_ring_uminus R a " a] Pring_zero[of I]
    by (metis P_ring_uminus_add P_ring_uminus_closed Pring_add_comm Pring_car)

text‹Pring R I is a ring›

lemma Pring_is_abelian_group:
  shows "abelian_group (Pring R I)"
  apply(rule abelian_groupI)
  apply (simp add: Pring_add_closed)
    apply (simp add: local.ring_axioms ring.Pring_zero_closed)
      apply (simp add: Pring_add_assoc)
        apply (simp add: Pring_add_comm)
          apply (simp add: local.ring_axioms ring.Pring_add_zero(2))
  using Pring_subtract(1) Pring_uminus
  by blast

lemma Pring_is_monoid:
"Group.monoid (Pring R I)"
  apply(rule monoidI)
  using Pring_mult_closed apply blast
    apply (simp add: Pring_one_closed)
      apply (metis Pring_mult_assoc)
        using Pring_mult_one'
        apply blast
  using Pring_mult_one by blast

lemma Pring_is_ring:
  shows "ring (Pring R I)"
  apply(rule ringI)
  apply (simp add: Pring_is_abelian_group)
    apply (simp add: Pring_is_monoid)
      apply (simp add: Pring_mult_ldistr)
  by (simp add: Pring_mult_rdistr)

lemma Pring_is_cring:
  assumes "cring R"
  shows "cring (Pring R I)"
  apply(rule cringI)
  apply (simp add: Pring_is_abelian_group)
   apply (simp add: Pring_is_monoid assms local.ring_axioms
          monoid.monoid_comm_monoidI ring.Pring_mult_comm)
     by (simp add: Pring_mult_ldistr)

lemma Pring_a_inv:
  assumes "P  carrier (Pring R I)"
  shows "Pring R IP = P_ring_uminus R P"
proof-
  have 0: "P_ring_uminus R P  carrier (Pring R I)"
    using Pring_uminus assms
    by blast
  have 1: "P_ring_uminus R P Pring R IP = 𝟬Pring R I⇙"
    using Pring_subtract(1) assms
    by blast
  show ?thesis using 0 1 assms Pring_is_ring
  by (simp add: Pring_car Pring_is_abelian_group abelian_group.minus_equality)
qed


end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Defining the R-Algebra of Indexed Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

context cring
begin

lemma Pring_smult_cfs:
  assumes "a  carrier R"
  assumes "P  carrier (Pring R I)"
  shows "(a Pring R IP) m = a  (P m)"
  using assms Pring_smult
  by (simp add: Pring_smult poly_scalar_mult_def)

lemma Pring_smult_closed:
      " a x. [| a  carrier R; x  carrier (Pring R I) |] ==> a (Pring R I)x  carrier (Pring R I)"
  by (simp add: Pring_car Pring_smult poly_scalar_mult_closed)

lemma Pring_smult_l_distr:
      "!!a b x. [| a  carrier R; b  carrier R; x  carrier (Pring R I) |] ==>
      (a  b) (Pring R I)x = (a (Pring R I)x) (Pring R I)(b (Pring R I)x)"
proof- fix a b x assume A: "a  carrier R" "b  carrier R" "x  carrier (Pring R I)"
  show "(a  b) Pring R Ix = a Pring R Ix Pring R Ib Pring R Ix"
  proof fix m
    have "(a  b)  x m = (a  (x m))  (b  (x m))"
      by (meson A(1) A(2) A(3) Pring_carrier_coeff' local.semiring_axioms semiring.l_distr)
    thus "((a  b) Pring R Ix) m = (a Pring R Ix Pring R Ib Pring R Ix) m"
      using Pring_smult_cfs[of "a  b" x I m]
            Pring_smult_cfs[of "a" x I m]
            Pring_smult_cfs[of "b" x I m]
            Pring_smult_closed[of a x I]
            Pring_smult_closed[of b x I]
            Pring_add  A
      by (simp add: (a  b)  x m = a  x m  b  x m Pring_def indexed_padd_def poly_scalar_mult_def)
  qed
qed

lemma Pring_smult_r_distr:
      "!!a x y. [| a  carrier R; x  carrier (Pring R I); y  carrier (Pring R I) |] ==>
      a (Pring R I)(x (Pring R I)y) = (a (Pring R I)x) (Pring R I)(a (Pring R I)y)"
proof fix a x y m assume A: "a  carrier R" "x  carrier (Pring R I)" "y  carrier (Pring R I)"
  show "(a Pring R I(x Pring R Iy)) m =
       (a Pring R Ix Pring R Ia Pring R Iy) m"
    using Pring_smult_cfs[of a "x Pring R Iy" I m]
            Pring_smult_cfs[of a x I m]
            Pring_smult_cfs[of a y I m]
            Pring_smult_closed[of a x I]
            Pring_smult_closed[of a y I]
            Pring_add  A
  by (metis (no_types, lifting) Pring_add_closed Pring_carrier_coeff' indexed_padd_def l_distr m_comm)
qed

lemma Pring_smult_assoc1:
      "!!a b x. [| a  carrier R; b  carrier R; x  carrier (Pring R I) |] ==>
      (a  b) Pring R Ix = a Pring R I(b Pring R Ix)"
proof fix a b x m  assume A: "a  carrier R" "b  carrier R" "x  carrier (Pring R I)"
  show " (a  b Pring R Ix) m = (a Pring R I(b Pring R Ix)) m"
      using Pring_smult_cfs[of "a  b" x I m]
            Pring_smult_cfs[of a "b Pring R Ix" I m]
            Pring_smult_cfs[of "b" x I m]
            Pring_smult_closed[of a "b Pring R Ix" I]
            Pring_smult_closed[of b x I]
            A(1) A(2) A(3) m_assoc m_closed by auto
qed

lemma Pring_smult_one:
      "!!x. x  carrier (Pring R I) ==> (one R) Pring R Ix = x"
  by (simp add: Pring_car Pring_smult poly_scalar_mult_one)


lemma Pring_smult_assoc2:
      "!!a x y. [| a  carrier R; x  carrier (Pring R I); y  carrier (Pring R I) |] ==>
      (a Pring R Ix) Pring R Iy = a Pring R I(x Pring R Iy)"
  by (simp add: Pring_def poly_scalar_mult_times)

lemma Pring_algebra:
"algebra R (Pring R I)"
  apply(rule algebraI)
         apply (simp add: is_cring)
        apply (simp add: Pring_is_cring is_cring)
       apply (simp add: Pring_smult_closed)
      apply (simp add: Pring_smult_l_distr)
     apply (simp add: Pring_smult_r_distr)
    apply (simp add: Pring_smult_assoc1)
   apply (simp add: Pring_smult_one)
  by (simp add: Pring_smult_assoc2)


end



(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Evaluation of Polynomials and Subring Structure›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  In this section the aim is to define the evaluation of a polynomial over its base ring. We define
  both total evaluation of a polynomial at all variables, and partial evaluation at only a subset
  of variables. The basic input for evaluation is a variable assignment function mapping variables
  to ring elements. Once we can evaluate a polynomial $P$ in variables $I$ over a ring $R$ at an
  assignment $f: I \to R$, this can be generalized to evaluation of $P$ in some other ring $S$,
  given a variable assignment $f: I \to S$ and a ring homomorphism $\phi: R \to S$. We chose to
  define this by simply applying $\phi$ to the coefficients of $P$, and then using the first
  evaluation function over $S$. This could also have been done the other way around: define
  general polynomial evaluation over any ring, given a ring hom $\phi$, and then defining
  evaluation over the base ring $R$ as the specialization of this function to the case there
  $\phi = \mathit{id}_R$.›

definition remove_monom ::
"('a,'b) ring_scheme  'c monomial  ('a, 'c) mvar_poly  ('a, 'c) mvar_poly" where
"remove_monom R m P = ring.indexed_padd R P (poly_scalar_mult R (RP m) (mset_to_IP R m))"

context cring
begin

lemma remove_monom_alt_def:
  assumes "P  Pring_set R I"
  shows "remove_monom R m P n = (if n = m then 𝟬 else P n)"
  unfolding remove_monom_def
  apply(cases "n = m")
  using assms
   apply (metis Pring_cfs_closed cring.cring_simprules(3) poly_scalar_mult_def
          indexed_padd_def is_cring mset_to_IP_simp r_neg r_one)
   using assms
   by (metis Pring_cfs_closed add.l_cancel_one cring.cring_simprules(3)
       poly_scalar_mult_def indexed_padd_def is_cring mset_to_IP_simp' r_null zero_closed)

lemma remove_monom_zero:
  assumes "m  monomials_of R P"
  assumes "P  Pring_set R I"
  shows "remove_monom R m P = P"
proof
  fix x
  show "remove_monom R m P x = P x "
    apply(cases "x  monomials_of R P")
    using assms unfolding monomials_of_def remove_monom_def
  apply (metis cring.remove_monom_alt_def is_cring remove_monom_def)
    using assms unfolding monomials_of_def remove_monom_def
    by (metis assms(1) cring.remove_monom_alt_def is_cring local.ring_axioms
        remove_monom_def ring.complement_of_monomials_of)
qed

lemma remove_monom_closed:
  assumes "P  Pring_set R I"
  shows "remove_monom R m P  Pring_set R I"

  apply(cases "m  monomials_of R P")
  using assms poly_scalar_mult_closed[of "( P m)" "(mset_to_IP R m)" I] mset_to_IP_closed[of m I]
  unfolding remove_monom_def
  apply (meson Pring_cfs_closed add.inv_closed indexed_pset.indexed_padd mset_to_IP_closed')
  by (metis assms remove_monom_def remove_monom_zero)

lemma remove_monom_monomials:
  assumes "P  Pring_set R I"
  shows "monomials_of R (remove_monom R m P) = monomials_of R P - {m}"
proof
  show "monomials_of R (remove_monom R m P)  monomials_of R P - {m}"
    using assms remove_monom_alt_def[of P I m]
    unfolding monomials_of_def
    by auto
  show "monomials_of R P - {m}  monomials_of R (remove_monom R m P)"
    using assms remove_monom_alt_def[of P I m]
    unfolding monomials_of_def
    by auto
qed

text‹The additive decomposition of a polynomial by a monomial›

lemma remove_monom_eq:
  assumes "P  Pring_set R I"
  shows "P = (remove_monom R a P)  poly_scalar_mult R  (P a) (mset_to_IP R a)"
  unfolding remove_monom_def poly_scalar_mult_def
proof
  fix x
  show "P x = (P  (λm.  P a  mset_to_IP R a m)  (λm. P a  mset_to_IP R a m)) x"
  apply(cases "x = a")
     apply (metis Pring_cfs_closed assms l_minus l_zero local.ring_axioms mset_to_IP_simp one_closed r_neg r_one ring.indexed_padd_def)
  proof-
    assume A: "x  a"
    show "P x = (P  (λm.  P a  mset_to_IP R a m)  (λm. P a  mset_to_IP R a m)) x"
      using assms A
      unfolding mset_to_IP_def indexed_padd_def
      using Pring_cfs_closed by fastforce
  qed
qed

lemma remove_monom_restrict_poly_to_monom_set:
  assumes "P  Pring_set R I"
  assumes "monomials_of R P = insert a M"
  assumes "a  M"
  shows "(remove_monom R a P) = restrict_poly_to_monom_set R P M"
proof
  fix m
  show "remove_monom R a P m= restrict_poly_to_monom_set R P M m"
    apply(cases "m = a")
    using assms apply (metis remove_monom_alt_def restrict_poly_to_monom_set_def)
    using assms
    by (metis complement_of_monomials_of insert_iff remove_monom_alt_def restrict_poly_to_monom_set_def)
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Nesting of Polynomial Rings According to Nesting of Index Sets›
(**************************************************************************************************)
(**************************************************************************************************)

lemma(in ring) Pring_carrier_subset:
  assumes "J  I"
  shows "(Pring_set R J)  (Pring_set R I)"
proof
  fix P
  show "P  Pring_set R J P  Pring_set R I"
    apply(erule indexed_pset.induct)
    using indexed_pset.indexed_const apply blast
      using indexed_pset.indexed_padd apply blast
        by (meson assms indexed_pset.indexed_pmult subsetD)
qed

lemma(in cring) Pring_set_restrict_induct:
  shows "finite S  P. monomials_of R P = S   P  Pring_set R I   ( m  monomials_of R P. set_mset m  J)   P  Pring_set R J"
proof(erule finite.induct)
  show "P. monomials_of R P = {}  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J"
  proof
    fix P
    show "monomials_of R P = {}  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J"
    proof
      assume A0: "monomials_of R P = {}  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)"
      show "P  Pring_set R J "
        by (metis A0 card_eq_0_iff indexed_pset.indexed_const monomials_of_card_zero zero_closed)
    qed
  qed
  show "A a. finite A 
           P. monomials_of R P = A  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J 
           P. monomials_of R P = insert a A  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J"
  proof
    fix A :: "('c monomial) set" fix a fix P
    assume A1: "finite A"
    assume A2: "P. monomials_of R P = A  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J"
    show "monomials_of R P = insert a A  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)  P  Pring_set R J "
    proof
      assume A3: "monomials_of R P = insert a A  P  Pring_set R I  (mmonomials_of R P. set_mset m  J)"
      show "P  Pring_set R  J"
        apply(cases "a    A")
         apply (metis A2 A3 insert_absorb)
      proof-
        assume N: "a  A"
        obtain Q where Q_def: "Q = remove_monom R a P"
          by simp
        have Q0: "monomials_of R Q = A"
        proof-
          have 0:  "monomials_of R P = insert a A"
            by (simp add: A3)
          have 1: "P  Pring_set R I"
            using A3 by blast
          have 2: "monomials_of R (remove_monom R a P) = monomials_of R P - {a}"
            using A3 remove_monom_monomials by blast
          then show ?thesis
            using Q_def 0
            by (simp add: N)
        qed
        have Q1: "Q  Pring_set R I"
          using A3 Q_def remove_monom_closed by blast
        have Q2: "(mmonomials_of R Q. set_mset m  J)"
          using Q0 A3
          by blast
        have "Q  Pring_set R J"
          using A2 Q0 Q1 Q2 by blast
        then show "P  Pring_set R J"
        proof-
          have "P = Q  poly_scalar_mult R  (P a) (mset_to_IP R a)"
            using Q_def remove_monom_eq
            using A3 by blast
          then show ?thesis
            by (metis A3 Pring_cfs_closed Q  Pring_set R J indexed_pset.indexed_padd insertCI mset_to_IP_closed poly_scalar_mult_closed)
        qed
      qed
    qed
  qed
qed

lemma(in cring) Pring_set_restrict:
  assumes "P  Pring_set R I"
  assumes "(m.  m  monomials_of R P   set_mset m  J)"
  shows " P  Pring_set R J"
  using assms Pring_set_restrict_induct[of "monomials_of R P"]
  by (metis monomials_finite)

lemma(in ring) Pring_mult_eq:
  fixes I:: "'c set"
  fixes J:: "'c set"
  shows "(⊗Pring R I) = (⊗Pring R J)"
  by (simp add: Pring_def)

lemma(in ring) Pring_add_eq:
  fixes I:: "'c set"
  fixes J:: "'c set"
  shows "(⊕Pring R I) = (⊕Pring R J)"
  using Pring_def
  by (simp add: Pring_def)

lemma(in ring) Pring_one_eq:
  fixes I:: "'c set"
  fixes J:: "'c set"
  shows "(𝟭Pring R I) = (𝟭Pring R J)"
  using Pring_def
  by (simp add: Pring_def)

lemma(in ring) Pring_zero_eq:
  fixes I:: "'c set"
  fixes J:: "'c set"
  shows "(𝟬Pring R I) = (𝟬Pring R J)"
  using Pring_def
  by (simp add: Pring_def)

lemma(in ring) index_subset_Pring_subring:
  assumes "J  I"
  shows "subring (carrier (Pring R J)) (Pring R I)"
  apply(rule ring.subringI)
  apply (simp add: Pring_is_ring; fail)
    using assms
   apply (simp add: Pring_car Pring_carrier_subset; fail)
   using Pring_def
         apply (simp add: Pring_def indexed_pset.indexed_const; fail)
    proof-
      show "h. h  carrier (Pring R J)  Pring R Ih  carrier (Pring R J)"
      proof-
        fix h
        assume A: "h  carrier (Pring R J)"
        then have A0: "Pring R Jh = P_ring_uminus R h"
          using Pring_a_inv[of h J]
          by auto
        have A1: "Pring R Ih = P_ring_uminus R h"
          using assms A Pring_carrier_subset[of J I] Pring_a_inv[of h I] Pring_car
          by blast
        show "Pring R Ih  carrier (Pring R J)"
          using A0 A1 A
          by (simp add: Pring_uminus)
      qed
      show " h1 h2. h1  carrier (Pring R J)  h2  carrier (Pring R J)  h1 Pring R Ih2  carrier (Pring R J)"
        using assms Pring_mult_eq
        by (metis Pring_mult_closed)
      show " h1 h2. h1  carrier (Pring R J)  h2  carrier (Pring R J)  h1 Pring R Ih2  carrier (Pring R J)"
        using assms Pring_add_eq
        by (metis Pring_add_closed)
    qed


(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Inclusion Maps›
(**************************************************************************************************)
(**************************************************************************************************)

definition Pring_inc :: "('a, 'c) mvar_poly  ('a, 'c) mvar_poly" where
"Pring_inc  (λP. P)"

lemma(in ring) Princ_inc_is_ring_hom:
  assumes "J  I"
  shows "ring_hom_ring (Pring R J) (Pring R I) Pring_inc"
  unfolding Pring_inc_def
  apply(rule ring_hom_ringI)
  apply (simp add: Pring_is_ring)
    using Pring_is_ring apply blast
       using index_subset_Pring_subring[of I J] assms index_subset_Pring_subring subringE(1)
       apply blast
          using Pring_mult_eq[of I J]
          apply (simp add: Pring_mult)
             using Pring_add_eq[of I J]
             apply (simp add: Pring_add)
             using Pring_one_eq
  by (simp add: Pring_one_eq)

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Restricting a Multiset to a Subset of Variables›
(**************************************************************************************************)
(**************************************************************************************************)

definition restrict_to_indices :: "'c monomial  'c set  'c monomial" where
"restrict_to_indices m S = filter_mset (λi. i  S) m"

lemma restrict_to_indicesE:
  assumes "i ∈# restrict_to_indices m S"
  shows "i  S"
  using assms
  unfolding restrict_to_indices_def
  by simp

lemma restrict_to_indicesI[simp]:
  assumes "i ∈# m"
  assumes "i  S"
  shows "i ∈# restrict_to_indices m S"
  using assms
  unfolding restrict_to_indices_def
  by simp

lemma restrict_to_indices_not_in[simp]:
  assumes "i ∈# m"
  assumes "i  S"
  shows "i ∉# restrict_to_indices m S"
  using assms
  unfolding restrict_to_indices_def
  by (meson count_eq_zero_iff count_filter_mset)

lemma restrict_to_indices_submultiset[simp]:
"restrict_to_indices m S ⊆# m"
  unfolding restrict_to_indices_def
  using multiset_filter_subset
  by blast

lemma restrict_to_indices_add_element:
"restrict_to_indices (add_mset x m) S = (if x  S then (add_mset x (restrict_to_indices m S) )
                                            else (restrict_to_indices m S) )"
  unfolding restrict_to_indices_def
  by (metis filter_mset_add_mset)

lemma restrict_to_indices_count[simp]:
"count (restrict_to_indices m S) i = (if (i  S) then (count m i) else 0)"
  unfolding restrict_to_indices_def
  by (metis count_filter_mset)

lemma restrict_to_indices_subset:
"restrict_to_indices m S = restrict_to_indices m (set_mset m  S)"
proof(induction m)
  case empty
  then show ?case unfolding restrict_to_indices_def
    by (metis filter_empty_mset)
next
  case (add x m)
  assume IH: "restrict_to_indices m S = restrict_to_indices m (set_mset m  S)"
  show "restrict_to_indices (add_mset x m) S = restrict_to_indices (add_mset x m) (set_mset (add_mset x m)  S)"
  proof-
    have "i. count (restrict_to_indices (add_mset x m) S) i  =
              count (restrict_to_indices (add_mset x m) (set_mset (add_mset x m)  S)) i  "
    proof-
      fix i
      show "count (restrict_to_indices (add_mset x m) S) i  = count (restrict_to_indices (add_mset x m) (set_mset (add_mset x m)  S)) i"
        apply(cases "i  S")
        using restrict_to_indices_count
        apply (metis IntI count_inI)
        by (metis restrict_to_indices_count Int_iff)
    qed
    then show ?thesis
      using multiset_eq_iff by blast
  qed
qed

text‹\texttt{Restrict\_to\_indices} only depends on the intersection
     of the index set with the set of indices in m:›

lemma restrict_to_indices_subset':
  assumes "(set_mset m)  S = (set_mset m)  S'"
  shows "restrict_to_indices m S = restrict_to_indices m S'"
  by (metis restrict_to_indices_subset assms)

lemma mset_add_plus:
  assumes "m = n + k"
  shows "add_mset x m = (add_mset x n) + k"
  using assms
  by simp

text‹Restricting to S› and the complement of S› partitions m›:›

lemma restrict_to_indices_decomp:
  "m = (restrict_to_indices m S) + (restrict_to_indices m ((set_mset m) - S))"
  apply(induction m)
  apply (metis add.right_neutral empty_Diff restrict_to_indices_submultiset set_mset_empty subset_mset.le_zero_eq)
proof-
  fix x
  fix m
  assume A: "m = restrict_to_indices m S + restrict_to_indices m (set_mset m - S)"
  show "add_mset x m = restrict_to_indices (add_mset x m) S + restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S)"
  proof(cases "x  S")
    case True
    then have T0: "restrict_to_indices (add_mset x m) S = (add_mset x (restrict_to_indices m S) ) "
      by (simp add: True restrict_to_indices_add_element)
    have T1: "restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) =
            restrict_to_indices m (set_mset (add_mset x m) - S)"
      using True by (metis DiffD2 restrict_to_indices_add_element)
    have T2: "restrict_to_indices (add_mset x m) S +  restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S)
            = (add_mset x (restrict_to_indices m S) )  + restrict_to_indices m (set_mset (add_mset x m) - S)"
      using T0 T1
      by presburger
    have T3: " add_mset x m = add_mset x (restrict_to_indices m S) + restrict_to_indices m (set_mset m - S)"
      using T2 A using mset_add_plus[of m "restrict_to_indices m S" " restrict_to_indices m (set_mset m - S)" x]
      by blast
    have T4: "set_mset m  (set_mset (add_mset x m) - S) = set_mset m  (set_mset m - S)"
    proof
      show "set_mset m  (set_mset (add_mset x m) - S)  set_mset m  (set_mset m - S)"
        by blast
      show "set_mset m  (set_mset m - S)  set_mset m  (set_mset (add_mset x m) - S)"
        by (simp add: True)
    qed
    have T5: "restrict_to_indices m (set_mset (add_mset x m) - S) = restrict_to_indices m (set_mset m - S)"
      using T4 restrict_to_indices_subset'[of m "(set_mset (add_mset x m) - S)" " (set_mset m - S)" ]
      by blast
    then show ?thesis using T4
      by (metis T0 T1 T3)
  next
    case False
    then have F0: "restrict_to_indices (add_mset x m) S = (restrict_to_indices m S) "
      by (simp add: False restrict_to_indices_add_element)
    have F1: "restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) =
            (add_mset x (restrict_to_indices m (set_mset (add_mset x m) - S)))"
      using False
      by (meson DiffI restrict_to_indices_add_element union_single_eq_member)

    have F2: "restrict_to_indices (add_mset x m) S +  restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S)
            = (restrict_to_indices m S) + (add_mset x (restrict_to_indices m (set_mset (add_mset x m) - S)))"
      using F0 F1
      by presburger
    have F3: " add_mset x m =  (restrict_to_indices m S) +  (add_mset x (restrict_to_indices m (set_mset m - S)))"
      using F2 A  mset_add_plus[of m "restrict_to_indices m (set_mset m - S)" "restrict_to_indices m S" x]
      by (metis union_mset_add_mset_right)
    have F4: " add_mset x m =  restrict_to_indices (add_mset x m) S+  (add_mset x (restrict_to_indices m (set_mset m - S)))"
      using F0 F3 by auto
    have F5: "add_mset x (restrict_to_indices m (set_mset m - S)) = restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) "
    proof(cases "x  set_mset m")
      case True
      then show ?thesis
        by (metis F1 add_mset_remove_trivial more_than_one_mset_mset_diff)
    next
      case F: False
      have "set_mset m  (set_mset (add_mset x m) - S) = set_mset m (set_mset m - S)"
      proof
        show "set_mset m  (set_mset (add_mset x m) - S)  set_mset m  (set_mset m - S)"
          using F False
          by blast
        show "set_mset m  (set_mset m - S)  set_mset m  (set_mset (add_mset x m) - S)"
          using F False
          by (metis Diff_mono mset_add_plus Int_mono add_cancel_right_left
              set_eq_subset subsetI subset_iff union_commute union_iff)
      qed
      then show ?thesis
        by (metis F1 restrict_to_indices_subset')
    qed
    then show ?thesis
      using False F4
      by presburger
  qed
qed

definition remove_indices :: "'c monomial  'c set  'c monomial" where
"remove_indices m S = (restrict_to_indices m (set_mset m - S))"

lemma remove_indices_decomp:
"m = (restrict_to_indices m S) + (remove_indices m S)"
  unfolding remove_indices_def
  using restrict_to_indices_decomp
  by blast

lemma remove_indices_indices[simp]:
  assumes "set_mset m  I"
  shows "set_mset (remove_indices m S)  I - S"
  unfolding remove_indices_def using assms
  by (meson Diff_iff restrict_to_indicesE subsetD subsetI)

subsubsection‹Total evaluation of a monomial›

text‹
  We define total evaluation of a monomial first, and then define the partial evaluation of a
  monomial in terms of this.
›

abbreviation(input) closed_fun where
"closed_fun R g  g  UNIV  carrier R"

definition monom_eval :: "('a, 'b) ring_scheme  'c monomial  ('c  'a)  'a"  where
"monom_eval R (m:: 'c monomial) g = fold_mset (λ x . λ y. if y  carrier R then (g x) Ry else 𝟬R) 𝟭Rm"

context cring
begin

lemma closed_fun_simp:
  assumes "closed_fun R g"
  shows "g n  carrier R"
  using assms
  by blast

lemma closed_funI:
  assumes "x. g x  carrier R"
  shows "closed_fun R g"
  by (meson Pi_I assms)


text‹The following are necessary technical lemmas to prove properties of about folds over multisets:›

lemma monom_eval_comp_fun:
  fixes g:: "'c  'a"
  assumes "closed_fun R g"
  shows "comp_fun_commute (λ x . λy. if y  carrier R then  (g x)  y else 𝟬)"
  unfolding comp_fun_commute_def
proof-
  have " x y. (λya. if ya  carrier R then g y  ya else 𝟬)  (λy. if y  carrier R then g x  y else 𝟬) =
          (λy. if y  carrier R then g x  y else 𝟬)  (λya. if ya  carrier R then g y  ya else 𝟬)"
  proof
    fix x y a
    show "((λya. if ya  carrier R then g y  ya else 𝟬)  (λy. if y  carrier R then g x  y else 𝟬)) a =
       ((λy. if y  carrier R then g x  y else 𝟬)  (λya. if ya  carrier R then g y  ya else 𝟬)) a"
    proof(cases "a  carrier R")
      case True
      then show ?thesis
      proof-
        have LHS: "((λya. if ya  carrier R then g y  ya else 𝟬)  (λy. if y  carrier R then g x  y else 𝟬)) a =
                    ((λya. if ya  carrier R then g y  ya else 𝟬)  (g x  a))"
          using True assms(1)  m_closed m_lcomm
          unfolding o_def
          by presburger

      then have LHS': "((λya. if ya  carrier R then g y  ya else 𝟬)  (λy. if y  carrier R then g x  y else 𝟬)) a =   g y  (g x  a) "
        using True assms m_closed
        by (meson PiE UNIV_I)
      then show ?thesis
        unfolding o_def
        using True assms m_closed m_lcomm
        by (smt PiE UNIV_I)
      qed
    next
      case False
      then show ?thesis
        unfolding o_def
        using assms r_null closed_fun_simp
        by smt
    qed
  qed
  then show " y x. (λya. if ya  carrier R then g y  ya else 𝟬)  (λy. if y  carrier R then g x  y else 𝟬) =
          (λy. if y  carrier R then g x  y else 𝟬)  (λya. if ya  carrier R then g y  ya else 𝟬)"
    by blast
qed

lemma monom_eval_car:
  assumes "closed_fun R g"
  shows "monom_eval R (m:: 'c monomial) g  carrier R"
proof(induction m)
case empty
  then show ?case
    unfolding monom_eval_def
    by (metis fold_mset_empty one_closed)
next
  case (add x m)
  fix x m
  assume A: "monom_eval R m g  carrier R"
  obtain f where f_def: "f = (λ x . λy. if y  carrier R then (g x)  y else 𝟬)"
    by blast
  have 0: "comp_fun_commute f"
    using assms monom_eval_comp_fun[of g] f_def
    by blast
  have 1: "m. monom_eval R m g = fold_mset f 𝟭 m"
    using f_def monom_eval_def
    by blast
  have 2: "monom_eval R (add_mset x m) g = fold_mset f 𝟭 (add_mset x m)"
    using 1 by blast
  have 3: "g x  carrier R"
    using assms  by blast
  then show "monom_eval R (add_mset x m) g  carrier R"
    using assms 0 1 2 3
    by (metis A comp_fun_commute.fold_mset_add_mset f_def m_closed)
qed

text‹Formula for recursive (total) evaluation of a monomial:›
lemma monom_eval_add:
  assumes "closed_fun R g"
  shows "monom_eval R (add_mset x M) g = (g x)  (monom_eval R M g)"
proof-
  obtain f where f_def: "f = (λ x . λy. if y  carrier R then (g x)  y else 𝟬)"
    by blast
  have 0: "comp_fun_commute f"
    using assms monom_eval_comp_fun f_def
    by blast
  have 1: "m. monom_eval R m g = fold_mset f 𝟭 m"
    using f_def monom_eval_def
    by blast
  have 2: "monom_eval R (add_mset x M) g = fold_mset f 𝟭 (add_mset x M)"
    using 1 by blast
  have 3: "g x  carrier R"
    using assms  by blast
  have 4: "(g x)  (monom_eval R M g) = f x (monom_eval R M g)"
    using f_def 3
    by (meson assms monom_eval_car)
  then show ?thesis
    by (metis "0" "1" comp_fun_commute.fold_mset_add_mset)
qed

end

text‹
  This function maps a polynomial $P$ to the set of monomials in $P$ which, after evaluating all
  variables in the set $S$ to values in the ring $R$, reduce to the monomial $n$.
›

definition monomials_reducing_to ::
"('a,'b) ring_scheme  'c monomial  ('a,'c) mvar_poly  'c set  ('c monomial) set" where
"monomials_reducing_to R n P S = {m  monomials_of R P. remove_indices m S = n}"

lemma monomials_reducing_to_subset[simp]:
 "monomials_reducing_to R n P s  monomials_of R P"
  unfolding monomials_reducing_to_def
  by blast

context cring
begin

lemma monomials_reducing_to_finite:
  assumes "P  Pring_set R I"
  shows "finite (monomials_reducing_to R n P s)"
  by (meson assms monomials_finite monomials_reducing_to_subset rev_finite_subset)

lemma monomials_reducing_to_disjoint:
  assumes "n1  n2"
  shows "monomials_reducing_to R n1 P S  monomials_reducing_to R n2 P S = {}"
  unfolding monomials_reducing_to_def
  using assms
  by blast

lemma monomials_reducing_to_submset:
assumes "n ⊂# m"
shows "n  monomials_reducing_to R m P S"
proof(rule ccontr)
  assume C: "¬ n  monomials_reducing_to R m P S "
  then have "n   monomials_reducing_to R m P S "
    by blast
  then have "remove_indices n S = m"
    unfolding monomials_reducing_to_def
    by blast
  then show False
    by (metis (full_types) remove_indices_def restrict_to_indices_submultiset assms subset_mset.less_asym' subset_mset.less_irrefl subset_mset_def)
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Partial Evaluation of a Polynomial›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  This function takes as input a set $S$ of variables, an evaluation function $g$, and a polynomial
  to evaluate $P$. The output is a polynomial which is the result of evaluating the variables from
  the set $S$ which occur in $P$, according to the evaluation function $g$.
›

definition poly_eval ::
  "('a,'b) ring_scheme  'c set  ('c  'a)  ('a, 'c) mvar_poly  ('a, 'c) mvar_poly"where
"poly_eval R S g P m = (finsum R (λn. monom_eval R (restrict_to_indices n S) g  R(P n)) (monomials_reducing_to R m P S))"

context cring
begin

lemma finsum_singleton:
  assumes "S = {s}"
  assumes "f s  carrier R"
  shows "finsum R f S = f s"
proof-
  have "finsum R f S = finsum R f (insert s {})"
    using assms(1)
    by blast
  then show ?thesis using finsum_insert[of "{}" s f] assms
    by (metis Pi_I empty_iff finite.emptyI finsum_empty r_zero)
qed

lemma poly_eval_constant:
  assumes "k  carrier R"
  shows "poly_eval R S g (indexed_const k) = (indexed_const k)"
proof
  have S: "monomials_of R (indexed_const k)  {{#}}"
    unfolding indexed_const_def monomials_of_def
    by (metis (mono_tags, lifting) mem_Collect_eq singletonI subset_iff)
  fix x
  show "poly_eval R S g (indexed_const k) x = indexed_const k x "
  proof(cases "x = {#}")
    case True
    have "(monomials_reducing_to R x (indexed_const k) S)  {{#}}"
      using S monomials_reducing_to_subset
      by blast
    then show ?thesis
    proof(cases "k = 𝟬")
      case True
      then have "(monomials_reducing_to R x (indexed_const k) S) = {}"
        by (metis S monomials_reducing_to R x (indexed_const k) S  {{#}}
            monomials_reducing_to_subset monoms_of_const subset_antisym subset_singletonD)
      then show ?thesis
        unfolding poly_eval_def
        by (metis True finsum_empty indexed_const_def)
    next
      case False
      then have "monomials_of R (indexed_const k) = {{#}}"
        by (meson monoms_of_const)
      have "remove_indices {#} S = {#}"
        using remove_indices_decomp by blast
      then have "{#}  monomials_reducing_to R x (indexed_const k) S"
        using True False unfolding monomials_reducing_to_def
              monomials_of R (indexed_const k) = {{#}}
        by blast
      then have 0: "monomials_reducing_to R x (indexed_const k) S = {{#}}"
        using monomials_reducing_to R x (indexed_const k) S  {{#}}
        by blast
      have 1: "restrict_to_indices {#} S = {#}"
        using restrict_to_indices_submultiset remove_indices_decomp by blast
      have 2: "monom_eval R (restrict_to_indices {#} S) g = 𝟭"
        unfolding monom_eval_def
        using 1
        by (metis fold_mset_empty)
      have 3: "poly_eval R S g (indexed_const k)  x =
                (n{{#}}. monom_eval R (restrict_to_indices n S) g  indexed_const k n)"
        unfolding poly_eval_def
        using 0
        by presburger
      have 4: "(n{{#}}. monom_eval R (restrict_to_indices n S) g  indexed_const k n) = monom_eval R (restrict_to_indices {#} S) g  indexed_const k {#}"
        using finsum_singleton[of "{{#}}" "{#}" "λn. monom_eval R (restrict_to_indices n S) g  indexed_const k n" ]
        by (metis "2" assms indexed_const_def l_one)
      then show ?thesis unfolding poly_eval_def
        using 0 1 2
        by (metis True assms indexed_const_def l_one)
    qed
  next
    case False
    then have F0: "(indexed_const k) x = 𝟬"
      by (meson indexed_const_def)
    have "(monomials_reducing_to R x (indexed_const k) S) = {}"
      unfolding monomials_reducing_to_def
    proof(rule ccontr)
      assume A: "{m  monomials_of R (indexed_const k). remove_indices m S = x}  {}"
      then obtain m where m_def: "m  monomials_of R (indexed_const k)  remove_indices m S = x"
        by blast
        then show False using A F0
          by (metis False S empty_eq_union empty_iff
              remove_indices_decomp singletonD subset_singletonD)
    qed
    then show ?thesis
      unfolding poly_eval_def
      by (metis False finsum_empty indexed_const_def)
  qed
qed

lemma finsum_partition:
  assumes "finite S"
  assumes "f  S  carrier R"
  assumes "T  S"
  shows "finsum  R f S  = finsum R f T   finsum  R f (S - T) "
proof-
  have "U. finite U  U  S  finsum  R f S  = finsum R f U   finsum  R f (S - U) "
  proof-
    fix U
    show "finite U  U  S  finsum R f S = finsum R f U  finsum R f (S - U) "
    apply(erule finite.induct)
       apply (metis Diff_empty assms(2) finsum_closed finsum_empty l_zero)
    proof
      fix A :: "'c set" fix a
      assume A0: "finite A"
      show "A  S  finsum R f S = finsum R f A  finsum R f (S - A)  insert a A  S  finsum R f S = finsum R f (insert a A)  finsum R f (S - insert a A)"

      proof-
        assume A1: "A  S  finsum R f S = finsum R f A  finsum R f (S - A)"
        assume A2: "insert a A  S"
        show "finsum R f S = finsum R f (insert a A)  finsum R f (S - insert a A)"
          apply(cases "a  A")
           apply (metis A1 A2 insert_absorb)
        proof-
          assume A3: "a  A"
          have A4: "f a  carrier R"
            by (metis A2 Pi_iff assms(2) insert_subset)
          have A5: " finsum R f (insert a A) = f a  finsum R f A"

            using A0 A1 A2 finsum_insert[of A a f] assms  A3
            by blast
          have A6: "a  S"
            using A2 by blast
          have A7: "finsum R f S  carrier R"
            using assms(2) finsum_closed by blast
          have A8: " finsum R f (S - A) carrier R"
            using Diff_subset[of S A] Pi_iff assms(2) finsum_closed[of f] subsetD[of _ S ]
            by (meson Pi_anti_mono in_mono)
          have A9: "finsum R f A  carrier R"
            by (meson A2 Pi_anti_mono assms(2) finsum_closed insert_subset subsetD)
          have A10: "finsum R f A =  finsum R f S  finsum R f (S - A)"
            using A7 A8 A9
            by (metis A1 A2 add.inv_solve_right' insert_subset minus_eq)
          have A11: " finsum R f (insert a A) = f a  (finsum R f S  finsum R f (S - A))"
            using A5 A6 A1 A2 assms  A10
            by presburger
          then have A12: " finsum R f (insert a A) =finsum R f S ( f a   finsum R f (S - A))"
            using A4 A7 A8 add.inv_closed add.m_lcomm minus_eq by presburger
          have A13: "finsum R f (insert a A)  carrier R"
            using A4 A5 A9 add.m_closed
            by presburger
          have A14: " finsum R f (S - insert a A)  carrier R"
            by (meson Diff_subset Pi_anti_mono assms(2) finsum_closed in_mono)
          have A15: "finsum R f S = finsum R f (insert a A)  ( f a   finsum R f (S - A)) "
            by (metis A12 A13 A4 A7 A8 add.inv_solve_right' minus_closed minus_eq)
          have A16: "finsum R f S = finsum R f (insert a A)  finsum R f (S - A)  f a"
            using  A1 A2 A4 A5 A8 A9 add.inv_closed add.m_assoc add.m_comm insert_subset minus_closed minus_eq r_neg1
            unfolding a_minus_def
            by (metis add.m_closed)
          have A16: "finsum R f S = finsum R f (insert a A)  (finsum R f (S - A)  f a)"
            unfolding a_minus_def
            using A13 A16 A4 A8 add.inv_closed add.m_assoc minus_eq by presburger
          have A17: "(finsum R f (S - A)  f a) = finsum R f (S - insert a A) "
          proof-
            have A170: "(S - A) = insert a (S - insert a A)"
              using A3 A6 by blast
            have A171: "a  S - insert a A"
              by blast
            then have "finsum R f (S - A)  =(f a)  finsum R f (S - insert a A) "
              using A170 finsum_insert[of "(S - insert a A)" a f]
              by (metis A4 Diff_subset Pi_anti_mono assms(1) assms(2)
                  rev_finite_subset subsetD )
            then show ?thesis
              by (metis A1 A13 A14 A16 A2 A4 A5 A8 A9 add.l_cancel
                  add.m_assoc add.m_comm insert_subset minus_closed)
          qed
          then show "finsum R f S = finsum R f (insert a A)  finsum R f (S - insert a A) "
            using A16
            by presburger
        qed
      qed
    qed
  qed
  then show ?thesis
    by (meson assms(1) assms(3) rev_finite_subset)
qed

lemma finsum_eq_parition:
  assumes "finite S"
  assumes "f  S  carrier R"
  assumes "T  S"
  assumes "x. x  S - T  f x = 𝟬"
  shows "finsum  R f S  = finsum R f T"
  using assms
  by (metis add.finprod_mono_neutral_cong_right)

lemma poly_eval_scalar_mult:
  assumes "k  carrier R"
  assumes "closed_fun R g"
  assumes "P  Pring_set R I"
  shows "poly_eval R  S g (poly_scalar_mult R k P)=
         (poly_scalar_mult R k (poly_eval R S g P))"
proof
  fix m
  show "poly_eval R S g (poly_scalar_mult R k P) m = poly_scalar_mult R k (poly_eval R S g  P) m"
    unfolding poly_eval_def poly_scalar_mult_def
  proof-
    have 0: "(nmonomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g  (k  P n)) =
              (nmonomials_reducing_to R m (λm. k  P m) S. monom_eval R (restrict_to_indices n S) g  (k  P n)) "
    proof-
      have 00: "monomials_reducing_to R m (λm. k  P m) S  monomials_reducing_to R m P S"
      proof
        fix x show " x  monomials_reducing_to R m (λm. k  P m) S  x  monomials_reducing_to R m P S"
        unfolding monomials_reducing_to_def
        using assms  assms(1) monomials_ofE complement_of_monomials_of  r_null[of k]
        by (metis (no_types, lifting) mem_Collect_eq)
       have 01: "(λn. monom_eval R (restrict_to_indices n S) g  (k  P n))  monomials_reducing_to R m P S  carrier R"
        by (smt Pi_I Pring_cfs_closed assms(1) assms(2) assms(3) closed_fun_simp m_closed monom_eval_car)
       have 02: "finite (monomials_reducing_to R m P S)"
        using assms(3) monomials_reducing_to_finite
        by blast
       have 03: "  (x. x  monomials_reducing_to R m P S - monomials_reducing_to R m (λm. k  P m) S 
          monom_eval R (restrict_to_indices x S) g  (k  P x) = 𝟬)"
       proof-
        fix x
        assume A: "x  monomials_reducing_to R m P S - monomials_reducing_to R m (λm. k  P m) S "
        have "x  monomials_of R ((λm. k  P m))"
        proof
          assume "x  monomials_of R (λm. k  P m)"
          then have "x   monomials_reducing_to R m (λm. k  P m) S"
            using A
            unfolding monomials_reducing_to_def
            by blast
          then show False
            using A by blast
        qed
        then show "monom_eval R (restrict_to_indices x S) g  (k  P x) = 𝟬"
          by (metis assms(2) complement_of_monomials_of monom_eval_car r_null)
       qed
      qed
      have 01: " (x. x  monomials_reducing_to R m P S - monomials_reducing_to R m (λm. k  P m) S  monom_eval R (restrict_to_indices x S) g  (k  P x) = 𝟬)"
      proof- fix x assume A: "x  monomials_reducing_to R m P S - monomials_reducing_to R m (λm. k  P m) S"
        hence "x  monomials_reducing_to R m (λm. k  P m) S"
        by blast
         hence "(k  P x) = 𝟬" unfolding monomials_reducing_to_def
        by (metis (no_types, lifting) A DiffD1 complement_of_monomials_of mem_Collect_eq monomials_reducing_to_def)
        thus "monom_eval R (restrict_to_indices x S) g  (k  P x) = 𝟬"
        using monom_eval_car[of g]  assms(2) r_null by presburger
      qed
      have 02: "finite (monomials_reducing_to R m P S)"
        using assms(3) monomials_reducing_to_finite by blast
      have 04: "(λn. monom_eval R (restrict_to_indices n S) g  (k  P n))  monomials_reducing_to R m P S  carrier R"
        by (smt Pi_I assms(1) assms(2) assms(3) closed_fun_simp cring.axioms(1) is_cring m_closed monom_eval_car ring.Pring_cfs_closed)
      show ?thesis
        using 00 01 02 04
            finsum_eq_parition[of "monomials_reducing_to R m P S"
                                "(λn. monom_eval R (restrict_to_indices n S) g  (k  P n))"
                                "monomials_reducing_to R m (λm. k  P m) S"]
        by blast
    qed
    then have 1: " (nmonomials_reducing_to R m (λm. k  P m) S. monom_eval R (restrict_to_indices n S) g  (k  P n))
       = (nmonomials_reducing_to R m P S. k  (monom_eval R (restrict_to_indices n S) g  (P n)))      "
    proof-
      have "n. monom_eval R (restrict_to_indices n S) g  (k  P n) = k  (monom_eval R (restrict_to_indices n S) g  (P n))"
        by (metis Pring_cfs_closed assms(1) assms(2) assms(3) m_lcomm monom_eval_car)
      then show ?thesis
        using 0
        by presburger
    qed
    show " (nmonomials_reducing_to R m (λm. k  P m) S. monom_eval R (restrict_to_indices n S) g  (k  P n))
       = k (nmonomials_reducing_to R m P S.  monom_eval R (restrict_to_indices n S) g  (P n))      "
    proof-
      have "(λn. monom_eval R (restrict_to_indices n S) g  (P n))  (monomials_reducing_to R m P S)  carrier R"
        by (meson Pi_I Pring_cfs_closed assms(2) assms(3) m_closed monom_eval_car)
      then have "k  (imonomials_reducing_to R m P S. monom_eval R (restrict_to_indices i S) g  P i) =
      (imonomials_reducing_to R m P S. k  (monom_eval R (restrict_to_indices i S) g  P i))"
        using finsum_rdistr[of "monomials_reducing_to R m P S"
                                  k
                                  "(λn. monom_eval R (restrict_to_indices n S) g  (P n))"]
              assms  monomials_reducing_to_finite by blast
      then show ?thesis
        using 1
        by presburger
    qed
  qed
qed

lemma poly_eval_monomial:
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  shows "poly_eval R  S g (mset_to_IP R m)
           = poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (remove_indices m S))"
proof
  have 0: "monomials_of R (mset_to_IP R m) = {m}"
    using assms  monomials_of_mset_to_IP
    by blast

  fix x
  show "poly_eval R  S g (mset_to_IP R m) x =
         poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                           (mset_to_IP R (remove_indices m S)) x "
  proof(cases "x = (remove_indices m S)")
    case True
    have "monomials_reducing_to R x (mset_to_IP R m) S = {m}"
      unfolding monomials_reducing_to_def
      using True 0
      by auto
    then have "poly_eval R  S g (mset_to_IP R m) x = monom_eval R (restrict_to_indices m S) g  (mset_to_IP R m m)"
      unfolding poly_eval_def
      by (metis (mono_tags, lifting) assms(1) finsum_singleton monom_eval_car mset_to_IP_simp r_one)
    then show ?thesis
      by (metis True mset_to_IP_simp poly_scalar_mult_def)
  next
    case False
    then have "monomials_reducing_to R x (mset_to_IP R m) S = {}"
      unfolding monomials_reducing_to_def
      using 0
      by auto
    then have "poly_eval R S g (mset_to_IP R m) x = 𝟬"
      unfolding poly_eval_def
      by (metis finsum_empty)
    then show ?thesis
      using False
      by (metis assms(1) monom_eval_car mset_to_IP_simp' poly_scalar_mult_def r_null)
  qed
qed


lemma(in cring) poly_eval_monomial_closed:
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  assumes "set_mset m  I"
  shows "poly_eval R S g (mset_to_IP R m)  Pring_set R (I - S)"
proof-
  have "(mset_to_IP R (remove_indices m S))  Pring_set R (I - S)"
    using assms  mset_to_IP_closed[of "(remove_indices m S)" "I - S"]
    by (metis remove_indices_indices)
  then show ?thesis
    using assms poly_eval_monomial[of g S m  ]
          poly_scalar_mult_closed[of "(monom_eval R (restrict_to_indices m S) g)"
                                        "(mset_to_IP R (remove_indices m S))"]
    by (metis monom_eval_car)
qed

lemma poly_scalar_mult_iter:
  assumes "𝟭 𝟬"
  assumes "P  Pring_set R I"
  assumes "k  carrier R"
  assumes "n  carrier R"
  shows "poly_scalar_mult R k (poly_scalar_mult R n P) = poly_scalar_mult R (k  n) P"
  using assms
  unfolding poly_scalar_mult_def
  by (metis Pring_cfs_closed m_assoc)

lemma poly_scalar_mult_comm:
  assumes "𝟭 𝟬"
  assumes "P  Pring_set R I"
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "poly_scalar_mult R a (poly_scalar_mult R b P) = poly_scalar_mult R b (poly_scalar_mult R a P)"
  using assms  poly_scalar_mult_iter m_comm[of a b]
  by metis

lemma poly_eval_monomial_term:
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  assumes "set_mset m  I"
  assumes "k  carrier R"
  shows "poly_eval R  S g (poly_scalar_mult R k (mset_to_IP R m)) = poly_scalar_mult R (k(monom_eval R (restrict_to_indices m S) g))
                              (mset_to_IP R (remove_indices m S))"
proof-
  have 0: "poly_eval R  S g (poly_scalar_mult R k (mset_to_IP R m)) =
          poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m))"
    using assms poly_eval_scalar_mult[of k g "mset_to_IP R m" I S] mset_to_IP_closed
    by blast
  have 1: "poly_eval R  S g (poly_scalar_mult R k (mset_to_IP R m)) =
           poly_scalar_mult R k (poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (remove_indices m S))) "
    using 0 assms
    by (metis poly_eval_monomial)
  have 2: "mset_to_IP R (remove_indices m S)  Pring_set R I"
    using assms mset_to_IP_closed
    by (metis Diff_subset remove_indices_def restrict_to_indicesE subset_iff)
  have 3: "monom_eval R (restrict_to_indices m S) g  carrier R"
    using assms monom_eval_car by blast
  show ?thesis
    using 1 2 3 assms   poly_scalar_mult_iter[of  "mset_to_IP R (remove_indices m S)" I k "(monom_eval R (restrict_to_indices m S) g)"]
  by presburger
qed

lemma poly_eval_monomial_term_closed:
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  assumes "set_mset m  I"
  assumes "k  carrier R"
  shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m))  Pring_set R (I - S)"
proof-
  have "(mset_to_IP R (remove_indices m S))  Pring_set R (I - S)"
    using assms
    by (meson remove_indices_indices mset_to_IP_closed)
  then show ?thesis
    using assms poly_eval_monomial_term remove_indices_indices[of m I S]
    by (metis cring.cring_simprules(5) cring.monom_eval_car is_cring poly_scalar_mult_closed)
qed

lemma finsum_split:
  assumes "finite S"
  assumes "f  S  carrier R"
  assumes "g  S  carrier R"
  assumes "k  carrier  R"
  assumes "c  S"
  assumes "s. s  S  s  c  f s = g s"
  assumes "g c = f c  k"
  shows "finsum R g S = k  finsum R f S"
proof-
  have 0: "finsum R f S = f c  finsum R f (S - {c})"
  proof-
    have "f  S - {c}  carrier R"
      by (metis Pi_split_insert_domain assms(2) assms(5) insert_Diff)
    then show ?thesis
    using assms finsum_insert[of "S - {c}" c f]
    by (metis DiffD2 Pi_iff finite_Diff insert_Diff singletonI)
  qed
  have 1: "finsum R g S = g c  finsum R g (S - {c})"
  proof-
    have "g  S - {c}  carrier R"
      by (metis Pi_split_insert_domain  assms(3) assms(5) insert_Diff)
    then show ?thesis
    using assms finsum_insert[of "S - {c}" c g]
    by (metis DiffD2 Pi_iff finite_Diff insert_Diff singletonI)
  qed
  have "finsum R f (S - {c}) = finsum R g (S - {c})"
    using assms Diff_iff Pi_split_insert_domain finsum_cong'[of "S- {c}" "S - {c}" g f]
      insert_Diff singletonI
    by blast
  then have "finsum R g S = f c  k  finsum R g (S - {c})"
    using assms  finsum R g S = g c  finsum R g (S - {c})
    by presburger
  then have 1: "finsum R g S = f c  k  finsum R f (S - {c})"
    using finsum R f (S - {c}) = finsum R g (S - {c}) by presburger
  have "finsum R g S = k  ( f c  finsum R f (S - {c}))"
  proof-
    have "f c  carrier R"
      by (metis PiE assms(2) assms(5))
    have "finsum R f (S - {c})  carrier R"
      by (metis Pi_split_insert_domain assms(2) assms(5) finsum_closed insert_Diff)
    then show ?thesis using assms(4) 1
      using f c  carrier R add.m_assoc add.m_lcomm
      by presburger
  qed
  then show ?thesis
    using "0"
    by presburger
qed

lemma poly_monom_induction:
assumes "P (indexed_const 𝟬)"
assumes "m k. set_mset m  I  k  carrier R  P (poly_scalar_mult R k (mset_to_IP R m))"
assumes "Q m k. Q  Pring_set R I  (P Q)  set_mset m  I  k  carrier R P (Q  (poly_scalar_mult R k (mset_to_IP R m)))"
shows "Q. Q  Pring_set R I  P Q"
proof-
  have 0: "Ms. finite Ms  ( Q  Pring_set R I. monomials_of R Q = Ms  P Q)"
  proof-
    fix Ms
    show " finite Ms  ( Q  Pring_set R I. monomials_of R Q = Ms  P Q)"
    proof(erule finite.induct)
      show "QPring_set R I. monomials_of R Q = {}  P Q"
      proof
        fix Q
        assume "Q  Pring_set R I"
        show "monomials_of R Q = {}  P Q "
          using assms
          by (metis Q  Pring_set R I card_0_eq monomials_finite monomials_of_card_zero)
      qed
      show "A a. finite A  QPring_set R I. monomials_of R Q = A  P Q 
                    QPring_set R I. monomials_of R Q = insert a A  P Q"
      proof
        fix A :: "'c monomial set" fix  a fix  Q
        assume A0: "finite A"
        assume A1: "QPring_set R I. monomials_of R Q = A  P Q"
        assume A2: " Q  Pring_set R I "
        show "monomials_of R Q = insert a A  P Q"
        proof
          assume A3: "monomials_of R Q = insert a A"
          show "P Q"
            apply(cases "a  A")
             apply (metis A1 A2 A3 insert_absorb)
          proof-
            assume A4: "a  A"
            show "P Q"
            proof-
              have A5: "set_mset a  I"
                by (metis A2 A3 insert_iff mset_to_IP_indices)
              have A6: "set_mset a  I  Q a  carrier R"
                using A2 A5 Pring_cfs_closed by blast
              obtain Q' where Q'_def: "Q' = remove_monom R a Q"
                by simp
              then have "Q = Q'  poly_scalar_mult R (Q a) (mset_to_IP R a)"
                using A2 cring.remove_monom_eq is_cring by blast
              then show ?thesis using A6 assms
                by (metis A1 A2 A3 A4 Diff_empty Diff_insert0 Q'_def insert_Diff1
                    remove_monom_closed remove_monom_monomials singletonI)
            qed
          qed
        qed
      qed
    qed
  qed
  show "Q. Q  Pring_set R I  P Q"
  proof-
    fix Q
    assume "Q  Pring_set R I"
    show "P Q"
      using 0[of "monomials_of R Q"]  Q  Pring_set R I monomials_finite
      by blast
  qed
qed

lemma Pring_car_induct:
  assumes "q  carrier (Pring R I)"
  assumes "P 𝟬Pring R I⇙"
  assumes "m k. set_mset m  I  k  carrier R  P (k Pring R I(mset_to_IP R m))"
  assumes "Q m k. Q  carrier (Pring R I)  (P Q)  set_mset m  I  k  carrier R
             P (Q  (k Pring R I(mset_to_IP R m)))"
  shows "P q"
  using poly_monom_induction[of P I q] assms Pring_smult[of I] Pring_car[of I] Pring_zero
  by metis

lemma poly_monom_induction2:
assumes "P (indexed_const 𝟬)"
assumes "m k. set_mset m  I  k  carrier R  P (poly_scalar_mult R k (mset_to_IP R m))"
assumes "Q m k. Q  Pring_set R I  (P Q)  set_mset m  I  k  carrier R  P (Q  (poly_scalar_mult R k (mset_to_IP R m)))"
assumes "Q   Pring_set R I"
shows "P Q"
  using assms(1) assms(2) assms(3) assms(4) poly_monom_induction by blast

lemma poly_monom_induction3:
assumes "Q   Pring_set R I"
assumes "P (indexed_const 𝟬)"
assumes "m k. set_mset m  I  k  carrier R  P (poly_scalar_mult R k (mset_to_IP R m))"
assumes "p q. p  Pring_set R I  (P p)  q  Pring_set R I  (P q)   P (p  q)"
shows "P Q"
  apply(rule poly_monom_induction2[of _  I])
  using assms(2) apply blast
  using assms(3) apply blast
  apply (meson assms(3) assms(4) mset_to_IP_closed poly_scalar_mult_closed)
  using assms(1) by blast

lemma Pring_car_induct':
assumes "Q  carrier (Pring R I)"
assumes "P 𝟬Pring R I⇙"
assumes "m k. set_mset m  I  k  carrier R  P (k Pring R Imset_to_IP R m)"
assumes "p q. p  carrier (Pring R I)  (P p)  q  carrier (Pring R I)  (P q)   P (p Pring R Iq)"
shows "P Q"
  using poly_monom_induction3[of Q I P] assms Pring_smult Pring_add Pring_zero Pring_car
  by metis

lemma poly_eval_mono:
  assumes "P  Pring_set R I"
  assumes "closed_fun R g"
  assumes "finite F"
  assumes "monomials_reducing_to R m P S  F"
  assumes "n. n  F  remove_indices n S = m"
  shows "poly_eval R S g P m = (n F. monom_eval R (restrict_to_indices n S) g  P n)"
proof-
  have 0: "(n F. monom_eval R (restrict_to_indices n S) g  P n)=
          (n F - (monomials_reducing_to R m P S). monom_eval R (restrict_to_indices n S) g  P n)  poly_eval R S g P m"
  proof-
    have 00: "  (λn. monom_eval R (restrict_to_indices n S) g  P n)  F  carrier R"
      by (meson Pi_I Pring_cfs_closed assms(1) assms(2) m_closed monom_eval_car)
    have 01: "monomials_reducing_to R m P S  F"
      by (simp add: assms(4))
    have 02: "(nF. monom_eval R (restrict_to_indices n S) g  P n) =
    (nmonomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g  P n) 
    (nF - monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g  P n)"
      using "00" "01" assms(3) finsum_partition by blast
    have 03: " (nF - monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g  P n) carrier R"
      by (metis (mono_tags, lifting) "00" DiffD1 PiE Pi_I finsum_closed)
    have "    (nmonomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g  P n)  carrier R"
    proof -
      have "(λm. monom_eval R (restrict_to_indices m S) g  P m)  monomials_reducing_to R m P S  carrier R"
        using "00" "01" by blast
      then show ?thesis
        using finsum_closed by blast
    qed
    then show ?thesis
      unfolding poly_eval_def
      using  00 01 02 03  add.m_comm
      by presburger
  qed
  have "(n F - (monomials_reducing_to R m P S). monom_eval R (restrict_to_indices n S) g  P n) = 𝟬"
  proof-
    have "n. n  F - (monomials_reducing_to R m P S)   monom_eval R (restrict_to_indices n S) g  P n = 𝟬"
    proof-
      fix n
      assume A: "n   F - (monomials_reducing_to R m P S)"
      have "n  monomials_of R P"
      proof
        assume "n  monomials_of R P"
        then have "n  (monomials_reducing_to R m P S)"
          unfolding monomials_reducing_to_def
          using assms(5) A
          by blast
        then show False using A by blast
      qed
      then show " monom_eval R (restrict_to_indices n S) g  P n = 𝟬"
        by (metis assms(2) monom_eval_car complement_of_monomials_of r_null)
    qed
    then show ?thesis
      by (meson add.finprod_one_eqI)
  qed
  then have 1: "(nF. monom_eval R (restrict_to_indices n S) g  P n) =
    𝟬  poly_eval R S g P m"
    using 0  Pi_I Pring_cfs_closed add.r_cancel_one assms(1) assms(2)
    by presburger
  have " (λn. monom_eval R (restrict_to_indices n S) g  P n)  monomials_reducing_to R m P S  carrier R"
    by (meson Pi_I Pring_cfs_closed assms(1) assms(2) m_closed monom_eval_car)
  hence "poly_eval R S g P m  carrier R"
    using assms
    unfolding poly_eval_def
    using finsum_closed[of "λ n. monom_eval R (restrict_to_indices n S) g  P n"
                          "monomials_reducing_to R m P S"]
    by (meson Pi_I Pring_cfs_closed m_closed monom_eval_car)
  then have "(nF. monom_eval R (restrict_to_indices n S) g  P n) = poly_eval R S g P m"
    using "1" add.r_cancel_one zero_closed
    by presburger
  then show ?thesis
    by presburger
qed

lemma finsum_group:
  assumes "n. f n  carrier R"
  assumes "n. g n   carrier R"
  shows "finite S  finsum R f S  finsum R g S = finsum R (λn. f n  g n) S "
  apply(erule finite.induct)
  apply (metis finsum_empty r_zero zero_closed)
proof-
  fix A :: "'c set "
  fix a
  assume A0: "finite A"
  assume A1: "finsum R f A  finsum R g A = (nA. f n  g n)"
  show "finsum R f (insert a A)  finsum R g (insert a A) = (ninsert a A. f n  g n)"
  proof(cases "a  A")
    case True
    then show ?thesis
      by (metis A1 insert_absorb)
  next
    case False
    have LHS: "finsum R f (insert a A)  finsum R g (insert a A) =
                  (f a  finsum R f A)  (g a  finsum R g A)"
      using assms  finsum_insert[of A a f] finsum_insert[of A a g]
      by (metis A0 False Pi_I)
    have F0: "(λn. f n  g n)  A  carrier R"
      using assms
      by blast
    have F1: "(f a  g a)  carrier  R"
      using assms
      by blast
    have RHS: " (ninsert a A. f n  g n) = (f a  g a)  (nA. f n  g n)"
      using F0 F1 assms finsum_insert[of A a " (λn. f n  g n)"] False A0
      by blast
    have F2: " f a  finsum R f A  (g a  finsum R g A) = (f a  g a)  (finsum R f A   finsum R g A)"
    proof-
      have F20: "f a  carrier R"
        using assms(1) by blast
      have F21: "g a  carrier R"
        using assms(2) by blast
      have F22: "finsum R f A  carrier R"
        by (metis Pi_iff assms(1) finsum_closed)
      have F23: "finsum R g A  carrier R"
        by (metis Pi_I assms(2) finsum_closed)
      show ?thesis using F21 F20 F22 F23
        using add.m_assoc add.m_closed add.m_lcomm
        by presburger
    qed
    show ?thesis
      using RHS LHS assms A1 F2
      by presburger
  qed
qed

lemma poly_eval_add:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  assumes "closed_fun R g"
  shows "poly_eval R S g (P  Q)   = poly_eval R S g P   poly_eval R S g Q "
proof
  fix m
  show "poly_eval R S g (P  Q)  m = (poly_eval R S g P  poly_eval R S g Q) m"
  proof-
    obtain F where F_def: "F = monomials_reducing_to R m (P  Q) S  monomials_reducing_to R m P S 
                              monomials_reducing_to R m Q S"
      by simp
    have 0: "finite F"
    proof-
      have 00: "finite (monomials_reducing_to R m (P  Q) S)"
        using assms
        by (meson finite_subset monomials_finite monomials_of_add_finite monomials_reducing_to_subset)
      have 01: "finite (monomials_reducing_to R m P S)"
        using assms(1) monomials_reducing_to_finite by blast
      have 02: "finite (monomials_reducing_to R m Q S)"
        using assms(2) monomials_reducing_to_finite by blast
      show ?thesis
        using F_def 00 01 02
        by blast
    qed
    have 1: "n. n  F  remove_indices n S = m"
    proof-
      fix n
      assume A: "n  F"
      show "remove_indices n S = m"
        using F_def
        unfolding monomials_reducing_to_def
        using A
        by blast
    qed
    have 2: "poly_eval R S g (P  Q)  m  = (n F. monom_eval R (restrict_to_indices n S) g  (P  Q) n)"
      using assms 0 1 poly_eval_mono[of "P  Q" I g F m] F_def indexed_pset.indexed_padd
      by blast
    have 3: "poly_eval R S g P m  = (n F. monom_eval R (restrict_to_indices n S) g  P n)"
      using assms 0 1 poly_eval_mono[of "P" I g F m] F_def indexed_pset.indexed_padd
      by blast
    have 4: "poly_eval R S g Q m  = (n F. monom_eval R (restrict_to_indices n S) g  Q n)"
      using assms 0 1 poly_eval_mono[of "Q" I g F m] F_def indexed_pset.indexed_padd
      by blast
    have 5: "poly_eval R S g P m  poly_eval R S g Q m = (n F. monom_eval R (restrict_to_indices n S) g  P n
                                                           monom_eval R (restrict_to_indices n S) g  Q n)"
    proof-
      have 50: "(λn. monom_eval R (restrict_to_indices n S) g  P n)  F  carrier R"
        by (meson Pi_I Pring_cfs_closed assms(1) assms(3) m_closed monom_eval_car)
      have 51: "(λn. monom_eval R (restrict_to_indices n S) g  Q n)  F  carrier R"
        by (meson Pi_I Pring_cfs_closed assms(2) assms(3) m_closed monom_eval_car)
      then show ?thesis
        using 0 2 3 50 finsum_group[of "(λn. monom_eval R (restrict_to_indices n S) g  P n)"
                                       "(λn. monom_eval R (restrict_to_indices n S) g  Q n)" F]
        by (metis (mono_tags, lifting) "4" Pring_cfs_closed assms(1) assms(2) assms(3) m_closed monom_eval_car)
    qed
    have 6: "poly_eval R S g P m  poly_eval R S g Q m
           = (n F. monom_eval R (restrict_to_indices n S) g ( P n   Q n))"
    proof-
      have 0 : "(λn. monom_eval R (restrict_to_indices n S) g  P n  monom_eval R (restrict_to_indices n S) g  Q n)  F  carrier R"
        apply(rule Pi_I)
        by (meson Pring_cfs_closed add.m_closed assms(1) assms(2) assms(3) m_closed monom_eval_car)
      have 1: "(λn. monom_eval R (restrict_to_indices n S) g  (P n  Q n))  F  carrier R"
        apply(rule Pi_I)
        by (meson Pring_cfs_closed add.m_closed assms(1) assms(2) assms(3) m_closed monom_eval_car)
      have "n. n  F   monom_eval R (restrict_to_indices n S) g ( P n   Q n) = monom_eval R (restrict_to_indices n S) g  P n
                                                           monom_eval R (restrict_to_indices n S) g  Q n"
        using assms Pring_cfs_closed cring.monom_eval_car is_cring r_distr
        by metis
      then have "(λxF. monom_eval R (restrict_to_indices x S) g  P x  monom_eval R (restrict_to_indices x S) g  Q x)
                = (λxF. monom_eval R (restrict_to_indices x S) g  (P x  Q x))"
        by (metis (no_types, lifting) restrict_ext)
      then show ?thesis
        using 5 finsum_eq[of "(λn. monom_eval R (restrict_to_indices n S) g  P n   monom_eval R (restrict_to_indices n S) g  Q n)"
                            F "(λn. monom_eval R (restrict_to_indices n S) g ( P n   Q n))" ] 0 1
        by presburger
    qed
    have 7: "monomials_reducing_to R m (P  Q) S  F"
      using F_def
      by blast
    have 8: "poly_eval R S g (P  Q)  m  = (nF. monom_eval R (restrict_to_indices n S) g  (P  Q) n)"
      using 7 0 1 "2"
      by blast
    obtain f where f_def: "f = (λn. monom_eval R (restrict_to_indices n S) g  (P  Q) n)"
      by blast
    obtain h where h_def: "h = (λn.  monom_eval R (restrict_to_indices n S) g ( P n   Q n))"
      by blast
    have 9: "f  F  carrier R"
      using f_def
      by (metis (mono_tags, lifting) Pi_I Pring_cfs_closed add.m_closed
          assms(1) assms(2) assms(3) indexed_padd_def m_closed monom_eval_car)
    have 10: "h  F  carrier R"
      using h_def
      by (metis (mono_tags, lifting) "9" Pi_cong f_def indexed_padd_def)
    have 11: "restrict f F = restrict h F"
      using f_def h_def
      by (metis indexed_padd_def)
    have "finsum R f F = finsum R h F"
      using 9 10 11 finsum_eq[of f F h]
      by blast
    then show ?thesis
      using f_def h_def
      by (metis (no_types, lifting) "6" "8"  indexed_padd_def)
  qed
qed

lemma poly_eval_Pring_add:
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  assumes "closed_fun R g"
  shows "poly_eval R S g (P Pring R IQ)   = poly_eval R S g P Pring R Ipoly_eval R S g Q "
  using assms poly_eval_add[of P I Q g S]
  by (metis Pring_add Pring_car)

text‹Closure of partial evaluation maps:›
lemma(in cring) poly_eval_closed:
  assumes "closed_fun R g"
  assumes "P  Pring_set R I"
  shows "poly_eval R S g P  Pring_set R (I - S)"
proof-
  obtain Pr where Pr_def[simp]: "Pr = (λQ. poly_eval R S g Q  Pring_set R (I - S))"
    by blast
  have "Pr P"
    apply(rule poly_monom_induction2[of _ I _ ])
     apply (metis Pr_def indexed_pset.indexed_const poly_eval_constant zero_closed)
       apply (metis Pr_def assms(1) indexed_pset.indexed_const mset_to_IP_closed
        poly_eval_constant poly_eval_monomial_term_closed poly_scalar_mult_zero r_null r_one)
  proof-
    show "P  Pring_set R I" using assms by blast
    show "Q m k. Q  Pring_set R I  Pr Q  set_mset m  I  k  carrier R Pr (Q  poly_scalar_mult R k (mset_to_IP R m))"
    proof-
      fix Q
      fix m
      fix k
      assume A: "Q  Pring_set R I  Pr Q  set_mset m  I k  carrier R"
      then have 0: "poly_eval R S g Q  Pring_set R (I - S)"
        using Pr_def by blast
      have 1: "poly_scalar_mult R k (mset_to_IP R m)  Pring_set R I"
        using  A mset_to_IP_closed poly_scalar_mult_closed
        by blast
      have  "poly_eval R  S g (Q  poly_scalar_mult R k (mset_to_IP R m)) =
     poly_eval R S g Q  poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m))   "
        using assms poly_eval_add[of Q I " (poly_scalar_mult R k (mset_to_IP R m))" g]  "1" A
        by blast
      then show "Pr (Q  poly_scalar_mult R k (mset_to_IP R m))"
        using Pr_def
        by (metis "1" A assms(1) indexed_pset.indexed_padd poly_eval_monomial_term_closed poly_scalar_mult_one poly_scalar_mult_zero)
    qed
  qed
  then show ?thesis
    using Pr_def
    by blast
qed

lemma poly_scalar_mult_indexed_pmult:
  assumes "P  Pring_set R I"
  assumes "k  carrier R"
  shows " poly_scalar_mult R k (P  i) = (poly_scalar_mult R k P)  i"
proof-
  have 0: "mset_to_IP R {#i#}  Pring_set R (I  {i})"
    by (metis Un_upper2 mset_to_IP_closed set_mset_add_mset_insert set_mset_empty)
  have 1: "P  Pring_set R (I  {i})"
    using Pring_carrier_subset Un_upper1 assms(1) by blast
  show ?thesis
    using 0 1  poly_scalar_mult_times[of "mset_to_IP R {#i#}" "I  {i}" P  k]
        poly_index_mult[of P "I  {i}" i] assms
    by (metis Un_upper2 insert_subset poly_index_mult poly_scalar_mult_closed)
qed

lemma remove_indices_add_mset:
  assumes "i  S"
  shows "remove_indices (add_mset i m) S = add_mset i (remove_indices m S)"
  apply(induction m)
  apply (smt assms empty_eq_union remove_indices_decomp restrict_to_indicesE single_is_union union_single_eq_member)
  by (metis assms multi_union_self_other_eq remove_indices_decomp restrict_to_indices_add_element union_mset_add_mset_right)

lemma poly_eval_monom_insert:
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  assumes "i  S"
  shows "poly_eval R  S g (mset_to_IP R (add_mset i m))
           = poly_scalar_mult R (g i)
                    (poly_eval R  S g (mset_to_IP R m))"
proof-
  have 0: "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (monom_eval R (restrict_to_indices (add_mset i m) S) g)
                              (mset_to_IP R (remove_indices (add_mset i m) S))"
    using assms(1) assms(2) poly_eval_monomial by blast
  have 1: "(mset_to_IP R (remove_indices (add_mset i m) S)) =
                           (mset_to_IP R (remove_indices m S))"
    using assms
    by (metis (full_types) Diff_iff insert_Diff1 remove_indices_def restrict_to_indices_add_element set_mset_add_mset_insert)
  have 2: "(monom_eval R (restrict_to_indices (add_mset i m) S) g) =
            (g i)  ((monom_eval R (restrict_to_indices m S) g))"
    using assms
    by (metis monom_eval_add restrict_to_indices_add_element)
  have 3: "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R ((g i)  ((monom_eval R (restrict_to_indices m S) g)))
                              (mset_to_IP R (remove_indices m S))"
    using 0 1 2 assms
    by presburger
  hence "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (g i)
                  (poly_scalar_mult R ((monom_eval R (restrict_to_indices m S) g))
                              (mset_to_IP R (remove_indices m S)))"
    using assms poly_scalar_mult_iter[of "(mset_to_IP R (remove_indices m S))"
                                          UNIV "g i" "monom_eval R (restrict_to_indices m S) g"]
          mset_to_IP_closed[of "remove_indices m S" UNIV]
    by (metis PiE UNIV_I monom_eval_car subsetI)
  thus ?thesis using assms
    by (metis poly_eval_monomial)
qed

lemma poly_eval_monom_insert':
  assumes "closed_fun R g"
  assumes "𝟭 𝟬"
  assumes "i  S"
  shows "poly_eval R  S g (mset_to_IP R (add_mset i m))
           =  (poly_eval R  S g (mset_to_IP R m))  i"
proof-
  have 0: "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (monom_eval R (restrict_to_indices (add_mset i m) S) g)
                              (mset_to_IP R (remove_indices (add_mset i m) S))"
    using assms(1) assms(2) poly_eval_monomial by blast
  hence "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (remove_indices (add_mset i m) S))"
    by (metis assms(3) restrict_to_indices_add_element)
  hence "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (add_mset i (remove_indices m S)))"
    by (metis assms(3) remove_indices_add_mset)
  hence "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (remove_indices m S)  i)"
    by (metis monom_add_mset)
  hence "poly_eval R  S g (mset_to_IP R (add_mset i m)) =
                (poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g)
                              (mset_to_IP R (remove_indices m S)))  i"
    by (metis assms(1) cring.monom_eval_car is_cring local.ring_axioms
        poly_scalar_mult_indexed_pmult ring.mset_to_IP_closed set_eq_subset)
  thus ?thesis
    by (metis assms(1) assms(2) poly_eval_monomial)
qed

lemma poly_eval_indexed_pmult_monomial:
  assumes "closed_fun R g"
  assumes "k  carrier R"
  assumes "i  S"
  assumes "𝟭  𝟬"
  shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i) =
          poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)))"
proof-
  have 0: "poly_scalar_mult R k (mset_to_IP R m)  i =
              poly_scalar_mult R k (mset_to_IP R (add_mset i m ))"
    using monom_add_mset[of i m] poly_scalar_mult_indexed_pmult
    by (metis (no_types, opaque_lifting) assms(2) local.ring_axioms  ring.mset_to_IP_closed subsetD subset_refl)
  hence 1: "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i)  =
             poly_scalar_mult R k (poly_eval R S g (mset_to_IP R (add_mset i m )))"
    by (metis assms(1) assms(2) cring.poly_eval_scalar_mult is_cring local.ring_axioms ring.mset_to_IP_closed subsetI)
  have 2: "poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)))
            = poly_scalar_mult R (g i)
              (poly_scalar_mult R k  (poly_eval R S g (mset_to_IP R m))) "
    by (smt assms(1) assms(2) local.ring_axioms poly_eval_scalar_mult ring.mset_to_IP_closed subsetD subset_refl)
  hence 3: "poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)))
            = poly_scalar_mult R k
              (poly_scalar_mult R (g i) (poly_eval R S g (mset_to_IP R m))) "
    using assms poly_scalar_mult_comm[of "(poly_eval R S g (mset_to_IP R m))" UNIV]
          poly_eval_closed[of g "(mset_to_IP R m)" UNIV]
    by (metis UNIV_I closed_fun_simp cring.poly_scalar_mult_comm
        is_cring local.ring_axioms ring.mset_to_IP_closed subsetI)
  have 4: "(poly_eval R S g (mset_to_IP R (add_mset i m))) =  (poly_scalar_mult R (g i) (poly_eval R S g (mset_to_IP R m)))"
    using assms poly_eval_monom_insert[of g i S m]
    by blast
  thus ?thesis
    using 1 3 poly_scalar_mult_comm assms
    by presburger
qed

lemma poly_eval_indexed_pmult_monomial':
  assumes "closed_fun R g"
  assumes "k  carrier R"
  assumes "i  S"
  assumes "𝟭  𝟬"
  shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i) =
         (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)))  i"
proof-
  have "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i) =
      poly_scalar_mult R k (poly_eval R S g ( (mset_to_IP R m)  i))"
    using poly_eval_scalar_mult[of k g]
    by (metis UNIV_I assms(1) assms(2) local.ring_axioms poly_scalar_mult_indexed_pmult ring.monom_add_mset ring.mset_to_IP_closed subsetI)
  hence "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i) =
      poly_scalar_mult R k (poly_eval R S g (mset_to_IP R (add_mset i m)))"
    by (simp add: monom_add_mset)
  hence "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)  i) =
      (poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m)))  i"
    using poly_eval_monom_insert'[of g i S m]
    by (smt assms(1) assms(2) assms(3) assms(4) poly_eval_monomial_closed poly_scalar_mult_indexed_pmult subsetD subset_refl)
  thus ?thesis using assms poly_eval_scalar_mult[of k g _ UNIV S]
    by (metis UNIV_I mset_to_IP_closed subsetI)
qed

lemma indexed_pmult_add:
  assumes "p  Pring_set R I"
  assumes "q  Pring_set R I"
  shows "p  q  i = (p  i)  (q  i)"
  using assms poly_index_mult[of _ "I  {i}"]
  by (smt Pring_carrier_subset Set.basic_monos(1) Un_upper1 Un_upper2 cring.axioms(1)
      insert_subset is_cring mk_disjoint_insert mset_to_IP_closed ring.P_ring_ldistr
      ring.indexed_pset.intros(2) ring.indexed_pset_in_carrier set_mset_add_mset_insert
      set_mset_empty)

lemma poly_eval_indexed_pmult:
  assumes "P  Pring_set R I"
  assumes "closed_fun R g"
  shows "poly_eval R S g (P  i)   = (if i  S then poly_scalar_mult R (g i) (poly_eval R S g P) else (poly_eval R S g P) i )"
proof(cases "i  S")
  case True
  have "poly_eval R S g (P  i) = poly_scalar_mult R (g i) (poly_eval R S g P)"
  apply(rule poly_monom_induction3[of P I])
    using assms apply blast
      apply (metis PiE UNIV_I assms(2) indexed_pmult_zero poly_eval_constant poly_scalar_mult_const r_null zero_closed)
        apply (meson True assms(2) poly_eval_indexed_pmult_monomial)
          apply (smt PiE True UNIV_I assms(2) genideal_one genideal_zero indexed_pmult_zero mset_to_IP_closed poly_eval_constant poly_eval_indexed_pmult_monomial poly_scalar_mult_one poly_scalar_mult_zero singletonD)
  proof- fix p q assume A:
          "p  Pring_set R I"
          "poly_eval R S g (p  i) = poly_scalar_mult R (g i) (poly_eval R S g p)"
          "q  Pring_set R I"
          "poly_eval R S g (q  i) = poly_scalar_mult R (g i) (poly_eval R S g q)"
    have "poly_eval R S g (p  q  i) = poly_eval R S g (p  i)  poly_eval R S g (q  i)"
      using assms poly_eval_add[of "p  i" "I  {i}" "q  i" g S]
            indexed_pmult_add[of p  I q i]
      by (smt A(1) A(3) Pring_carrier_subset Un_insert_right Un_upper1 indexed_pset.indexed_pmult insert_iff insert_subset mk_disjoint_insert)
    hence "poly_eval R S g (p  q  i) = poly_scalar_mult R (g i) (poly_eval R S g p) 
                                           poly_scalar_mult R (g i) (poly_eval R S g q)"
      using A
      by presburger
    hence "poly_eval R S g (p  q  i) = poly_scalar_mult R (g i)
                                            ((poly_eval R S g p)  (poly_eval R S g q))"
      using Pring_smult poly_eval_closed[of g] A Pring_add Pring_car  poly_eval R S g (p  q  i) = poly_eval R S g (p  i)  poly_eval R S g (q  i) assms(1) assms(2) closed_fun_simp
            Pring_smult_r_distr[of "g i" "poly_eval R S g p" _ "poly_eval R S g q"] Pring_add Pring_car
      by metis
    thus "poly_eval R S g (p  q  i) = poly_scalar_mult R (g i) (poly_eval R S g (p  q))"
      by (metis A(1) A(3) assms(2) poly_eval_add)
  qed
  then show ?thesis
    using True by presburger
next
  case False
  have "poly_eval R S g (P  i) = (poly_eval R S g P)  i"
  apply(rule poly_monom_induction3[of P I])
       apply (simp add: assms(1))
        apply (metis indexed_pmult_zero poly_eval_constant zero_closed)
          apply (metis False assms(2) indexed_pmult_zero inv_unique l_null  mset_to_IP_closed
           one_closed one_mset_to_IP poly_eval_constant  poly_eval_indexed_pmult_monomial'
           poly_scalar_mult_zero  zero_closed)
  proof-fix p q assume A:
          "p  Pring_set R I"
          "poly_eval R S g (p  i) = poly_eval R S g p  i"
          "q  Pring_set R I"
          "poly_eval R S g (q  i) = poly_eval R S g q  i"
    have "poly_eval R S g (p  q  i) = poly_eval R S g (p  i)  poly_eval R S g (q  i)"
      using assms poly_eval_add[of "p  i" "I  {i}" "q  i" g S]
            indexed_pmult_add[of p  I q i]
      by (smt A(1) A(3) Pring_carrier_subset Un_insert_right Un_upper1 indexed_pset.indexed_pmult insert_iff insert_subset mk_disjoint_insert)
    thus "poly_eval R S g (p  q  i) = poly_eval R S g (p  q)  i"
      by (metis A(1) A(2) A(3) A(4) assms(2) indexed_pmult_add poly_eval_add poly_eval_closed)
  qed
  then show ?thesis
    by (simp add: False)
qed

lemma poly_eval_index:
  assumes "𝟭 𝟬"
  assumes "closed_fun R g"
  shows "poly_eval R S g (mset_to_IP R {#i#})= (if i  S then (indexed_const (g i)) else mset_to_IP R {#i#})"
proof-
  have 0: "poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (monom_eval R (restrict_to_indices {#i#} S) g)
                              (mset_to_IP R (remove_indices {#i#} S))"
    using poly_eval_monomial[of g S "{#i#}" ] assms(1) assms(2) by blast
  show ?thesis proof(cases "i  S")
    case True
    then have T0: "(restrict_to_indices {#i#} S) =  {#i#}"
      by (metis restrict_to_indices_add_element restrict_to_indices_submultiset add_mset_subseteq_single_iff)
    then have T1: "(monom_eval R (restrict_to_indices {#i#} S) g) = (monom_eval R {#i#} g)"
      by presburger
    then have "(monom_eval R (restrict_to_indices {#i#} S) g)  = g i   monom_eval R {#} g"
      using assms monom_eval_add[of g i "{#}"]
      by presburger
    then have T2: "(monom_eval R (restrict_to_indices {#i#} S) g)  = g i"
      unfolding monom_eval_def
      using T0
      by (metis PiE UNIV_I assms(2) fold_mset_empty r_one)
    have T3: "(remove_indices {#i#} S) = {#}"
      by (metis Diff_iff remove_indices_indices restrict_to_indicesE
          T0 multiset_cases subset_iff union_single_eq_member)
    then have T4: " (mset_to_IP R (remove_indices {#i#} S)) = indexed_const 𝟭"
      by (metis one_mset_to_IP)
    have T5: "poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (g i) (indexed_const 𝟭)"
      using "0" T2 T4
      by presburger
    then show ?thesis using True  poly_scalar_mult_const
      by (metis T2 assms(2) monom_eval_car one_closed r_one)
  next
    case False
    have F0: "(restrict_to_indices {#i#} S) = {#}"
      using False restrict_to_indices_def
      by (metis restrict_to_indices_add_element filter_empty_mset)
    have F1: "(monom_eval R (restrict_to_indices {#i#} S) g) = 𝟭"
      using F0
      unfolding monom_eval_def
      by (metis fold_mset_empty)
    have F2: "(remove_indices {#i#} S) = {#i#}"
      using False
      by (metis Diff_iff remove_indices_def restrict_to_indices_add_element
          restrict_to_indices_def filter_empty_mset set_mset_single singletonI)
    have F3: "(mset_to_IP R (remove_indices {#i#} S)) = mset_to_IP R {#i#}"
      by (simp add: F2)
    have F4: " poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R 𝟭 (mset_to_IP R {#i#})"
      using "0" F1 F3
      by presburger
    show ?thesis using False
      by (metis F4 mset_to_IP_closed poly_scalar_mult_one subset_iff)
  qed
qed

lemma poly_eval_indexed_pmult':
  assumes "P  Pring_set R I"
  assumes "closed_fun R g"
  assumes "i  I"
  shows "poly_eval R  S g (P p  (mset_to_IP R {#i#}))  =  poly_eval R S g P  p  poly_eval R S g (mset_to_IP R {#i#})"
proof(cases "i  S")
  case True
  have "(P p mset_to_IP R {#i#})=(P  i) "
    using assms poly_index_mult
    by metis
  then have 0: " poly_eval R  S g (P p mset_to_IP R {#i#}) = poly_scalar_mult R (g i) (poly_eval R S g P)"
    using True assms poly_eval_indexed_pmult[of P I g S i]
    by presburger
  have 1: "poly_eval R S g (mset_to_IP R {#i#}) = indexed_const (g i)"
    using assms True
    by (smt PiE UNIV_I genideal_one genideal_zero indexed_pmult_zero monom_add_mset one_mset_to_IP poly_eval_constant poly_eval_index singletonD)
  then have "poly_eval R S g P  p  poly_eval R S g (mset_to_IP R {#i#})=  poly_eval R S g Pp indexed_const (g i) "
    by presburger
  then have "poly_eval R S g P  p  poly_eval R S g (mset_to_IP R {#i#})= indexed_const (g i) p  poly_eval R S g P"
    using assms P_ring_mult_comm[of "indexed_const (g i)" "poly_eval R S g P"]
    unfolding carrier_coeff_def
    by (metis "1" Pring_cfs_closed cring.closed_fun_simp indexed_pset.indexed_const is_cring poly_eval_closed)
  then have "poly_eval R S g P  p  poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (g i)  (poly_eval R S g P)"
    using assms
    by (metis "0" "1" P p mset_to_IP R {#i#} = P  i cring.closed_fun_simp is_cring poly_eval_closed poly_scalar_mult_eq)
  then show ?thesis using 0
    by presburger
next
  case False
  then have 0: "poly_eval R  S g (P p  (mset_to_IP R {#i#})) = (poly_eval R S g P) i "
    using assms
    by (metis poly_eval_indexed_pmult poly_index_mult)
  have 1: "poly_eval R S g (mset_to_IP R {#i#})= mset_to_IP R {#i#}"
    using False
    by (metis assms(2) indexed_pmult_zero monom_add_mset one_closed one_mset_to_IP poly_eval_constant poly_eval_index)
  then show ?thesis
    using 0 False assms poly_eval_index[of g]
    by (metis UNIV_I cring.Pring_set_restrict is_cring local.ring_axioms poly_eval_closed ring.poly_index_mult subsetI)
qed

lemma poly_eval_monom_mult:
  assumes "P  Pring_set R I"
  assumes "closed_fun R g"
  shows "poly_eval R  S g (P  p  (mset_to_IP R m))  = poly_eval R S g P  p  poly_eval R S g (mset_to_IP R m) "
proof(induct m)
  case empty
  have 0:  "mset_to_IP R {#} = indexed_const 𝟭"
    using one_mset_to_IP by blast
  then have 1: "(P p mset_to_IP R {#}) = P"
    using assms
    by (metis P_ring_mult_comm Pring_cfs_closed carrier_coeff_def
        mset_to_IP_simp mset_to_IP_simp' one_closed one_mult_left zero_closed)
  have 2: "poly_eval R S g  (mset_to_IP R {#}) = indexed_const 𝟭"
    by (metis "0" one_closed poly_eval_constant)
  show ?case using 0 1 2
    by (metis assms(1) assms(2) cring.P_ring_mult_comm indexed_pset.indexed_const
        is_cring local.ring_axioms one_closed one_mult_left poly_eval_closed
        ring.indexed_pset_in_carrier set_eq_subset)
next
  case (add x m)
  fix x
  fix m
  assume A: "poly_eval R  S g (P p mset_to_IP R m) = poly_eval R S g P p poly_eval R S g (mset_to_IP R m)"
  show "poly_eval R  S g (P p mset_to_IP R (add_mset x m))= poly_eval R S g P p poly_eval R  S g (mset_to_IP R (add_mset x m))"
  proof-
    obtain J where J_def: "J = I  set_mset m {x}"
      by blast
    have I0: "P  Pring_set R J"
      using J_def assms
      by (meson Pring_carrier_subset Un_upper1 subsetD)
    have I1: "set_mset m  J"
      using J_def by blast
    have I2: "x  J"
      using J_def by blast
    have "mset_to_IP R (add_mset x m)= (mset_to_IP R m)  x"
      by (simp add: monom_add_mset)
    then have "(P p mset_to_IP R (add_mset x m)) = P p((mset_to_IP R m)  x)"
      by simp
    then have "(P p mset_to_IP R (add_mset x m)) = P p((mset_to_IP R m) p (mset_to_IP R {#x#}))"
      by (metis add_mset_add_single monom_mult)
    then have I3: "(P p mset_to_IP R (add_mset x m)) = (P p(mset_to_IP R m)) p (mset_to_IP R {#x#})"
      by (metis I0 P_ring_mult_assoc indexed_pset_in_carrier mset_to_IP_closed set_eq_subset)
    have "poly_eval R  S g (P p mset_to_IP R m p mset_to_IP R {#x#}) =
            poly_eval R  S g (P p mset_to_IP R m) p poly_eval R S g  (mset_to_IP R {#x#})"
      using poly_eval_indexed_pmult'[of "(P p(mset_to_IP R m))" J g x]
            I0 I1 I2 assms(2) assms(1) mset_to_IP_mult_closed
      by blast
    then have "poly_eval R S g (P p mset_to_IP R (add_mset x m)) =
          poly_eval R  S g (P p(mset_to_IP R m)) p  poly_eval R  S g (mset_to_IP R {#x#}) "
      using I3
      by simp
    then have "poly_eval R  S g (P p mset_to_IP R (add_mset x m)) =
         poly_eval R S g P p poly_eval R S g (mset_to_IP R m)  p  poly_eval R S g (mset_to_IP R {#x#}) "
      by (simp add: A)
    then have "poly_eval R S g  (P p mset_to_IP R (add_mset x m)) =
         poly_eval R S g P p (poly_eval R S g (mset_to_IP R m)  p  poly_eval R  S g (mset_to_IP R {#x#}))"
      using P_ring_mult_assoc[of "poly_eval R S g P " "poly_eval R S g (mset_to_IP R m)"  " poly_eval R  S g (mset_to_IP R {#x#})"]
      by (metis assms(1) assms(2)  indexed_pset_in_carrier mset_to_IP_closed poly_eval_closed set_eq_subset)
    then have "poly_eval R  S g (P p mset_to_IP R (add_mset x m)) =
         poly_eval R S g P p (poly_eval R  S g ((mset_to_IP R m) p (mset_to_IP R {#x#})))"
      by (metis I1 I2 assms(2)  mset_to_IP_closed poly_eval_indexed_pmult')
    then show ?thesis
      by (metis add_mset_add_single monom_mult)
  qed
qed

abbreviation mon_term ("Mt") where
"Mt k m  poly_scalar_mult R k (mset_to_IP R m)"

lemma poly_eval_monom_term_mult:
  assumes "P  Pring_set R I"
  assumes "closed_fun R g"
  assumes "k  carrier R"
  shows "poly_eval R S g (P p (Mt k m))   = poly_eval R S g P  p  poly_eval R S g (Mt k m) "
proof-
  obtain J where J_def: "J = I  (set_mset m)"
    by blast
  have J0: "P  Pring_set R J"
    using J_def Pring_carrier_subset Un_upper1 assms(1) by blast
  have J1: "mset_to_IP R m   Pring_set R J"
    by (metis J_def Un_upper2 mset_to_IP_closed)
  have 0: "(P  p  (Mt k m)) = poly_scalar_mult R k (P  p  mset_to_IP R m)"
    using times_poly_scalar_mult[of P J "mset_to_IP R m" k ] J0 J1 assms(3)
    by blast
  have 1: "poly_eval R S g (P p (Mt k m)) = poly_scalar_mult R k (poly_eval R  S g (P  p  mset_to_IP R m))"
    by (metis "0" J0 J1 P_ring_mult_closed' assms(2) assms(3) cring.poly_eval_scalar_mult is_cring)
  have 2: "poly_eval R S g (P p (Mt k m)) = poly_scalar_mult R k ((poly_eval R S g P)  p (poly_eval R S g (mset_to_IP R m)))"
    by (metis "1" assms(1) assms(2)  poly_eval_monom_mult)
  have 3: "poly_eval R S g (P p (Mt k m)) = (poly_eval R S g P)  p poly_scalar_mult R k  (poly_eval R S g (mset_to_IP R m))"
    by (metis "0" "2" J0 J1 assms(2) assms(3)  poly_eval_closed times_poly_scalar_mult)
  then show ?thesis
    by (metis J1 assms(3) assms(2) poly_eval_scalar_mult)
qed

lemma poly_eval_mult:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  assumes "closed_fun R g"
  shows "poly_eval R S g (P p Q)  = poly_eval R S g P  p  poly_eval R S g Q "
proof-
  obtain Pr where Pr_def: "Pr = (λQ. poly_eval R S g (P p Q)   = poly_eval R S g P  p  poly_eval R S g Q )"
    by blast
  have "Pr Q"
  proof(rule poly_monom_induction2[of _ I])
    show "Q  Pring_set R I"
      by (simp add: assms(2))
    show "Pr (indexed_const 𝟬)"
    proof-
      have 0: "(P p indexed_const 𝟬) = indexed_const 𝟬"
        by (metis assms(1) cring.P_ring_mult_comm indexed_pset.indexed_const
            indexed_pset_in_carrier is_cring poly_scalar_mult_eq poly_scalar_mult_zero
            set_eq_subset zero_closed)
      have 1: "poly_eval R  S g (indexed_const 𝟬) = indexed_const 𝟬"
        using poly_eval_constant by blast
      have 2: "poly_eval R S g P  p poly_eval R  S g (indexed_const 𝟬) = indexed_const 𝟬"
        using 1
        by (metis "0" assms(1) assms(3) local.ring_axioms one_closed poly_eval_monom_term_mult
            poly_scalar_mult_const r_one ring.one_mset_to_IP zero_closed)
      have 3: "poly_eval R S g (P p indexed_const 𝟬) = indexed_const 𝟬"
        using "0" "1" by presburger
      show ?thesis
        using 2 3 Pr_def
        by presburger
    qed
    show "m k. set_mset m  I  k  carrier R  Pr (poly_scalar_mult R k (mset_to_IP R m))"
      using Pr_def assms(1) assms(3)  poly_eval_monom_term_mult by blast
    show " Q m k. Q  Pring_set R I  Pr Q  set_mset m  I  k  carrier R  Pr (Q  Mt k m)"
    proof-
      fix Q
      fix m
      fix k
      assume A: "Q  Pring_set R I  Pr Q  set_mset m  I  k  carrier R"
      show "Pr (Q  Mt k m)"
      proof-
        have 0: "poly_eval R S g (P p Q)   = poly_eval R S g P  p  poly_eval R S g Q "
          using A Pr_def by blast
        have 1: "poly_eval R S g (P p  (Mt k m))   = poly_eval R S g P  p  poly_eval R S g (Mt k m)"
          using A assms(1) assms(3)  poly_eval_monom_term_mult
          by blast
        have 2: "P p  (Q  Mt k m) =(P p  Q) (P p  (Mt k m)) "
          by (meson A assms(1) local.ring_axioms mset_to_IP_closed poly_scalar_mult_closed
              ring.P_ring_rdistr ring.indexed_pset_in_carrier set_eq_subset)
        have 3: "poly_eval R S g (P p  (Q  Mt k m))= poly_eval R S g (P p Q)   poly_eval R S g (P p  (Mt k m))     "
          by (metis "2" A P_ring_mult_closed' assms(1) assms(3)
              mset_to_IP_closed  poly_eval_add poly_scalar_mult_closed)
        have 4: "poly_eval R S g (P p  (Q  Mt k m))=
        (poly_eval R S g P  p  poly_eval R S g Q)  ( poly_eval R S g P  p  poly_eval R S g (Mt k m) )"
          by (simp add: "0" "1" "3")
        have 5: " poly_eval R S g P p (poly_eval R S g Q  poly_eval R S g (Mt k m)) = poly_eval R S g P p poly_eval R S g Q  (poly_eval R S g P p poly_eval R S g (Mt k m))"
          using ring.P_ring_rdistr[of R "(poly_eval R S g P)" "(poly_eval R S g Q)" "( poly_eval R S g (Mt k m) )"]
          by (meson A assms(1) assms(3) indexed_pset_in_carrier local.ring_axioms
              poly_eval_closed poly_scalar_mult_closed ring.mset_to_IP_closed subset_refl)
        have 6: "poly_eval R S g (P p  (Q  Mt k m))=
        (poly_eval R S g P)  p  ((poly_eval R S g Q)  ( poly_eval R S g (Mt k m) ))"
          using 4 5
          by simp
        have 7: "poly_eval R S g (P p  (Q  Mt k m))=
        (poly_eval R S g P)  p  (poly_eval R  S g (Q(Mt k m)))"
          using 6 poly_eval_add[of "(poly_eval R S g Q)" I "(Mt k m)" g ]
          by (metis A assms(3) mset_to_IP_closed poly_eval_add poly_scalar_mult_closed)
        show ?thesis using 7
          using Pr_def by blast
      qed
    qed
  qed
  then show ?thesis
    using Pr_def by blast
qed

lemma poly_eval_Pring_mult:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  assumes "closed_fun R g"
  shows "poly_eval R S g (P Pring R IQ)  = poly_eval R S g P Pring R Ipoly_eval R S g Q "
  by (metis Pring_mult assms(1) assms(2) assms(3) poly_eval_mult)

lemma poly_eval_smult:
  assumes "P  Pring_set R I"
  assumes "a  carrier R"
  assumes "closed_fun R g"
  shows "poly_eval R S g (a Pring R IP)  =a Pring R Ipoly_eval R S g P"
  using poly_eval_scalar_mult assms
  by (metis Pring_smult)

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Partial Evaluation is a Homomorphism›
(**************************************************************************************************)
(**************************************************************************************************)

lemma poly_eval_ring_hom:
  assumes "I  J"
  assumes "closed_fun R g"
  assumes "J - S  I"
  shows "ring_hom_ring (Pring R J) (Pring R I) (poly_eval R S g)"
  apply(rule ring_hom_ringI)
  apply (simp add: Pring_is_ring)
  apply (simp add: Pring_is_ring)
  apply (metis (full_types) Pring_car Pring_carrier_subset assms(2) assms(3) poly_eval_closed subset_iff)
  apply (metis Pring_car assms(2) local.ring_axioms poly_eval_mult ring.Pring_mult)
  apply (metis Pring_add Pring_car assms(2) poly_eval_add)
  by (metis Pring_one one_closed poly_eval_constant)

text‹\texttt{poly\_eval} R at the zero function is an inverse to the inclusion of polynomial rings›

lemma poly_eval_zero_function:
  assumes "g = (λn. 𝟬)"
  assumes "J - S = I"
  shows "P  Pring_set R I  poly_eval R S g P = P"
  apply(erule indexed_pset.induct)
  using poly_eval_constant apply blast
  using  assms poly_eval_add[of _ I _ g S]  zero_closed
  apply (metis Pi_I)
  using Diff_iff assms(1) assms(2) Pi_I[of UNIV g ]
      poly_eval_indexed_pmult set_eq_subset subsetD zero_closed
  by smt

lemma poly_eval_eval_function_eq:
  assumes "closed_fun R g"
  assumes "closed_fun R g'"
  assumes "restrict g S = restrict g' S"
  assumes "P  Pring_set R I"
  shows "poly_eval R S g P = poly_eval R S g' P"
  apply(rule indexed_pset.induct[of P I "carrier R"])
     apply (simp add: assms(4))
    apply (metis poly_eval_constant)
   apply (metis assms(1) assms(2) poly_eval_add)
proof- fix P i assume A: "P  Pring_set R I" "poly_eval R S g P = poly_eval R S g' P" "i  I "
  show "poly_eval R S g (P  i) = poly_eval R S g' (P  i)"
  proof(cases "i  S")
    case True
    then have "g i = g' i"
      using assms
      unfolding  restrict_def
      by meson
    then show ?thesis
      using assms A poly_eval_indexed_pmult[of P I g S i] poly_eval_indexed_pmult[of P I g' S i]
      by presburger
  next
    case False
    then show ?thesis
      using assms A poly_eval_indexed_pmult[of P I g S i] poly_eval_indexed_pmult[of P I g' S i]
      by presburger
  qed
qed

lemma poly_eval_eval_set_eq:
  assumes "closed_fun R g"
  assumes "S  I = S'  I"
  assumes "P  Pring_set R I"
  assumes "𝟭 𝟬"
  shows "poly_eval R S g P = poly_eval R S' g P"
  apply(rule indexed_pset.induct[of P I "carrier R"])
     apply (simp add: assms(3))
  apply (metis poly_eval_constant)
   apply (metis assms(1) poly_eval_add)
proof- fix P i
  assume A: " P  Pring_set R I" "poly_eval R S g P = poly_eval R S' g P" "i  I "
  show "poly_eval R S g (P  i) = poly_eval R S' g (P  i)"
  using assms poly_eval_index[of g _ i] A
  by (metis Diff_Diff_Int Diff_iff poly_eval_indexed_pmult)
qed

lemma poly_eval_trivial:
  assumes "closed_fun R g"
  assumes "P  Pring_set R (I - S)"
  shows "poly_eval R S g P = P"
  apply(rule indexed_pset.induct[of P "I - S" "carrier R"])
  apply (simp add: assms(2))
  using poly_eval_constant apply blast
  apply (metis assms(1) poly_eval_add)
  by (metis Diff_iff assms(1) poly_eval_indexed_pmult)

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Total Evaluation of a Polynomial›
(**************************************************************************************************)
(**************************************************************************************************)

lemma zero_fun_closed:
"closed_fun R (λn. 𝟬)"
  by blast

lemma deg_zero_cf_eval:
  shows "P  Pring_set R I  poly_eval R I (λn. 𝟬) P = indexed_const (P {#})"
  apply(erule indexed_pset.induct)
  apply (metis indexed_const_def poly_eval_constant)
proof-
  show " P Q. P  Pring_set R I 
           poly_eval R I (λn. 𝟬) P = indexed_const (P {#}) 
           Q  Pring_set R I  poly_eval R I (λn. 𝟬) Q = indexed_const (Q {#})  poly_eval R I (λn. 𝟬) (P  Q) = indexed_const ((P  Q) {#})"
  proof-
    fix P
    fix Q
    assume A: "P  Pring_set R I"
    assume B: " poly_eval R I (λn. 𝟬) P = indexed_const (P {#})"
    assume C: "Q  Pring_set R I "
    assume D: " poly_eval R I (λn. 𝟬) Q = indexed_const (Q {#})"
    have "indexed_const ((P  Q) {#}) = indexed_const (P {#})   indexed_const (Q {#})"
      by (metis indexed_padd_const indexed_padd_def)
    thus "poly_eval R I (λn. 𝟬) (P  Q) = indexed_const ((P  Q) {#})"
      using A B C D poly_eval_add[of P I Q "λn. 𝟬" I]
      by (smt zero_fun_closed)
  qed
  show "P i. P  Pring_set R I  poly_eval R I (λn. 𝟬) P = indexed_const (P {#})  i  I  poly_eval R I (λn. 𝟬) (P  i) = indexed_const ((P  i) {#})"
  proof-
    fix P
    fix i
    assume A: "P  Pring_set R I" "poly_eval R I (λn. 𝟬) P = indexed_const (P {#})" "i  I"
    show "poly_eval R I (λn. 𝟬) (P  i) = indexed_const ((P  i) {#})"
    proof-
      have "poly_eval R I (λn. 𝟬) (P  i) = indexed_const 𝟬"
        using A(1) A(3) cring.poly_eval_constant is_cring poly_eval_indexed_pmult
              poly_eval_scalar_mult poly_scalar_mult_zero zero_closed
        by (metis Pi_I)
      have "(P  i) {#} = 𝟬"
        using indexed_pmult_def
        by (metis empty_iff set_mset_empty)
      then show ?thesis
        using poly_eval R I (λn. 𝟬) (P  i) = indexed_const 𝟬
        by presburger
    qed
  qed
qed

lemma deg_zero_cf_mult:
  assumes "P  Pring_set R I"
  assumes "Q  Pring_set R I"
  shows " (P p Q) {#} = P {#}  Q {#}"
proof-
  have "poly_eval R I (λn. 𝟬) (P p Q) = poly_eval R I (λn. 𝟬) P p  poly_eval R I (λn. 𝟬) Q"
    using zero_fun_closed assms
    by (metis poly_eval_mult)
  then have 0: "indexed_const ((P p Q) {#}) = (indexed_const (P {#})) p (indexed_const (Q {#}))"
    by (metis P_ring_mult_closed' Pring_cfs_closed assms(1) assms(2) deg_zero_cf_eval indexed_const_P_mult_eq indexed_const_def)
  have "indexed_const ((P p Q) {#}) = (indexed_const ((P {#}) (Q {#})))"
    apply(rule ccontr)
    using 0
    by (metis Pring_cfs_closed assms(1) assms(2) indexed_const_P_mult_eq)
  then show ?thesis
  proof -
    show ?thesis
      by (metis (no_types) indexed_const ((P p Q) {#}) = indexed_const (P {#}  Q {#})
        local.ring_axioms ring.indexed_const_def)
  qed
qed

definition deg_zero_cf :: "('a, 'c) mvar_poly  'a" where
"deg_zero_cf P = P {#}"

lemma deg_zero_cf_ring_hom:
  shows "ring_hom_ring (Pring R I) R (deg_zero_cf)"
  apply(rule ring_hom_ringI)
  using Pring_is_ring apply blast
  apply (simp add: local.ring_axioms)
  apply (metis deg_zero_cf_def Pring_car Pring_cfs_closed)
  apply (metis deg_zero_cf_def Pring_car Pring_mult deg_zero_cf_mult)
  apply (metis deg_zero_cf_def Pring_add indexed_padd_def)
  by (metis deg_zero_cf_def Pring_one indexed_const_def)

end

definition eval_in_ring ::
 "('a,'b) ring_scheme  'c set  ('c  'a)  ('a, 'c) mvar_poly  'a" where
"eval_in_ring R S g P = (poly_eval R S g P) {#}"

definition total_eval ::
"('a,'b) ring_scheme  ('c  'a)  ('a, 'c) mvar_poly  'a"  where
"total_eval R g P = eval_in_ring R UNIV g P"

context cring
begin

lemma eval_in_ring_ring_hom:
  assumes "closed_fun R g"
  shows "ring_hom_ring (Pring R I) R (eval_in_ring R S g)"
  unfolding eval_in_ring_def
  apply(rule ring_hom_ringI)
  apply (simp add: Pring_is_ring)
  apply (simp add: local.ring_axioms)
    using Pring_car Pring_carrier_coeff' assms poly_eval_closed
     apply (metis )
        using Pring_mult[of I] poly_eval_mult[of _ I _ g S] poly_eval_closed[of g _ I S] deg_zero_cf_mult[of _ I] assms
          Pring_car[of I]
        apply (metis deg_zero_cf_mult)

          using Pring_add[of I ] poly_eval_add[of _ I _ g S] Pring_car[of I] assms indexed_padd_def
          apply metis
            unfolding deg_zero_cf_def
            by (metis Pring_one indexed_const_def one_closed poly_eval_constant)

lemma eval_in_ring_smult:
  assumes "P  carrier (Pring R I)"
  assumes "a  carrier R"
  assumes "closed_fun R g"
  shows "eval_in_ring R S g (a Pring R IP) = a  eval_in_ring R S g P "
  using assms unfolding eval_in_ring_def
  by (smt Pring_car Pring_smult poly_eval_scalar_mult poly_scalar_mult_def)


lemma total_eval_ring_hom:
  assumes "closed_fun R g"
  shows "ring_hom_ring (Pring R I) R (total_eval R g)"
  using assms unfolding total_eval_def
  using eval_in_ring_ring_hom by blast

lemma total_eval_smult:
  assumes "P  carrier (Pring R I)"
  assumes "a  carrier R"
  assumes "closed_fun R g"
  shows "total_eval R g (a Pring R IP) = a  total_eval R g P"
  using assms unfolding total_eval_def
  using eval_in_ring_smult by blast

lemma total_eval_const:
  assumes "k  carrier R"
  shows "total_eval R g (indexed_const k) = k"
  unfolding total_eval_def eval_in_ring_def
  using assms
  by (metis indexed_const_def poly_eval_constant)

lemma total_eval_var:
  assumes "closed_fun R g"
  shows "(total_eval R g (mset_to_IP R {#i#})) = g i"
  unfolding total_eval_def eval_in_ring_def
  using UNIV_I assms indexed_const_def indexed_pmult_zero monom_add_mset one_closed
      one_mset_to_IP one_zeroD poly_eval_constant poly_eval_index  singletonD
  by (smt PiE iso_tuple_UNIV_I monom_eval_add monom_eval_car mset_to_IP_simp poly_scalar_mult_const r_one)

lemma total_eval_indexed_pmult:
  assumes "P  carrier (Pring R I)"
  assumes "i  I"
  assumes "closed_fun R g"
  shows "total_eval R g (P  i) = total_eval R g P Rg i"
proof-
  have "P  i = P p mset_to_IP R ((add_mset i) {#})"
    using assms poly_index_mult[of P I i] Pring_car
    by blast
  then have 0: "(P  i) =  P Pring R I(mset_to_IP R {#i#})"
    by (simp add: Pring_mult)
  then have "total_eval R g (P  i) = (total_eval R g P) R(total_eval R g (mset_to_IP R {#i#}))"
  proof-
    have "(mset_to_IP R {#i#})  carrier (Pring R I)"
      by (metis Pring_car assms(2) mset_to_IP_closed set_mset_single singletonD subset_iff)
    then show ?thesis
      using assms total_eval_ring_hom[of g I] ring_hom_mult
      by (metis "0" ring_hom_ring.homh)
  qed
  then show ?thesis
    by (metis assms(3) cring.total_eval_var is_cring)
qed

lemma total_eval_mult:
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  assumes "closed_fun R g"
  shows "total_eval R g (P Pring R IQ) = (total_eval R g P) R(total_eval R g Q) "
  by (metis assms(1) assms(2) assms(3) ring_hom_mult ring_hom_ring.homh total_eval_ring_hom)

lemma total_eval_add:
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  assumes "closed_fun R g"
  shows "total_eval R g (P Pring R IQ) = (total_eval R g P) R(total_eval R g Q) "
  by (metis assms(1) assms(2) assms(3) ring_hom_add ring_hom_ring.homh total_eval_ring_hom)

lemma total_eval_one:
  assumes "closed_fun R g"
  shows "total_eval R g 𝟭Pring R I= 𝟭"
  by (metis Pring_one one_closed total_eval_const)

lemma total_eval_zero:
  assumes "closed_fun R g"
  shows "total_eval R g 𝟬Pring R I= 𝟬"
  by (metis Pring_zero total_eval_const zero_closed)

lemma total_eval_closed:
  assumes "P  carrier (Pring R I)"
  assumes "closed_fun R g"
  shows "total_eval R g P  carrier R"
  using assms total_eval_ring_hom[of g]
  by (metis ring_hom_closed ring_hom_ring.homh)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Constructing Homomorphisms from Indexed Polynomial Rings and a Universal Property›
(**************************************************************************************************)
(**************************************************************************************************)

text‹The inclusion of R› into its polynomial ring›

lemma indexed_const_ring_hom:
"ring_hom_ring R (Pring R I) (indexed_const)"
  apply(rule ring_hom_ringI)
  apply (simp add: local.ring_axioms)
  apply (simp add: Pring_is_ring)
  using Pring_car indexed_pset.indexed_const apply blast
  apply (metis Pring_mult indexed_const_P_mult_eq)
  apply (metis indexed_padd_const local.ring_axioms ring.Pring_add)
  by (metis Pring_one)

lemma indexed_const_inj_on:
"inj_on (indexed_const) (carrier R)"
  by (metis cring.total_eval_const inj_onI is_cring)

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Mapping $R[x] \to S[x]$ along a homomorphism $R \to S$›
(**************************************************************************************************)
(**************************************************************************************************)

definition ring_hom_to_IP_ring_hom ::
"('a, 'e) ring_hom  ('a, 'c) mvar_poly  'c monomial  'e"  where
"ring_hom_to_IP_ring_hom φ P m = φ (P m)"

context cring
begin

lemma ring_hom_to_IP_ring_hom_one:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  shows "ring_hom_to_IP_ring_hom φ 𝟭Pring R I= 𝟭Pring S I⇙"
  unfolding ring_hom_to_IP_ring_hom_def
proof
  fix m
  show " φ (𝟭Pring R Im) = 𝟭Pring S Im"
  proof(cases "m = {#}")
    case True
    then have "(𝟭Pring R Im) = 𝟭R⇙"
      by (metis Pring_one indexed_const_def)
    then have "φ (𝟭Pring R Im) = 𝟭S⇙"
      using assms
      unfolding ring_hom_ring_def
      by (metis assms(2) ring_hom_one ring_hom_ring.homh)
    then show ?thesis using assms True ring.indexed_const_def ring_hom_ring_def
      by (metis ring.Pring_one)
  next
    case False
      then have "(𝟭Pring R Im) = 𝟬R⇙"
        by (metis indexed_const_def local.ring_axioms ring.Pring_one)

    then have "φ (𝟭Pring R Im) = 𝟬S⇙"
      using assms
      unfolding ring_hom_ring_def cring_def
      by (metis assms(2) ring_hom_zero ring_hom_ring.homh)
    then show ?thesis using assms False ring.indexed_const_def ring_hom_ring_def
      by (metis ring.Pring_one)
  qed
qed

lemma ring_hom_to_IP_ring_hom_constant:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  assumes "a  carrier R"
  shows "ring_hom_to_IP_ring_hom φ ((indexed_const a):: 'c monomial  'a) = ring.indexed_const S (φ a)"
  unfolding ring_hom_to_IP_ring_hom_def indexed_const_def
proof
  fix m:: "'c monomial"
  show "φ (if m = {#} then a else 𝟬) = ring.indexed_const S (φ a) m"
    apply(cases "m = {#}")
  apply (simp add: assms(1) cring.axioms(1) ring.indexed_const_def)
  proof-
    assume "m  {#} "
    then have "φ (if m = {#} then a else 𝟬) = 𝟬S⇙"
      by (metis (full_types) assms(1) assms(2) cring_def local.ring_axioms
          ring_hom_ring.homh ring_hom_zero)
    then show " φ (if m = {#} then a else 𝟬) = ring.indexed_const S (φ a) m"
      using assms
      by (metis m  {#} cring.axioms(1) ring.indexed_const_def)
  qed
qed

lemma ring_hom_to_IP_ring_hom_add:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  shows "ring_hom_to_IP_ring_hom φ (P Pring R IQ) =
       (ring_hom_to_IP_ring_hom φ P) Pring S I(ring_hom_to_IP_ring_hom φ Q)"
  unfolding ring_hom_to_IP_ring_hom_def
proof
  fix m
  show " φ ((P Pring R IQ) m) = ((λm. φ (P m)) Pring S I(λm. φ (Q m))) m"
  proof-
    have RHS: "((λm. φ (P m)) Pring S I(λm. φ (Q m))) m = φ (P m) Sφ (Q m)"
      using ring.Pring_add[of  S I _ _ ] assms
      by (metis cring.axioms(1) ring.indexed_padd_def)
    have LHS: "φ ((P Pring R IQ) m) = φ ((P m)R(Q m))"
      by (metis Pring_add indexed_padd_def)
    then show ?thesis
      using assms unfolding ring_hom_ring_def
      using  ring_hom_add[of φ R S "P m" "Q m"]
      by (metis Pring_carrier_coeff' RHS assms(2) ring_hom_ring.homh)
  qed
qed

lemma ring_hom_to_IP_ring_hom_closed:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  assumes "P  carrier (Pring R I)"
  shows "ring_hom_to_IP_ring_hom φ P  carrier (Pring S I)"
  apply(rule indexed_pset.induct[of P I "carrier R"])
  using Pring_car assms(3) apply blast
proof-
  show "k. k  carrier R  ring_hom_to_IP_ring_hom φ (indexed_const k)  carrier (Pring S I)"
  proof-
    fix k
    show " k  carrier R  ring_hom_to_IP_ring_hom φ (indexed_const k)  carrier (Pring S I)"
    proof-
      assume A: "k  carrier R"
      have "(φ k)  carrier S"
        by (meson A assms(2) ring_hom_closed ring_hom_ring.homh)
      then have 0: "ring.indexed_const S (φ k)  carrier (Pring S I)"
        by (metis assms(1) cring.axioms(1) ring.Pring_car ring.indexed_pset.indexed_const)
      then show ?thesis
        using  assms(2)
        by (simp add: "0" A ring_hom_to_IP_ring_hom_constant[of S φ k] assms(1))
    qed
  qed
  show "P Q. P  Pring_set R I 
           ring_hom_to_IP_ring_hom φ P  carrier (Pring S I) 
           Q  Pring_set R I 
           ring_hom_to_IP_ring_hom φ Q  carrier (Pring S I)  ring_hom_to_IP_ring_hom φ (P  Q)  carrier (Pring S I)"
    using ring_hom_to_IP_ring_hom_add[of S φ _ I] Pring_car[of I] ring.Pring_car[of S I]
    by (metis Pring_add assms(1) assms(2) cring.axioms(1) ring.Pring_add_closed )
  show "P i. P  Pring_set R I 
           ring_hom_to_IP_ring_hom φ P  carrier (Pring S I)  i  I  ring_hom_to_IP_ring_hom φ (P  i)  carrier (Pring S I)"
  proof-
    fix P
    fix i
    assume A: "P  Pring_set R I " "ring_hom_to_IP_ring_hom φ P  carrier (Pring S I) " "i  I"
    have 0: "(λm. φ ((P  i) m)) = ring.indexed_pmult S (ring_hom_to_IP_ring_hom φ P) i"
    proof
      fix m
      show "φ ((P  i) m) = ring.indexed_pmult S (ring_hom_to_IP_ring_hom φ P) i m"
      proof(cases "i ∈# m")
        case True
        then have LHS: "φ ((P  i) m) = φ  (P (m - {#i#}))"
          by (metis indexed_pmult_def)
        then show ?thesis
          using True
          by (metis assms(1) cring.axioms(1) ring.indexed_pmult_def ring_hom_to_IP_ring_hom_def)
      next
        case False
        then have "φ ((P  i) m) = φ 𝟬R⇙"
          by (metis indexed_pmult_def)
        then have LHS: "φ ((P  i) m) = 𝟬S⇙"
          by (metis assms(1) assms(2) cring.axioms(1) local.ring_axioms
              ring_hom_ring.homh ring_hom_zero)
        then show ?thesis
          using False assms ring.indexed_pmult_def
          by (metis cring.axioms(1))
      qed
    qed
    then show "ring_hom_to_IP_ring_hom φ (P  i)  carrier (Pring S I)"
      using assms
      unfolding ring_hom_to_IP_ring_hom_def
      by (metis "0" A(2) A(3) cring.axioms(1) ring.Pring_car ring.indexed_pset.simps)
  qed
qed

lemma ring_hom_to_IP_ring_hom_monom:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  shows "ring_hom_to_IP_ring_hom φ (mset_to_IP R m) = mset_to_IP S m"
proof
  fix x
  show "ring_hom_to_IP_ring_hom φ (mset_to_IP R m) x = mset_to_IP S m x"
    unfolding ring_hom_to_IP_ring_hom_def mset_to_IP_def apply( cases "x = m" )
     apply (metis (mono_tags, lifting) assms(2) ring_hom_one ring_hom_ring.homh)
    by (metis (full_types) assms(1) assms(2) cring.axioms(1) local.ring_axioms
        ring_hom_ring.homh ring_hom_zero)
qed

lemma Pring_morphism:
  assumes "cring S"
  assumes "φ  (carrier (Pring R I))  (carrier S)"
  assumes "φ 𝟭Pring R I= 𝟭S⇙"
  assumes "φ 𝟬Pring R I= 𝟬S⇙"
  assumes "P Q. P  carrier (Pring R I)  Q  carrier (Pring R I) 
          φ (P Pring R IQ) = (φ P) S(φ Q)"
  assumes " i .  P. i  I  P  carrier (Pring R I)   φ (P  i) =  (φ P) S(φ (mset_to_IP R {#i#}))"
  assumes "k Q. k  carrier R  Q  carrier (Pring R I)   φ (poly_scalar_mult R k Q) =
                                     (φ (indexed_const k)) S(φ Q)"
  shows "ring_hom_ring (Pring R I) S φ"
  apply(rule ring_hom_ringI)
  apply (simp add: Pring_is_ring; fail)
  apply (simp add: assms(1) cring.axioms(1); fail)
  using assms(2) apply blast
proof-

  show "x y. x  carrier (Pring R I)  y  carrier (Pring R I)  φ (x Pring R Iy) = φ x Sφ y"
  proof-
    fix P Q
    assume A0: "P  carrier (Pring R I)"
    assume A1: "Q  carrier (Pring R I)"
    show "φ (P Pring R IQ) = φ P Sφ Q"
    proof(rule mpoly_induct'[of ])
      show "φ (P Pring R Iindexed_const 𝟬) = φ P Sφ (indexed_const 𝟬)"
      proof-
        have 0: "(P Pring R Iindexed_const 𝟬) = 𝟬Pring R I⇙"
          using assms
          by (metis A0 Pring_is_cring Pring_zero cring.cring_simprules(27) is_cring)
        have 1: " φ P Sφ (indexed_const 𝟬) = 𝟬S⇙"
        proof-
          have "φ P  carrier S"
            using assms(2) A0
            by blast
          then show ?thesis
            using assms(4) Pring_zero[of I]
            by (metis assms(1) cring.cring_simprules(27))
        qed
        then show ?thesis using 0 1
          using assms(4) by presburger
      qed
      show "n Q. (Q. Q  Pring_set R I  card (monomials_of R Q)  n  φ (P Pring R IQ) = φ P Sφ Q) 
           Q  Pring_set R I  card (monomials_of R Q) = Suc n  φ (P Pring R IQ) = φ P Sφ Q"
      proof- fix n fix Q
        assume IH: " (Q. Q  Pring_set R I  card (monomials_of R Q)  n  φ (P Pring R IQ) = φ P Sφ Q)"
        assume A: "Q  Pring_set R I  card (monomials_of R Q) = Suc n"
        show "φ (P Pring R IQ) = φ P Sφ Q"
        proof-
          obtain m M where m_M_def: "monomials_of R Q = insert m M  m  M"
            using A
            by (metis card_0_eq ex_in_conv finite.emptyI mk_disjoint_insert nat.distinct(1))
          have "Q = (restrict_poly_to_monom_set R Q M) Pring R I(poly_scalar_mult R (Q m) (mset_to_IP R m))"
            by (metis A Pring_add m_M_def remove_monom_eq remove_monom_restrict_poly_to_monom_set)
          obtain Q' where Q'_def: "Q' = (restrict_poly_to_monom_set R Q M)"
            by simp
          have Q'_fact: " Q'  Pring_set R I  card (monomials_of R Q')  n"
            proof-
              have 0: "Q'  Pring_set R I"
                using Q'_def  A restriction_closed
                by blast
              have "monomials_of R Q' = M"
                using A m_M_def Q'_def
                by (metis restrict_poly_to_monom_set_monoms subset_insertI)
              then have "card (monomials_of R Q') = n" using A m_M_def
                by (metis "0" card_insert_disjoint diff_Suc_1 monomials_finite)
              then show ?thesis
                by (simp add: "0")
            qed
          have 0:"(P Pring R IQ) = (P Pring R I(Q' Pring R I(poly_scalar_mult R (Q m) (mset_to_IP R m))))"
            using Q'_def Q = restrict_poly_to_monom_set R Q M Pring R IMt (Q m) m by presburger
          have 1: "(P Pring R IQ) = (P Pring R IQ') Pring R I(P Pring R I(poly_scalar_mult R (Q m) (mset_to_IP R m)))"
          proof-
            have 10: "P  carrier (Pring R I)"
              by (simp add: A0)
            have 11: "Q  carrier (Pring R I)"
              by (simp add: A Pring_car)
            have 12: "Q'  carrier (Pring R I)"
              using A Pring_car Q'_def restriction_closed by blast
            have 13: "(poly_scalar_mult R (Q m) (mset_to_IP R m))  carrier (Pring R I)"
              by (metis A Pring_car Pring_carrier_coeff' insert_iff m_M_def mset_to_IP_closed' poly_scalar_mult_closed)
            then show ?thesis
              using 0 10 11 12 13
              by (metis Pring_mult_rdistr)
          qed
          then have "φ (P Pring R IQ) = (φ (P Pring R IQ')) Sφ (P Pring R I(poly_scalar_mult R (Q m) (mset_to_IP R m)))"
            by (metis A A0 Pring_car Pring_mult_closed Pring_cfs_closed Q'_def assms(5) insert_iff
                local.ring_axioms m_M_def poly_scalar_mult_closed ring.mset_to_IP_closed'
                ring.restriction_closed)
          then have "φ (P Pring R IQ) = (φ P) S(φ Q') Sφ (P Pring R I(poly_scalar_mult R (Q m) (mset_to_IP R m)))"
            using IH[of Q'] Q'_fact
            by presburger
          then have 2: "φ (P Pring R IQ) = (φ P) S(φ Q') Sφ (poly_scalar_mult R (Q m) (P Pring R I(mset_to_IP R m)))"
            by (metis A A0 Pring_car Pring_mult Pring_cfs_closed insert_iff m_M_def mset_to_IP_closed' times_poly_scalar_mult)
          have 3: "k. k  carrier R  set_mset m  I  φ (poly_scalar_mult R k (P Pring R I(mset_to_IP R m))) =
                    φ (indexed_const k) S(φ P) Sφ (mset_to_IP R m)"
          proof(induction m)
            case empty
            assume A: "set_mset {#}  I"
            then have E0: "(P Pring R Imset_to_IP R {#}) = P"
              by (metis A0 Pring_mult_one Pring_one one_mset_to_IP)
            have E1: "k  carrier R"
              using A Pring_cfs_closed empty.prems(1)
              by linarith
            have E2: "φ (mset_to_IP R {#}) = 𝟭S⇙"
              by (metis Pring_one assms(3) one_mset_to_IP)
            have E3: " φ (poly_scalar_mult R k (P Pring R Imset_to_IP R {#})) =
                       φ (indexed_const k) Sφ (P Pring R Imset_to_IP R {#})"
              using E1 E2 E0 assms(7)[of "k" "P Pring R Imset_to_IP R {#}"]  A0
              by (simp add: A0 E1)
            have E4: " φ (P Pring R Imset_to_IP R {#}) =  (φ P) "
              using E0
              by auto
            have E5: " φ (poly_scalar_mult R k (P Pring R Imset_to_IP R {#})) =
                       φ (indexed_const k) Sφ P"
              using E0 E3 by presburger
            have E6: " φ (indexed_const k) Sφ P =  φ (indexed_const k) Sφ P S𝟭S⇙"
            proof-
              have 0: "φ P  carrier S"
                using assms A0
                by blast
              have "indexed_const k  carrier (Pring R I)"
                using assms(2) E1 Pring_car indexed_pset.indexed_const
                by blast
              then have 1: "φ (indexed_const k)  carrier S"
                using assms(2)
                by blast
              show ?thesis
                using assms(1) 0 1
                by (metis cring.cring_simprules(12) cring.cring_simprules(14)
                    cring.cring_simprules(5) cring.cring_simprules(6))
            qed
            then show ?case
              using E0 E2 E3 by presburger
          next
            case (add x m)
            assume AA: "k. k  carrier R  set_mset m  I  φ (poly_scalar_mult R k (P Pring R Imset_to_IP R m)) = φ (indexed_const k) Sφ P Sφ (mset_to_IP R m)"
            assume P0: "set_mset (add_mset x m)  I"
            then have IA: "set_mset m I"
              by (metis insert_subset set_mset_add_mset_insert)
            have IH: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R m)) =
                        φ (indexed_const k) Sφ P Sφ (mset_to_IP R m)"
              using AA IA add.prems(1)
              by blast
            then have x_mem: "x  I"
              by (meson P0 subsetD union_single_eq_member)
            have 0: " mset_to_IP R (add_mset x m) =  mset_to_IP R m  x"
              by (simp add: monom_add_mset)
            then have 1: "P Pring R Imset_to_IP R (add_mset x m) = P Pring R I(mset_to_IP R m  x)"
              by simp
            have 2: "P Pring R Imset_to_IP R (add_mset x m) =( P Pring R I(mset_to_IP R m))  x"
            proof-
              have "mset_to_IP R (add_mset x m) = (mset_to_IP R m)  x"
                using "0" by blast
              then have "P Pring R Imset_to_IP R (add_mset x m) = P Pring R I((mset_to_IP R m)  x)"
                by simp
              then have "P Pring R Imset_to_IP R (add_mset x m) = P Pring R I((mset_to_IP R m) Pring R Imset_to_IP R {#x#})"
                by (metis IA Pring_mult mset_to_IP_closed poly_index_mult x_mem)
              then have "P Pring R Imset_to_IP R (add_mset x m) = (P Pring R I(mset_to_IP R m)) Pring R Imset_to_IP R {#x#}"
                by (metis A0 IA Pring_car Pring_mult_assoc Pring_one Pring_one_closed
                    indexed_pset.indexed_pmult monom_add_mset mset_to_IP_closed one_mset_to_IP x_mem)
              then show ?thesis
                by (metis A0 IA Pring_car Pring_mult Pring_mult_closed mset_to_IP_closed poly_index_mult x_mem)
            qed
            have 3: "poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m)) =
                      poly_scalar_mult R k ( P Pring R I(mset_to_IP R m)  x)"
              using "2" by presburger
            have 4: "poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m)) =
                      poly_scalar_mult R k ( P Pring R I(mset_to_IP R m))  x"
            proof-
              have "poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m)) =
                      poly_scalar_mult R k ( P Pring R I(mset_to_IP R m) Pring R Imset_to_IP R {#x#})"
                by (metis "2" A0 IA Pring_car Pring_mult Pring_mult_closed mset_to_IP_closed poly_index_mult x_mem)
              have "poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m)) =
                      (poly_scalar_mult R k ( P Pring R Imset_to_IP R m)) Pring R Imset_to_IP R {#x#}"
                by (metis A0 IA Pring_car Pring_mult Pring_mult_closed Pring_one Pring_one_closed
                    poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m)) = poly_scalar_mult R k (P Pring R Imset_to_IP R m Pring R Imset_to_IP R {#x#})
                    add.prems(1) indexed_pset.indexed_pmult monom_add_mset mset_to_IP_closed
                    one_mset_to_IP poly_scalar_mult_times x_mem)
              then show ?thesis
                by (metis A0 IA Pring_car Pring_mult Pring_mult_closed add.prems(1)
                    cring.poly_scalar_mult_closed is_cring mset_to_IP_closed
                    poly_index_mult x_mem)
            qed
            have 5: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
                      φ (poly_scalar_mult R k ( P Pring R I(mset_to_IP R m))  x)"
              using 4 by metis
            have 6: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
                      φ (poly_scalar_mult R k ( P Pring R I(mset_to_IP R m))) Sφ (mset_to_IP  R {#x#})"
              using assms
              by (metis "5" A0 IA Pring_car Pring_mult_closed add.prems(1)
                  cring.poly_scalar_mult_closed is_cring mset_to_IP_closed x_mem)
            have 7: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
                       φ (indexed_const k) Sφ P Sφ (mset_to_IP R m) Sφ (mset_to_IP  R {#x#})"
              using assms 6 IH
              by presburger
            have 8: " φ (mset_to_IP R m) Sφ (mset_to_IP  R {#x#}) = φ (mset_to_IP R (add_mset x m))"
            proof-
              have "φ (mset_to_IP R m) Sφ (mset_to_IP  R {#x#}) =
                    φ ((mset_to_IP R m) Pring R I(mset_to_IP  R {#x#}))"
                by (metis IA Pring_car Pring_mult assms(6) mset_to_IP_closed poly_index_mult x_mem)
              then show ?thesis
                by (metis "0" IA Pring_car assms(6) mset_to_IP_closed x_mem)
            qed
            have 9: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
             φ (indexed_const k) Sφ P S( φ (mset_to_IP R m) Sφ (mset_to_IP  R {#x#}))"
            proof-
              have 0: "φ (indexed_const k)  carrier S"
              proof-
                have "(indexed_const k)  carrier (Pring R I)"
                  using A Pring_car Pring_cfs_closed indexed_pset.indexed_const add.prems(1)
                  by blast
                then show ?thesis
                  using assms(2)
                  by blast
              qed
              have 1: "φ P  carrier S"
                using A0 assms(2)
                by blast
              have 2: "φ (mset_to_IP R m)  carrier S"
              proof-
                have "(mset_to_IP R m)  carrier (Pring R I)"
                  using IA Pring_car mset_to_IP_closed by blast
                then show ?thesis using assms(2)
                  by blast
              qed
              have 3: " φ (mset_to_IP  R {#x#})  carrier S"
              proof-
                have "(mset_to_IP  R {#x#})  carrier (Pring R I)"
                  by (metis Pring_car Pring_one Pring_one_closed indexed_pset.indexed_pmult
                      monom_add_mset one_mset_to_IP x_mem)
                then show ?thesis
                  using assms(2)
                  by blast
              qed
              have "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
                  φ (indexed_const k) S(φ P Sφ (mset_to_IP R m) Sφ (mset_to_IP  R {#x#}))"
                using 1 2 3 7
                by (metis "0" φ (mset_to_IP R m)  carrier S assms(1)
                    cring.cring_simprules(11) cring.cring_simprules(5))
              then show ?thesis
                by (metis "0" "1" "2" "3" "8" Pring_mult assms(1) cring.cring_simprules(11)
                    cring.cring_simprules(5))
            qed
            have 10: "φ (poly_scalar_mult R k (P Pring R Imset_to_IP R (add_mset x m))) =
             φ (indexed_const k) Sφ P Sφ (mset_to_IP R (add_mset x m))"
              using 8 9
              by presburger
            then show ?case
              by blast
          qed
          have 4: "φ (poly_scalar_mult R (Q m) (P Pring R I(mset_to_IP R m))) =
                    φ (indexed_const (Q m)) S(φ P) Sφ (mset_to_IP R m)"
            using 3
            by (metis A Pring_mult insert_iff local.ring_axioms m_M_def mset_to_IP_indices
                ring.Pring_cfs_closed)
          have 5: "φ (P Pring R IQ) = (φ P) S(φ Q') Sφ (indexed_const (Q m)) S(φ P) Sφ (mset_to_IP R m)"
            using 2 4
            by presburger
          have " φ (indexed_const (Q m)) S(φ P) Sφ (mset_to_IP R m) =
                    (φ P) Sφ (indexed_const (Q m)) Sφ (mset_to_IP R m)"
          proof-
            have 0: "(φ P)   carrier S"
              using A0 assms(2)
              by blast
            have 1: "φ (indexed_const (Q m))  carrier S"
            proof-
              have "indexed_const (Q m )  carrier (Pring R I)"
                using A Pring_car Pring_cfs_closed indexed_pset.indexed_const by blast
              then show ?thesis
                using assms(2)
                by blast
            qed
            have 2: "φ (mset_to_IP R m)  carrier S"
            proof-
              have "(mset_to_IP R m)  carrier (Pring R I)"
                by (metis A Pring_car insert_iff m_M_def mset_to_IP_closed')
              then show ?thesis using assms(2)
                by blast
            qed
            show ?thesis using assms(1) 0 1 2
              by (metis cring.cring_simprules(14))
          qed
          then  have 6: "φ (P Pring R IQ) = (φ P) S(φ Q') S(φ P) Sφ (indexed_const (Q m)) Sφ (mset_to_IP R m)"
            using 5
            by presburger
          then have 7: "φ (P Pring R IQ) = (φ P) S((φ Q') Sφ (indexed_const (Q m)) Sφ (mset_to_IP R m))"
          proof-
            have 0: "(φ P)   carrier S"
              using A0 assms(2)
              by blast
            have 1: "φ (indexed_const (Q m))  carrier S"
            proof-
              have "indexed_const (Q m )  carrier (Pring R I)"
                using A Pring_car Pring_cfs_closed indexed_pset.indexed_const by blast
              then show ?thesis
                using assms(2)
                by blast
            qed
            have 2: "φ (mset_to_IP R m)  carrier S"
            proof-
              have "(mset_to_IP R m)  carrier (Pring R I)"
                by (metis A Pring_car insert_iff m_M_def mset_to_IP_closed')
              then show ?thesis using assms(2)
                by blast
            qed
            have 3: "φ Q'  carrier S"
            proof-
              have "Q'  carrier (Pring R I)"
                using Q'_fact Pring_car
                by blast
              then show ?thesis
                using assms(2)
                by blast
            qed
            show ?thesis using 0 1 2 3 6
              by (metis assms(1) cring.cring_simprules(11) cring.cring_simprules(25) cring.cring_simprules(5))
          qed
          then show ?thesis
            by (metis A Pring_car Pring_cfs_closed Q'_def Q'_fact
                Q = restrict_poly_to_monom_set R Q M Pring R IMt (Q m) m assms(5) assms(7)
                cring.poly_scalar_mult_closed insert_iff is_cring m_M_def mset_to_IP_closed')
        qed
      qed
      show "Q  Pring_set R I "
        using A1 Pring_car
        by blast
    qed
  qed
  show "x y. x  carrier (Pring R I)  y  carrier (Pring R I)  φ (x Pring R Iy) = φ x Sφ y"
    using assms(5)
    by blast
  show "φ 𝟭Pring R I= 𝟭S⇙"
    by (simp add: assms(3))
qed

lemma(in cring) indexed_const_Pring_mult:
  assumes "k  carrier R"
  assumes "P  carrier (Pring R I)"
  shows "(indexed_const k Pring R IP) m = k R(P m)"
        "(P Pring R Iindexed_const k) m = k R(P m)"
  apply (metis Pring_car Pring_mult assms(1) assms(2) poly_scalar_mult_def poly_scalar_mult_eq)
  by (metis Pring_car Pring_carrier_coeff' Pring_mult assms(1) assms(2) indexed_const_P_multE m_comm)

lemma(in cring) ring_hom_to_IP_ring_hom_is_hom:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  shows "ring_hom_ring (Pring R I) (Pring S I) (ring_hom_to_IP_ring_hom φ)"
proof(rule Pring_morphism)
  show 0: "cring (Pring S I)"
    by (simp add: assms(1) cring.axioms(1) ring.Pring_is_cring)
  show 1: "ring_hom_to_IP_ring_hom φ  carrier (Pring R I)  carrier (Pring S I)"
    by (meson Pi_I assms(1) assms(2) ring_hom_to_IP_ring_hom_closed)
  show 2: "ring_hom_to_IP_ring_hom φ 𝟭Pring R I= 𝟭Pring S I⇙"
    by (simp add: assms(1) assms(2) ring_hom_to_IP_ring_hom_one)
  show 3: "ring_hom_to_IP_ring_hom φ 𝟬Pring R I= 𝟬Pring S I⇙"
  proof-
    have "m. φ (𝟬Pring R Im) = 𝟬S⇙"
      using assms
      by (metis Pring_carrier_coeff' Pring_zero Pring_zero_closed cring.axioms(1)
          cring.cring_simprules(26) indexed_const_Pring_mult(2) is_cring ring.Pring_is_cring
          ring_hom_ring.homh ring_hom_zero zero_closed)
    then have "m. φ (𝟬Pring R Im) = 𝟬Pring S Im"
      by (metis assms(1) cring.axioms(1) ring.Pring_zero ring.indexed_const_def)
    then show ?thesis unfolding ring_hom_to_IP_ring_hom_def
      by blast
  qed
  show 4: " P Q. P  carrier (Pring R I) 
           Q  carrier (Pring R I)  ring_hom_to_IP_ring_hom φ (P Pring R IQ) = ring_hom_to_IP_ring_hom φ P Pring S Iring_hom_to_IP_ring_hom φ Q"
    using assms(1) assms(2) ring_hom_to_IP_ring_hom_add by blast
  show 5: " i P. i  I 
           P  carrier (Pring R I) 
           ring_hom_to_IP_ring_hom φ (P  i) = ring_hom_to_IP_ring_hom φ P Pring S Iring_hom_to_IP_ring_hom φ (mset_to_IP R {#i#})"
  proof-
    fix i P
    assume i0: "i  I"
    assume P0: "P  carrier (Pring R I)"
    show  "ring_hom_to_IP_ring_hom φ (P  i) = ring_hom_to_IP_ring_hom φ P Pring S Iring_hom_to_IP_ring_hom φ (mset_to_IP R {#i#})"
      unfolding ring_hom_to_IP_ring_hom_def
    proof
      fix m
      show " φ ((P  i) m) = ((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m"
      proof(cases "i ∈# m")
        case True
        have "(λm. φ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}"
        proof
          fix m
          show "φ (mset_to_IP R {#i#} m) = mset_to_IP S {#i#} m"
            apply(cases "{#i#} = m")
             apply (metis (mono_tags, lifting) assms(1) assms(2) cring.axioms(1) local.ring_axioms
                mset_to_IP_def one_mset_to_IP ring.Pring_one ring.Pring_one ring.one_mset_to_IP
                ring_hom_to_IP_ring_hom_def ring_hom_to_IP_ring_hom_one)
          proof-
            assume "{#i#}  m"
            then have LHS: "(mset_to_IP R {#i#} m) = 𝟬R⇙"
              by (metis mset_to_IP_simp')
            have  RHS :"mset_to_IP S {#i#} m = 𝟬S⇙"
              by (metis {#i#}  m mset_to_IP_def)
            have "φ 𝟬R= 𝟬S⇙"
              using assms(1) assms(2) cring.axioms(1) local.ring_axioms
                ring_hom_ring.homh ring_hom_zero by blast
            then show ?thesis
              using  LHS RHS
              by presburger
          qed
        qed
        then have RHS: "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
              ((λm. φ (P m)) Pring S Imset_to_IP S {#i#}) m"
          by presburger
        then have RHS': "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
              (ring.indexed_pmult S (λm. φ (P m)) i) m"
        proof-
          have 0: "(λm. φ (P m))  Pring_set S I"
            using ring_hom_to_IP_ring_hom_closed[of S φ P I] ring.Pring_car[of S I] assms
            unfolding ring_hom_to_IP_ring_hom_def
            using P0 cring.axioms(1) by blast
          show ?thesis using ring.poly_index_mult[of S "(λm. φ (P m))" I i]
            by (metis "0" (λm. φ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}
                assms(1) cring.axioms(1) i0 ring.Pring_mult)
        qed
        then have RHS'': "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
            (λm. φ (P m)) (m - {#i#})" using ring.indexed_pmult_def[of S "(λm. φ (P m))" i] True
          by (metis assms(1) cring.axioms(1))
        then show ?thesis
          by (metis True indexed_pmult_def)
      next
        case False
        have LHS: "((P  i) m) = 𝟬R⇙"
          using False
          by (meson indexed_pmult_def)
        have "(λm. φ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}"
        proof
          fix m
          show "φ (mset_to_IP R {#i#} m) = mset_to_IP S {#i#} m"
            apply(cases "{#i#} = m")
             apply (metis (mono_tags, lifting) assms(1) assms(2) cring.axioms(1) local.ring_axioms
                mset_to_IP_def one_mset_to_IP ring.Pring_one ring.Pring_one ring.one_mset_to_IP
                ring_hom_to_IP_ring_hom_def ring_hom_to_IP_ring_hom_one)
          proof-
            assume "{#i#}  m"
            then have LHS: "(mset_to_IP R {#i#} m) = 𝟬R⇙"
              by (metis mset_to_IP_simp')
            have  RHS :"mset_to_IP S {#i#} m = 𝟬S⇙"
              by (metis {#i#}  m mset_to_IP_def)
            have "φ 𝟬R= 𝟬S⇙"
              using assms(1) assms(2) cring.axioms(1) local.ring_axioms
                ring_hom_ring.homh ring_hom_zero by blast
            then show ?thesis
              using  LHS RHS
              by presburger
          qed
        qed
        then have RHS: "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
              ((λm. φ (P m)) Pring S Imset_to_IP S {#i#}) m"
          by presburger
        then have RHS': "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
              (ring.indexed_pmult S (λm. φ (P m)) i) m"
        proof-
          have 0: "(λm. φ (P m))  Pring_set S I"
            using ring_hom_to_IP_ring_hom_closed[of S φ P I] ring.Pring_car[of S I] assms
            unfolding ring_hom_to_IP_ring_hom_def
            using P0 cring.axioms(1) by blast
          show ?thesis using ring.poly_index_mult[of S "(λm. φ (P m))" I i]
            by (metis "0" (λm. φ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}
                assms(1) cring.axioms(1) i0 ring.Pring_mult)
        qed
        then have RHS'': "((λm. φ (P m)) Pring S I(λm. φ (mset_to_IP R {#i#} m))) m =
              𝟬S⇙"
          using False
          by (metis assms(1) cring.axioms(1) ring.indexed_pmult_def)
        then show ?thesis
          using LHS False assms ring_hom_zero[of φ R S]
          by (metis cring.axioms(1) local.ring_axioms ring_hom_ring.homh)
      qed
    qed
  qed
  show "k Q. k  carrier R 
           Q  carrier (Pring R I) 
           ring_hom_to_IP_ring_hom φ (poly_scalar_mult R k Q) = ring_hom_to_IP_ring_hom φ (indexed_const k) Pring S Iring_hom_to_IP_ring_hom φ Q"
  proof-
    fix k Q
    assume A0: "k  carrier R"
    assume A1: " Q  carrier (Pring R I)"
    show "ring_hom_to_IP_ring_hom φ (poly_scalar_mult R k Q) =
          ring_hom_to_IP_ring_hom φ (indexed_const k) Pring S Iring_hom_to_IP_ring_hom φ Q"
    proof
      fix x
      show "ring_hom_to_IP_ring_hom φ (poly_scalar_mult R k Q) x = (ring_hom_to_IP_ring_hom φ (indexed_const k) Pring S Iring_hom_to_IP_ring_hom φ Q) x"
      proof-
        have LHS: "ring_hom_to_IP_ring_hom φ (poly_scalar_mult R k Q) x = φ (k RQ x)"
          by (metis poly_scalar_mult_def ring_hom_to_IP_ring_hom_def)
        then have LHS': "ring_hom_to_IP_ring_hom φ (poly_scalar_mult R k Q) x = (φ k) Sφ (Q x)"
        proof-
          have "Q x  carrier  R"
            using A1 Pring_car local.ring_axioms ring.Pring_cfs_closed by blast
          then show ?thesis
            using LHS assms ring_hom_mult[of φ R S k "Q x"]
            by (metis A0 ring_hom_ring.homh)
        qed
        have RHS: "(ring_hom_to_IP_ring_hom φ (indexed_const k) Pring S Iring_hom_to_IP_ring_hom φ Q) x =
               (φ k) S(φ (Q x))"
        proof-
          have 0: "ring_hom_to_IP_ring_hom φ (indexed_const k)= ring.indexed_const S (φ k)"
            by (simp add: A0 assms(1) assms(2) ring_hom_to_IP_ring_hom_constant)
          have 1: "(φ k)  carrier S"
            by (meson A0 assms(2) ring_hom_closed ring_hom_ring.homh)
          have 2: "ring_hom_to_IP_ring_hom φ Q  carrier (Pring S I)"
            using A1 assms ring_hom_to_IP_ring_hom_closed
            by blast
          have 3: "(ring.indexed_const S (φ k) Pring S Iring_hom_to_IP_ring_hom φ Q) x = (φ k) S(φ (Q x))"
            using assms(1) 1 2
                 cring.indexed_const_Pring_mult(1)[of S "φ k" "ring_hom_to_IP_ring_hom φ Q" I x]
                 ring_hom_to_IP_ring_hom_def[of φ Q x]
            by presburger
          then show ?thesis
            by (metis "0")
        qed
        then show ?thesis
          using LHS'
          by metis
      qed
    qed
  qed
qed

lemma ring_hom_to_IP_ring_hom_smult:
  assumes "cring S"
  assumes "ring_hom_ring R S φ"
  assumes "P  carrier (Pring R I)"
  assumes "a  carrier R"
  shows "ring_hom_to_IP_ring_hom φ (a Pring R IP) =
                 φ a Pring S I(ring_hom_to_IP_ring_hom φ P)"
proof fix m
  have 0: "φ ((a Pring R IP) m) = φ (a  (P m))"
    using assms
    by (metis Pring_smult_cfs)
  hence 1: "φ ((a Pring R IP) m) = φ a Sφ (P m)"
    using assms ring_hom_mult[of φ R S]
    by (metis Pring_carrier_coeff' ring_hom_ring.homh)
  have 2: "(φ a Pring S I(λm. φ (P m))) m = φ a Sφ (P m)"
    using assms ring.Pring_smult[of S I]
    unfolding poly_scalar_mult_def cring_def
    by presburger
  show "ring_hom_to_IP_ring_hom φ (a Pring R IP) m =
         (φ a Pring S Iring_hom_to_IP_ring_hom φ P) m"
    unfolding ring_hom_to_IP_ring_hom_def using assms 1 2
    by presburger
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹A Universal Property for Indexed Polynomial Rings›
(**************************************************************************************************)
(**************************************************************************************************)

lemma Pring_universal_prop_0:
  assumes a_cring: "cring S"
  assumes index_map: "closed_fun S g"
  assumes ring_hom: "ring_hom_ring R S φ"
  assumes "ψ = (total_eval S g)  (ring_hom_to_IP_ring_hom φ)"
  shows "(ring_hom_ring (Pring R I) S ψ)"
        "(i  I. ψ (mset_to_IP R {#i#}) = g i)"
        "(a  carrier R. ψ (indexed_const a) = φ a)"
        " ρ. (ring_hom_ring (Pring R I) S ρ) 
          (i  I. ρ (mset_to_IP R {#i#}) = g i) 
          (a  carrier R. ρ (indexed_const a) = φ a) 
          (x  carrier (Pring R I). ρ x = ψ x)"
proof-
  have 0: " (ring_hom_to_IP_ring_hom φ)  ring_hom (Pring R I) (Pring S I)"
    using a_cring ring_hom ring_hom_ring.homh ring_hom_to_IP_ring_hom_is_hom
    by blast
  have 1: "(total_eval S g)  ring_hom (Pring S I) S "
    using a_cring cring.total_eval_ring_hom index_map ring_hom_ring.homh
    by blast
  show P0: "ring_hom_ring (Pring R I) S ψ "
    using ring_hom_trans 0 1 Pring_is_ring a_cring assms(4) cring.axioms(1) ring_hom_ringI2
    by blast
  show P1: "iI. ψ (mset_to_IP R {#i#}) = g i"
  proof
    fix i assume Pi: "i  I"
    show "ψ (mset_to_IP R {#i#}) = g i"
    proof-
      have 0: "ψ (mset_to_IP R {#i#}) = (total_eval S g) ( (ring_hom_to_IP_ring_hom φ) (mset_to_IP R {#i#}))"
        by (simp add: assms(4))
      have "( (ring_hom_to_IP_ring_hom φ) (mset_to_IP R {#i#})) = (mset_to_IP S {#i#})"
        by (simp add: a_cring ring_hom ring_hom_to_IP_ring_hom_monom)
      then have "ψ (mset_to_IP R {#i#}) = (total_eval S g) (mset_to_IP S {#i#})"
        by (simp add: 0)
      then show ?thesis
        by (simp add: a_cring cring.total_eval_var index_map)
    qed
  qed
  show P2: "acarrier R. ψ (indexed_const a) = φ a"
  proof
    fix a assume A: "a  carrier R"
    show "ψ (indexed_const a) = φ a"
    proof-
      have 0: "ring_hom_to_IP_ring_hom φ (indexed_const a) = ring.indexed_const S (φ a)"
        by (simp add: A a_cring ring_hom ring_hom_to_IP_ring_hom_constant)
      have 1: "total_eval S g (ring.indexed_const S (φ a)) = φ a"
        by (meson A a_cring cring.total_eval_const ring_hom ring_hom_closed ring_hom_ring.homh)
      show ?thesis
        using assms 0 1
        by (simp add: "0" index_map)
    qed
  qed
  show " ρ. (ring_hom_ring (Pring R I) S ρ) 
          (i  I. ρ (mset_to_IP R {#i#}) = g i) 
          (a  carrier R. ρ (indexed_const a) = φ a) 
          (x  carrier (Pring R I). ρ x = ψ x)"
  proof
    fix ρ
    show "ring_hom_ring (Pring R I) S ρ  (iI. ρ (mset_to_IP R {#i#}) = g i)  (acarrier R. ρ (indexed_const a) = φ a) 
         (xcarrier (Pring R I). ρ x = ψ x)"
    proof
        assume A: "(ring_hom_ring (Pring R I) S ρ) 
          (i  I. ρ (mset_to_IP R {#i#}) = g i) 
          (a  carrier R. ρ (indexed_const a) = φ a)"
        show "(x  carrier (Pring R I). ρ x = ψ x)"
        proof
          fix x assume B: "x  carrier (Pring R I)"
          show "ρ x = ψ x"
            apply(rule indexed_pset.induct[of x I "carrier R"])
            using B Pring_car apply blast
              apply (metis A P2)
          proof-
            show "P Q. P  Pring_set R I  ρ P = ψ P  Q  Pring_set R I  ρ Q = ψ Q  ρ (P  Q) = ψ (P  Q)"
            proof-
              fix P Q
              assume A0: "P  Pring_set R I " "ρ P = ψ P" " Q  Pring_set R I" " ρ Q = ψ Q "
              show "ρ (P  Q) = ψ (P  Q)"
                using A A0 ring_hom_add[of ψ "Pring R I" S P Q] ring_hom_add[of ρ "Pring R I" S P Q] P0
                by (metis Pring_add Pring_car ring_hom_ring.homh)
            qed
            show "P i. P  Pring_set R I  ρ P = ψ P  i  I  ρ (P  i) = ψ (P  i)"
            proof-
              fix P i
              assume A0: " P  Pring_set R I" " ρ P = ψ P" "i  I"
              show "ρ (P  i) = ψ (P  i)"
              proof-
                have "ρ (P Pring R I(mset_to_IP R {#i#})) = ψ (P Pring R I(mset_to_IP R {#i#}))"
                  using A A0 ring_hom_mult[of ψ "Pring R I" S P "(mset_to_IP R {#i#})"]
                        ring_hom_mult[of ρ "Pring R I" S P "(mset_to_IP R {#i#})"] P0
                  by (metis P1 Pring_car Pring_one Pring_one_closed indexed_pset.indexed_pmult
                      monom_add_mset one_mset_to_IP ring_hom_ring.homh)
                then show ?thesis
                  by (metis A0(1) A0(3) Pring_mult poly_index_mult)
              qed
            qed
          qed
        qed
    qed
  qed
qed

end

definition close_fun :: "'c set  ('e, 'f) ring_scheme  ('c  'e)  ('c  'e)" where
"close_fun I S g = (λi. (if i  I then g i else  𝟬S))"

context cring
begin

lemma close_funE:
  assumes "cring S"
  assumes "g  I  carrier S"
  shows "closed_fun S (close_fun I S g)"
  apply(rule cring.closed_funI)
  apply (simp add: assms(1))
  by (metis close_fun_def PiE assms(1) assms(2) cring.cring_simprules(2))

end

definition indexed_poly_induced_morphism ::
 "'c set  ('e, 'f) ring_scheme  ('a, 'e) ring_hom  ('c  'e)  (('a,'c) mvar_poly, 'e) ring_hom" where
"indexed_poly_induced_morphism I S φ g = (total_eval S (close_fun I S g))  (ring_hom_to_IP_ring_hom φ)"

context cring
begin

lemma Pring_universal_prop:
  assumes a_cring: "cring S"
  assumes index_map: "g  I  carrier S"
  assumes ring_hom: "ring_hom_ring R S φ"
  assumes "ψ = indexed_poly_induced_morphism I S φ g"
  shows "(ring_hom_ring (Pring R I) S ψ)"
        "(i  I. ψ (mset_to_IP R {#i#}) = g i)"
        "(a  carrier R. ψ (indexed_const a) = φ a)"
        " ρ. (ring_hom_ring (Pring R I) S ρ) 
          (i  I. ρ (mset_to_IP R {#i#}) = g i) 
          (a  carrier R. ρ (indexed_const a) = φ a) 
          (x  carrier (Pring R I). ρ x = ψ x)"
proof-
  obtain g' where g'_def: "g' =  (close_fun I S g)"
    by simp
  have 0: "closed_fun S g'"
    using close_funE a_cring g'_def index_map
    by blast
  show "(ring_hom_ring (Pring R I) S ψ)" using assms 0
    using cring.Pring_universal_prop_0(1) indexed_poly_induced_morphism_def g'_def is_cring
    by blast
  show "(i  I. ψ (mset_to_IP R {#i#}) = g i)"
  proof-
    have "(i  I. ψ (mset_to_IP R {#i#}) = g' i)"
      apply(intro Pring_universal_prop_0[of S _ φ] assms)
      unfolding assms indexed_poly_induced_morphism_def g'_def
      using assms 0 g'_def apply fastforce
      by auto
    thus ?thesis unfolding g'_def using assms
      by (simp add: close_fun_def)
  qed
  show "acarrier R. ψ (indexed_const a) = φ a"
    using 0  indexed_poly_induced_morphism_def Pring_universal_prop_0(3) a_cring assms(4) g'_def ring_hom
    by blast
  show "ρ. ring_hom_ring (Pring R I) S ρ  (iI. ρ (mset_to_IP R {#i#}) = g i)  (acarrier R. ρ (indexed_const a) = φ a) 
        (xcarrier (Pring R I). ρ x = ψ x)"
  proof-
    have "ρ. ring_hom_ring (Pring R I) S ρ  (iI. ρ (mset_to_IP R {#i#}) = g' i)  (acarrier R. ρ (indexed_const a) = φ a) 
        (xcarrier (Pring R I). ρ x = ψ x)"
      using assms 0 Pring_universal_prop_0(4)  g'_def
      unfolding indexed_poly_induced_morphism_def
      by blast
    then show ?thesis
      using g'_def
      unfolding close_fun_def
      by meson
  qed
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Mapping Mulitvariate Polynomials over a Single Variable to Univariate Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

text‹Constructor for multisets which have one distinct element›

definition nat_to_mset :: "'c  nat  'c monomial" where
"nat_to_mset i n = Abs_multiset (λj. if (j = i) then n else 0)"

lemma nat_to_msetE: "count (nat_to_mset i n) i = n"
  unfolding nat_to_mset_def by simp

lemma nat_to_msetE':
  assumes "j  i"
  shows "count (nat_to_mset i n) j = 0"
  unfolding nat_to_mset_def using assms by simp

lemma nat_to_mset_add: "nat_to_mset i (n + m) = (nat_to_mset i n) + (nat_to_mset i m)"
  apply(rule multiset_eqI)
  by (metis add.right_neutral nat_to_msetE nat_to_msetE' count_union)

lemma nat_to_mset_inj:
  assumes "n  m"
  shows "(nat_to_mset i n)  (nat_to_mset i m)"
  using assms
  by (metis nat_to_msetE)

lemma nat_to_mset_zero: "nat_to_mset i 0 = {#}"
  by (metis add.right_neutral add_cancel_right_right nat_to_mset_add)

lemma nat_to_mset_Suc: "nat_to_mset i (Suc n) = add_mset i (nat_to_mset i n)"
  using nat_to_mset_add[of i n 1]
  by (simp add: multiset_eqI nat_to_msetE nat_to_msetE')

lemma nat_to_mset_Pring_singleton:
  assumes "cring R"
  assumes "P  carrier (Pring R {i})"
  assumes "m  monomials_of R P"
  shows "m = nat_to_mset i (count m i)"
proof-
  have " j. count m j = count (nat_to_mset i (count m i)) j"
  proof-
    fix j
    show "count m j = count (nat_to_mset i (count m i)) j"
      apply(cases "j = i")
       apply (simp add: nat_to_msetE; fail)
    proof-
      assume A: "j i"
      have P0: "set_mset m  {i}"
        using assms
        by (metis cring.axioms(1) ring.Pring_car ring.mset_to_IP_indices)
      then have "count m j = 0"
        using A assms
        by (metis count_inI empty_iff singletonD subset_singletonD)
      then show " count m j = count (nat_to_mset i (count m i)) j"
        by (simp add: A nat_to_msetE')
    qed
  qed
  then show ?thesis
    using multiset_eqI by blast
qed

definition IP_to_UP :: "'d  ('e, 'd) mvar_poly  'e u_poly" where
"IP_to_UP i P  = (λ (n::nat). P (nat_to_mset i n))"

lemma IP_to_UP_closed:
  assumes "cring R"
  assumes "P  carrier (Pring R {i::'c})"
  shows "IP_to_UP i P  carrier (UP R)"
proof-
  have "IP_to_UP i P  up R"
    apply(rule mem_upI)
    using assms
     apply (metis cring_def IP_to_UP_def ring.Pring_carrier_coeff')
       unfolding bound_def IP_to_UP_def
       apply(rule ccontr)
     proof-
       assume A: "n. m>n. P (nat_to_mset i m) = 𝟬R⇙"
       then have 0: "n. m> n. (nat_to_mset i m)  monomials_of R P"
         by (meson assms(1) cring_def ring.complement_of_monomials_of)
       have "¬ finite {m. (nat_to_mset i m)  monomials_of R P }"
       proof
         assume "finite {m. nat_to_mset i m  monomials_of R P}"
         then have "n. m>n.  m  {m. nat_to_mset i m  monomials_of R P}"
           by (meson finite_nat_set_iff_bounded nat_less_le order.strict_trans)
         then have "n. m>n.  nat_to_mset i m  monomials_of R P"
           by (simp add: monomials_of_def)
         then show False using 0
           by blast
       qed
       then obtain S where S_def: "infinite (S::nat set)  (m  S. (nat_to_mset i m)  monomials_of R P)"
         by blast
       have "inj_on (nat_to_mset i) S"
         using inj_onI[of S "nat_to_mset i"]
         by (meson nat_to_mset_inj)
       then have 1: "infinite (nat_to_mset i ` S)"
         using S_def finite_imageD
         by blast
       have 2: "(nat_to_mset i ` S)  (monomials_of R P)"
         using S_def
         by blast
       then have "infinite (monomials_of R P)"
         using S_def "1" finite_subset
         by blast
       then show False using assms
         by (metis Pring_def cring_def partial_object.select_convs(1) ring.monomials_finite)
     qed
     then show ?thesis
       by (metis (no_types, lifting) UP_def partial_object.select_convs(1))
qed

lemma IP_to_UP_var:
  shows "IP_to_UP i (mset_to_IP R {#i#}) = X_poly R"
proof
  have UP: "UP_cring R"
    by (simp add: UP_cring_def cring_axioms)
  fix x
  show "IP_to_UP i (mset_to_IP R {#i#}) x = X_poly R x"
  proof(cases "x = 1")
    case True
    then have RHS: "X_poly R x = 𝟭R⇙"
      unfolding X_poly_def using UP UP_ring.cfs_monom[of R]
      unfolding UP_ring_def
      using local.ring_axioms one_closed by presburger
    have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x
        = mset_to_IP R ((add_mset i) {#}) (nat_to_mset i x)"
      unfolding IP_to_UP_def
      by blast
    have "(nat_to_mset i x) = {#i#}"
    proof-
      have "n. count (nat_to_mset i x) n = count {#i#} n"
        by (simp add: True nat_to_msetE nat_to_msetE')
      then show ?thesis
        using multi_count_eq
        by blast
    qed
    then have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x
       =  mset_to_IP R ((add_mset i) {#}) {#i#}"
      using LHS by presburger
    then show ?thesis
      unfolding deg_zero_cf_def
      by (metis RHS mset_to_IP_simp)
  next
    case False
    then have RHS: "X_poly R x = 𝟬R⇙"
      unfolding X_poly_def
      using UP UP_ring.cfs_monom[of R]
      unfolding UP_ring_def
      using local.ring_axioms one_closed by presburger
    have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x
        = mset_to_IP R ((add_mset i) {#}) (nat_to_mset i x)"
      unfolding IP_to_UP_def
      by blast
    then show ?thesis
      unfolding deg_zero_cf_def
      by (metis False RHS count_single mset_to_IP_def nat_to_msetE)
  qed
qed

end

context UP_cring
begin

lemma IP_to_UP_monom:
  shows "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) = ((X_poly R)[^]UP Rn) "
proof
  fix x
  show "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = (X_poly R [^]UP Rn) x"
  proof(cases "x = n")
    case True
    have RHS: "(X_poly R [^]UP Rn) x = 𝟭R⇙"
      unfolding X_poly_def
      by (metis P.nat_pow_closed P.nat_pow_eone P_def R.one_closed True UP_cring.X_closed
          UP_cring.monom_coeff UP_one_closed UP_r_one deg_one is_UP_cring monom_one monom_rep_X_pow
          to_poly_inverse to_poly_mult_simp(2))
    have LHS:  "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = 𝟭R⇙"
      by (metis R.mset_to_IP_simp True IP_to_UP_def)
    then show ?thesis
      using RHS by presburger
  next
    case False
    have 0: "x y::nat. nat_to_mset x = nat_to_mset y  x = y"
    proof-
      fix a b::nat assume A:  "nat_to_mset a = nat_to_mset b"
      then show "a = b" unfolding nat_to_mset_def
        by (metis A nat_to_msetE nat_to_msetE' zero_neq_one)
    qed
    have 1: "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x  = (if nat_to_mset i x = nat_to_mset i n then 𝟭 else 𝟬) "
      unfolding IP_to_UP_def mset_to_IP_def
      by blast
    hence 2: "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x  = (if x = n then 𝟭 else 𝟬) "
      using 0
      by (meson False nat_to_mset_inj)
    have 3: "(X_poly R [^]UP Rn) x = 𝟬"
      unfolding X_poly_def using False
      by (smt ctrm_degree P.nat_pow_closed P.nat_pow_eone P.r_null P_def R.one_closed
          UP_cring.ltrm_of_X UP_cring.ltrm_rep_X_pow UP_cring.X_closed UP_cring.monom_coeff
          UP_r_one UP_zero_closed X_mult_cf cfs_closed cfs_monom deg_nzero_nzero is_UP_cring
          monom_closed monom_one to_poly_inverse to_poly_mult_simp(2))
    thus ?thesis using 2 1
      using False by presburger
  qed
qed

lemma IP_to_UP_one:
 "IP_to_UP i 𝟭Pring R {i}= 𝟭UP R⇙"
proof
  fix x
  show "IP_to_UP i 𝟭Pring R {i}x = 𝟭UP Rx"
  proof(cases "x = 0")
    case True
    have RHS: "𝟭UP Rx = 𝟭R⇙"
      using P_def True cfs_one by presburger
    have "𝟭Pring R {i}= (λ m. if m = {#} then 𝟭Relse 𝟬R)"
      by (metis R.Pring_one R.indexed_const_def)
    then have "IP_to_UP i 𝟭Pring R {i}= IP_to_UP i (λ m. if m = {#} then 𝟭Relse 𝟬R)"
      by presburger
    then have LHS: "IP_to_UP i 𝟭Pring R {i}x = 𝟭R⇙"
      by (smt True count_empty IP_to_UP_def multi_count_eq nat_to_msetE nat_to_msetE')
    then show ?thesis
      using RHS by presburger
  next
    case False
    have RHS: "𝟭UP Rx = 𝟬R⇙"
      by (smt False UP_def monoid.simps(2))
    show ?thesis
      using False count_empty
            nat_to_msetE
            ring.indexed_const_def
      unfolding IP_to_UP_def
      by (metis R.Pring_one R.ring_axioms RHS)
  qed
qed

lemma IP_to_UP_zero:
 "IP_to_UP i 𝟬Pring R {i}= 𝟬UP R⇙"
proof
  fix x
  show "IP_to_UP i 𝟬Pring R {i}x = 𝟬UP Rx"
    unfolding IP_to_UP_def using R.Pring_zero
    by (metis P_def R.indexed_zero_def cfs_zero)
qed

lemma IP_to_UP_add:
  assumes " x  carrier (Pring R {i})"
  assumes " y  carrier (Pring R {i})"
  shows " IP_to_UP i (x Pring R {i}y) =
          IP_to_UP i x UP RIP_to_UP i y"
proof
  fix n
  have LHS: "IP_to_UP i (x Pring R {i}y) n = (x Pring R {i}y) (nat_to_mset i n)"
    by (meson IP_to_UP_def)
  then have LHS: "IP_to_UP i (x Pring R {i}y) n = x (nat_to_mset i n) Ry (nat_to_mset i n)"
    using assms  unfolding IP_to_UP_def
    by (metis R.Pring_add R.indexed_padd_def)
  have RHS: "(IP_to_UP i x UP RIP_to_UP i y) n =
   (IP_to_UP i x) n R(IP_to_UP i y) n"
    using assms UP_ring.cfs_add IP_to_UP_closed
    by (simp add: UP_ring.cfs_add R_cring cring.IP_to_UP_closed is_UP_ring)
  then   show "IP_to_UP i (x Pring R {i}y) n = (IP_to_UP i x UP RIP_to_UP i y) n"
    using assms
    by (metis LHS IP_to_UP_def)
qed

lemma IP_to_UP_indexed_const:
  assumes "k  carrier R"
  shows "IP_to_UP i (ring.indexed_const R k) = to_polynomial R k"
proof
  fix x
  show "IP_to_UP i (ring.indexed_const R k) x = to_polynomial R k x"
  proof(cases "x = 0")
    case True
    have LHS: "IP_to_UP i (ring.indexed_const R k) x = k"
      using True unfolding IP_to_UP_def
      by (metis R.indexed_const_def nat_to_mset_zero)
    then show ?thesis
      using assms
      unfolding to_polynomial_def
      using True to_polynomial_def
      by (metis UP_ring.cfs_monom is_UP_ring)
  next
    case False
    have LHS: "IP_to_UP i (ring.indexed_const R k) x = 𝟬R⇙"
      using False unfolding IP_to_UP_def
      by (metis R.indexed_const_def nat_to_mset_inj nat_to_mset_zero)
    then show ?thesis
      using assms
      unfolding to_polynomial_def
      using False UP_cring.intro UP_cring.monom_coeff UP_cring.monom_rep_X_pow
      using P_def cfs_monom by presburger
  qed
qed

lemma IP_to_UP_indexed_pmult:
  assumes "p  carrier (Pring R {i})"
  shows "IP_to_UP i (ring.indexed_pmult R p i) = (IP_to_UP i p) UP R(X_poly R)"
proof
  fix n
  have 0: "IP_to_UP i p  carrier (UP R)"
    by (simp add: R_cring assms cring.IP_to_UP_closed)
  show "IP_to_UP i (ring.indexed_pmult R p i) n = (IP_to_UP i p UP RX_poly R) n"
  proof(cases "n = 0")
    case True
    then have RHS: "(IP_to_UP i p UP RX_poly R) n = 𝟬R⇙"
      by (metis (no_types, lifting) "0" lcf_closed One_nat_def P.r_null P_def R.r_null
          UP_cring.ltrm_of_X UP_cring.cfs_monom_mult UP_cring.cfs_monom_mult_l UP_zero_closed
          X_closed cfs_times_X deg_leE deg_nzero_nzero is_UP_cring lessI neq0_conv plus_1_eq_Suc to_poly_inverse)
    have LHS: "IP_to_UP i (ring.indexed_pmult R p i) n = ring.indexed_pmult R p i (nat_to_mset i n)"
      unfolding IP_to_UP_def
      by blast
    then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n =
                  (p Pring R {i}(mset_to_IP R {#i#})) (nat_to_mset i n)"
      using  assms(1) ring.Pring_car ring.Pring_mult
                ring.poly_index_mult singletonI
      by (metis R.ring_axioms)
  then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n =
                  (p Pring R {i}(mset_to_IP R {#i#})) {#}"
    using True
    by (metis nat_to_mset_zero)
  then show ?thesis using RHS  LHS True assms(1) nat_to_mset_zero ring.indexed_pmult_def
    by (metis R.ring_axioms empty_iff set_mset_empty)
  next
    case False
    then have RHS: "(IP_to_UP i p UP RX_poly R) n = (IP_to_UP i p ) (n -1)"
      using "0" Suc_diff_1 Suc_eq_plus1
           assms(1) bot_nat_def IP_to_UP_def nat_neq_iff not_less0
      by (metis (no_types, lifting) P_def UP_cring X_closed cfs_times_X cring.cring_simprules(14))
    have LHS: "IP_to_UP i (ring.indexed_pmult R p i) n = ring.indexed_pmult R p i (nat_to_mset i n)"
      unfolding IP_to_UP_def
      by blast
    then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n =
                  (p Pring R {i}(mset_to_IP R {#i#})) (nat_to_mset i n)"
      using assms(1) ring.Pring_car ring.Pring_mult
      ring.poly_index_mult singletonI
      by (metis R.ring_axioms)
    then have LHS'': "IP_to_UP i (ring.indexed_pmult R p i) n =
                  (p Pring R {i}(mset_to_IP R {#i#})) (add_mset i (nat_to_mset i (n-1))) "
      by (metis False Suc_diff_1 nat_to_mset_Suc neq0_conv)
    then show ?thesis using RHS unfolding IP_to_UP_def
      by (metis (no_types, lifting) False R.indexed_pmult_def Suc_diff_1 add_mset_remove_trivial add_mset_remove_trivial_If multi_self_add_other_not_self nat_to_mset_Suc neq0_conv)
  qed
qed

lemma IP_to_UP_ring_hom:
  shows "ring_hom_ring (Pring R {i}) (UP R) (IP_to_UP i)"
  apply(rule cring.Pring_morphism)
      apply (simp add: R_cring; fail)
  using P_def UP_cring apply blast
      apply (simp add: R.IP_to_UP_closed R_cring; fail)
      apply (meson IP_to_UP_one)
            apply (meson IP_to_UP_zero)
    apply (meson IP_to_UP_add)
  apply (metis R.IP_to_UP_var IP_to_UP_indexed_pmult singletonD)
proof-
  fix k Q
  assume A0: "k  carrier R"
  assume A1: "Q  carrier (Pring R {i})"
  show "IP_to_UP i (poly_scalar_mult R k Q) =
           IP_to_UP i (ring.indexed_const R k) UP RIP_to_UP i Q"
    unfolding poly_scalar_mult_def
  proof
    fix x
    show "IP_to_UP i (λm. k RQ m) x =
         (IP_to_UP i (ring.indexed_const R k) UP RIP_to_UP i Q) x"
    proof-
      have LHS: "IP_to_UP i (λm. k RQ m) x = k RQ (nat_to_mset i x)"
        unfolding IP_to_UP_def
        by blast
      have RHS: "(IP_to_UP i (ring.indexed_const R k) UP RIP_to_UP i Q) x =
                (to_polynomial R k UP RIP_to_UP i Q) x"
        by (metis A0 IP_to_UP_indexed_const)
      have RHS': "(IP_to_UP i (ring.indexed_const R k) UP RIP_to_UP i Q) x =
                k R((IP_to_UP i Q) x)"
      proof-
        have 0: "deg R (to_polynomial R k) = 0"
          using A0 degree_to_poly by blast
        have 1: "(IP_to_UP i Q)  carrier (UP R)"
          using IP_to_UP_closed  unfolding P_def
          by (simp add: A1 R.IP_to_UP_closed R_cring)
        then show ?thesis
        proof -
          have "UP_cring R  IP_to_UP i Q  carrier (UP R)"
            using "1" is_UP_cring by blast
          then show ?thesis
            by (metis A0 UP_cring.to_poly_mult_simp(1) UP_ring.UP_mult_closed UP_ring.coeff_simp UP_ring.coeff_smult UP_ring.monom_closed IP_to_UP_indexed_const is_UP_ring to_polynomial_def)
        qed
      qed
      then show ?thesis
        by (metis IP_to_UP_def)
    qed
  qed
qed

lemma IP_to_UP_ring_hom_inj:
  shows "inj_on  (IP_to_UP i) (carrier (Pring R {i}))"
proof
  fix x y
  assume A: "x  carrier (Pring R {i})" "y  carrier (Pring R {i}) "
  assume B: "IP_to_UP i x = IP_to_UP i y"
  show "x = y"
  proof
    fix a
    show "x a = y a"
    proof(cases "set_mset a  {i}")
      case True
      then obtain n where "a = (nat_to_mset i n)"
        by (metis count_eq_zero_iff insert_subset multiset_eqI nat_to_msetE nat_to_msetE'
            set_eq_subset singletonD singleton_insert_inj_eq' subset_insertI subset_refl)
      then have LHS: "x a = IP_to_UP i x n"
        by (metis IP_to_UP_def)
      then show ?thesis
        by (metis B a = nat_to_mset i n IP_to_UP_def)
    next
      case False
      then show ?thesis
        using ring.Pring_set_zero[of R y "{i}" a] ring.Pring_set_zero[of R x "{i}" a] A
        by (metis R.Pring_car R.ring_axioms)
    qed
  qed
qed

lemma IP_to_UP_scalar_mult:
  assumes "a  carrier R"
  assumes "p  carrier (Pring R {i})"
  shows "(IP_to_UP i (a Pring R {i}p)) = aUP R(IP_to_UP i p)"
  apply(rule ring.indexed_pset.induct[of R p "{i}" "carrier R"])
  apply (simp add: R.ring_axioms; fail)
  using R.Pring_car assms(2) apply blast
  apply (metis IP_to_UP_indexed_const P_def R.m_closed R.poly_scalar_mult_const R.ring_axioms assms(1) ring.Pring_smult to_poly_closed to_poly_mult to_poly_mult_simp(1))
proof-
  show "P Q. P  Pring_set R {i} 
           IP_to_UP i (a Pring R {i}P) = a UP RIP_to_UP i P 
           Q  Pring_set R {i}  IP_to_UP i (a Pring R {i}Q) = a UP RIP_to_UP i Q  IP_to_UP i (a Pring R {i}(P  Q)) = a UP RIP_to_UP i (P  Q)"
  proof-
    fix p Q
    assume A0: "p  Pring_set R {i}"
              "IP_to_UP i (a Pring R {i}p) = a UP RIP_to_UP i p"
              " Q  Pring_set R {i}"
              "IP_to_UP i (a Pring R {i}Q) = a UP RIP_to_UP i Q"
    show "IP_to_UP i (a Pring R {i}(p  Q)) = a UP RIP_to_UP i (p  Q)"
    proof-
      have "(a Pring R {i}(p  Q)) = a Pring R {i}p Pring R {i}a Pring R {i}Q"
        by (metis A0(1) A0(3) R.Pring_add R.Pring_car R.Pring_smult_r_distr assms(1))
      then show ?thesis using A0
        by (metis IP_to_UP_add P_def R.IP_to_UP_closed R.Pring_add R.Pring_car R.Pring_smult_closed R_cring UP_smult_r_distr assms(1))
      qed
  qed
  show " P ia. P  Pring_set R {i}  IP_to_UP i (a Pring R {i}P) = a UP RIP_to_UP i P  ia  {i}  IP_to_UP i (a Pring R {i}(P  ia)) = a UP RIP_to_UP i (P  ia)"
  proof
    fix P j x
    assume A0: "P  Pring_set R {i}"
    assume A1: "IP_to_UP i (a Pring R {i}P) = a UP RIP_to_UP i P"
    assume A2: "j  {i}"
    then have A3: "j = i"
      by blast
    have "IP_to_UP i (ring.indexed_pmult R P j)  carrier (UP R)"
      by (simp add: A0 A3 R.Pring_car R.indexed_pset.indexed_pmult R_cring cring.IP_to_UP_closed)
    then have "(a UP R(λn. ring.indexed_pmult R P j (nat_to_mset i n))) x = a R((λn. ring.indexed_pmult R P j (nat_to_mset i n))) x"
      using A0 A1 A3 assms unfolding IP_to_UP_def
      using P_def cfs_smult by blast
    then show " IP_to_UP i (a Pring R {i}(P  j)) x = (a UP RIP_to_UP i (P  j)) x"
      by (metis A0 A2 IP_to_UP_def P_def R.Pring_car R.Pring_smult_cfs R.indexed_pset.indexed_pmult IP_to_UP i (P  j)  carrier (UP R) assms(1) cfs_smult)
  qed
qed

end

text‹Evaluation of indexed polynomials commutes with evaluation of univariate polynomials:›

lemma pvar_closed:
  assumes "cring R"
  assumes "i  I"
  shows "(pvar R i)  carrier (Pring R I)"
  by (meson assms(1) assms(2) cring.axioms(1) ring.Pring_var_closed)

context UP_cring
begin

lemma pvar_mult:
  assumes "i  I"
  assumes "j   I"
  shows "(pvar R i) Pring R I(pvar R j) = mset_to_IP R {#i, j#}"
proof-
  have "{#i#} + {#j#}  = {# i, j#}"
    by auto
  then show ?thesis
    unfolding var_to_IP_def
    by (metis R.Pring_mult R.monom_mult)
qed

lemma pvar_pow:

  assumes "i  I"
  shows "(pvar R i)[^]Pring R I(n::nat) = mset_to_IP R (nat_to_mset i n)"
  apply(induction n)
  apply (metis Group.nat_pow_0 R.one_mset_to_IP R.ring_axioms nat_to_mset_zero ring.Pring_one)
proof-
  fix n
  assume IH: "pvar R i [^]Pring R In = mset_to_IP R (nat_to_mset i n)"
  show "pvar R i [^]Pring R ISuc n = mset_to_IP R (nat_to_mset i (Suc n)) "
  proof-
    have "mset_to_IP R (nat_to_mset i (Suc n)) = mset_to_IP R (nat_to_mset i n) Pring R Ipvar R i"
      using R.monom_mult[of "nat_to_mset i n" "nat_to_mset i 1"]
      by (metis One_nat_def R.Pring_mult Suc_eq_plus1 nat_to_mset_Suc nat_to_mset_add nat_to_mset_zero var_to_IP_def)
    then show ?thesis
      using IH
      by simp
  qed
qed

lemma IP_to_UP_poly_eval:
  assumes "p  Pring_set R {i}"
  assumes "closed_fun R g"
  shows "total_eval R g p = to_function R (IP_to_UP i p) (g i)"
  apply(rule R.indexed_pset.induct[of p "{i}" "carrier R" ])
  apply (simp add: assms(1); fail)
proof-
  show "k. k  carrier R  total_eval R g (R.indexed_const k) = to_function R (IP_to_UP i (R.indexed_const k)) (g i)"
  proof-
    fix k
    assume A: "k  carrier R"
    have P0: "total_eval R g (ring.indexed_const R k) = k"
      unfolding total_eval_def eval_in_ring_def
      using cring.poly_eval_constant[of R k UNIV g]
      by (metis A R.indexed_const_def R_cring)
    have P1: "(IP_to_UP i (ring.indexed_const R k)) = to_polynomial R k"
      by (meson A IP_to_UP_indexed_const)
    have P2: "to_function R (IP_to_UP i (ring.indexed_const R k)) (g i) =
              to_function R (to_polynomial R k) (g i)"
      using P1 by presburger
    have P3: "to_function R (to_polynomial R k) (g i) = k"
      using A assms(2) to_fun_to_poly[of k "g i"] unfolding to_fun_def by blast
    then show "total_eval R g (R.indexed_const k) = to_function R (IP_to_UP i (R.indexed_const k)) (g i)"
      using P0 P2 by presburger
  qed
  show "P Q. P  Pring_set R {i} 
           total_eval R g P = to_function R  (IP_to_UP i P) (g i) 
           Q  Pring_set R {i} 
           total_eval R g Q = to_function R  (IP_to_UP i Q) (g i)  total_eval R g (P  Q) = to_function R (IP_to_UP i (P  Q)) (g i)"
  proof-
    fix p Q
    assume A0: "p  Pring_set R {i}"
    assume A1: " total_eval R g p = to_function R (IP_to_UP i p) (g i)"
    assume A2: " Q  Pring_set R {i}"
    assume A3: "total_eval R g Q = to_function R (IP_to_UP i Q) (g i)"
    have "total_eval R g (R.indexed_padd p Q) = (total_eval R g p) R(total_eval R g Q)"
      using R.total_eval_add[of p "{i}" Q g] A0 A1
      by (metis A2 R.Pring_add R.Pring_car assms(2))
    then
    have 0: "total_eval R g (p  Q) = total_eval R g p  total_eval R g Q "
      by blast
    have 1: "IP_to_UP i (p  Q) = IP_to_UP i p UP RIP_to_UP i Q"
      using A0 A1 A3 assms A2 R.ring_axioms R_cring IP_to_UP_add
      by (metis R.Pring_add R.Pring_car)
    have "g i  carrier R"
      using assms by blast
    hence 2: "to_function R (IP_to_UP i (p  Q)) (g i) = to_function R (IP_to_UP i p) (g i)  to_function R (IP_to_UP i Q) (g i)"
      using A0 A1 A3 assms A2 R.ring_axioms R_cring to_fun_plus[of "IP_to_UP i p" "IP_to_UP i Q" "g i"]
            IP_to_UP_closed is_UP_cring UP_cring.to_fun_def
            to_fun_def 0 1
      unfolding to_fun_def P_def
      by (smt (z3) P_def R.IP_to_UP_closed R.Pring_car to_fun_plus)
    show "total_eval R g (R.indexed_padd p Q) = to_function R (IP_to_UP i (ring.indexed_padd R p Q)) (g i) "
      using A0 A1 A3 assms A2 R.ring_axioms R_cring is_UP_cring to_fun_def 0 1 2
      unfolding to_fun_def   by metis
  qed
  show "P ia.
       P  Pring_set R {i} 
       total_eval R g P = to_function R (IP_to_UP i P) (g i)  ia  {i}  total_eval R g (P  ia) = to_function R (IP_to_UP i (P  ia)) (g i)"
  proof-
    fix P
    fix j
    assume A0: "P  Pring_set R {i}"
    assume A1: "total_eval R g P = to_function R (IP_to_UP i P) (g i)"
    assume A2: "j  {i}"
    then have A3: "j = i"
      by blast
    show "total_eval R g (P  j) = to_function R (IP_to_UP i (P  j)) (g i)"
    proof-
      have LHS: "total_eval R g (P  j) = (total_eval R g P) R(g i)"
        using assms A0 A3
        by (metis R.Pring_car R_cring cring.total_eval_indexed_pmult insertI1)
      have RHS: "IP_to_UP i (P  j)=  IP_to_UP i P  UP RX_poly R"
        by (metis A0 A3 IP_to_UP_indexed_pmult R.Pring_car)
      have "g i  carrier R"
        using assms by blast
      then show ?thesis
          using A0 A1 A3 X_closed to_fun_X[of "g i"] to_fun_mult[of "IP_to_UP i P" "X_poly R" "g i"] LHS RHS
              assms  cring.axioms(1) domain.axioms(1)
              IP_to_UP_indexed_pmult IP_to_UP_closed
              Pring_car unfolding to_fun_def P_def
          by (smt (z3) P.m_comm P_def R.m_comm R_cring cring.IP_to_UP_closed ring.Pring_car to_fun_closed to_fun_def)
      qed
   qed
qed
end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Mapping Univariate Polynomials to Multivariate Polynomials over a Singleton Variable Set›
(**************************************************************************************************)
(**************************************************************************************************)

definition UP_to_IP :: "('a,'b) ring_scheme  'c  'a u_poly  ('a, 'c) mvar_poly" where
"UP_to_IP R i P = (λ m. if (set_mset m)   {i} then P (count m i) else 𝟬R)"

context UP_cring
begin

lemma UP_to_IP_inv:
  assumes "p  Pring_set R {i}"
  shows "UP_to_IP R i (IP_to_UP i p) = p"
proof
  fix x
  show "UP_to_IP R i (IP_to_UP i p) x = p x"
  proof(cases "(set_mset x) = {i}")
    case True
    have "{a. 0 < (λj. if j = i then count x i else 0) a} = {i}"
      by (smt Collect_cong True count_eq_zero_iff neq0_conv singletonI singleton_conv)
  then have "finite {j. (if j = i then count x i else 0)  0}"
      by auto
    have "(λj. if j = i then count x i else 0) = count x"
    proof
      fix j
      show "(if j = i then count x i else 0) = count x j "
        apply(cases "j = i")
        using True
        apply (simp; fail)
        using True
        by (metis count_inI singletonD)
    qed
    then have "(Abs_multiset (λj. if j = i then count x i else 0)) = x"
      using count_inverse
      by simp
    then show ?thesis
      unfolding UP_to_IP_def IP_to_UP_def nat_to_mset_def
      by (metis True set_eq_subset)
  next
    case False
    then show ?thesis
      apply(cases "x = {#}")
       apply (metis count_empty empty_subsetI IP_to_UP_def nat_to_mset_zero set_mset_empty UP_to_IP_def)
     unfolding UP_to_IP_def IP_to_UP_def nat_to_mset_def
     using False  assms
     by (metis R.Pring_set_zero set_mset_eq_empty_iff subset_singletonD)
  qed
qed

lemma UP_to_IP_const:
  assumes "a  carrier R"
  shows "UP_to_IP R i (to_polynomial R a) = ring.indexed_const R a"
proof
  fix x
  show "UP_to_IP R i (to_polynomial R a) x = ring.indexed_const R a x"
    apply(cases "x = {#}")
    unfolding UP_to_IP_def
    apply (metis R.indexed_const_def UP_ring.cfs_monom assms count_eq_zero_iff insert_absorb insert_not_empty is_UP_ring set_mset_empty subset_insert subset_refl to_polynomial_def)
  by (metis R.indexed_const_def UP_ring.cfs_monom assms count_eq_zero_iff is_UP_ring set_mset_eq_empty_iff subset_empty subset_insert to_polynomial_def)
qed

lemma UP_to_IP_add:
  assumes "p  carrier (UP R)"
  assumes "Q  carrier (UP R)"
  shows "UP_to_IP R i (p UP RQ) =
         UP_to_IP R i p Pring R {i}UP_to_IP R i Q"
proof
  fix x
  show "UP_to_IP R i (p UP RQ) x = (UP_to_IP R i p Pring R {i}UP_to_IP R i Q) x"
  proof(cases "set_mset x  {i}")
    case True
    have "(UP_to_IP R i p Pring R {i}UP_to_IP R i Q) x =
          (UP_to_IP R i p) x R(UP_to_IP R i Q) x"
      using True assms
      by (metis R.Pring_add R.indexed_padd_def)
    then show ?thesis     using assms True
      unfolding UP_to_IP_def UP_def
      by (smt partial_object.select_convs(1) restrict_def ring_record_simps(12))
  next
    case False
    have "(UP_to_IP R i p Pring R {i}UP_to_IP R i Q) x =
          (UP_to_IP R i p) x R(UP_to_IP R i Q) x"
            using False assms
            by (metis R.Pring_add R.indexed_padd_def)
    then show ?thesis using False assms
    unfolding UP_to_IP_def UP_def
    using R.l_zero R.zero_closed by presburger
  qed
qed

lemma UP_to_IP_var:
  shows "UP_to_IP R i (X_poly R) = pvar R i"
proof
  have 0: "(count {#i#} i) = 1"
    by simp
  have 1: "set_mset {#i#}  {i}"
    by simp
  have 2: "pvar R i {#i#} = 𝟭"
    by (metis R.mset_to_IP_simp var_to_IP_def)
  fix x
  show "UP_to_IP R i (X_poly R) x = pvar R i x"
    apply(cases "x = {#i#}")
    using X_poly_def[of R] cfs_monom[of 𝟭 1 "count x i"] 0 1 2
    unfolding UP_to_IP_def P_def
    using R.one_closed apply presburger
  proof-
    assume A: "x {#i#}"
    then show "(if set_mset x  {i} then X_poly R (count x i) else 𝟬R) = pvar R i x"
    proof(cases "set_mset x  {i}")
      case True
      have "count x i  1" using True A
        by (metis One_nat_def count_empty count_inI count_single empty_iff multiset_eqI
            set_mset_add_mset_insert set_mset_empty set_mset_eq_empty_iff singletonD singletonI subset_singletonD)
      then have 0: "X_poly R (count x i) = 𝟬R⇙"
        using A UP_cring.X_closed UP_cring.degree_X UP_cring.intro True
        unfolding X_poly_def
        using P_def R.one_closed 𝟭  carrier R  up_ring.monom P 𝟭 1 (count x i) = (if 1 = count x i then 𝟭 else 𝟬) by presburger
      have "pvar R i x = 𝟬R⇙"
        using A var_to_IP_def
        by (metis R.mset_to_IP_simp')
      then show ?thesis
        using A "0" by presburger
    next
      case False
      have "pvar R i x = 𝟬R⇙" using A var_to_IP_def False
        by (metis "1" R.Pring_set_zero R.mset_to_IP_closed)
      then show ?thesis
        unfolding UP_to_IP_def
        using False by presburger
    qed
  qed
qed

lemma UP_to_IP_var_pow:
  shows "UP_to_IP R i ((X_poly R)[^]UP R(n::nat)) = (pvar R i)[^]Pring R {i}n"
proof
  fix x
  show "UP_to_IP R i (X_poly R [^]UP Rn) x = (pvar R i [^]Pring R {i}n) x "
  proof(cases "set_mset x  {i}")
    case True
    show ?thesis
    proof(cases "count x i = n")
      case T: True
      then have 0: "x = nat_to_mset i n"
        using True
        by (metis count_inI emptyE insert_iff multiset_eqI nat_to_msetE
            nat_to_msetE' subsetD)
      have 1: "x = nat_to_mset i (count x i)"
        using "0" T by auto
      then have LHS: "(pvar R i [^]Pring R {i}n) x = 𝟭R⇙"
        using T True  0 1
        by (metis R.mset_to_IP_simp insertI1 pvar_pow)
      have 2: "UP_to_IP R i (X_poly R [^]UP Rn) x = (up_ring.monom (UP R) 𝟭 n) (count x i)"
        unfolding UP_to_IP_def X_poly_def using True
        by (metis ctrm_degree P.nat_pow_closed P.nat_pow_eone P_def R.one_closed UP_cring.monom_coeff
            UP_one_closed UP_r_one X_closed is_UP_cring monom_one monom_rep_X_pow to_poly_inverse
            to_poly_mult_simp(2))
      then show ?thesis
        using True T LHS P_def R.one_closed cfs_monom
        by presburger
    next
      case False
      have "(pvar R i [^]Pring R {i}n) = mset_to_IP R (nat_to_mset i n)"
        by (simp add: pvar_pow)
      hence 0: "(pvar R i [^]Pring R {i}n) x = 𝟬"
        by (metis False R.mset_to_IP_simp' nat_to_msetE)
      have 1: "UP_to_IP R i (X_poly R [^]UP Rn) x = (up_ring.monom (UP R) 𝟭 n) (count x i)"
        unfolding UP_to_IP_def X_poly_def  using False  True
        by (metis ctrm_degree P.nat_pow_closed P.nat_pow_eone P_def R.one_closed UP_one_closed
            UP_r_one X_closed cfs_monom monom_one monom_rep_X_pow to_poly_inverse to_poly_mult_simp(2))
      thus ?thesis using True False
        unfolding UP_to_IP_def X_poly_def 0
        by (metis P_def R.one_closed cfs_monom)
    qed
  next
    case False
    then have 0: "UP_to_IP R i (X_poly R [^]UP Rn) x = 𝟬"
      unfolding UP_to_IP_def
      by meson
    have "(pvar R i [^]Pring R {i}n) = mset_to_IP R (nat_to_mset i n)"
      by (simp add: pvar_pow)
    hence "(pvar R i [^]Pring R {i}n) x = 𝟬"
      by (metis False R.mset_to_IP_simp' count_eq_zero_iff nat_to_msetE' singleton_iff subsetI)
    then show ?thesis using 0
      by presburger
  qed
qed

lemma one_var_indexed_poly_monom_simp:
  assumes "a  carrier R"
  shows "(a Pring R {i}((pvar R i) [^]Pring R {i}n)) x = (if x = (nat_to_mset i n) then a else 𝟬)"
proof-
  have 0: "(a Pring R {i}((pvar R i) [^]Pring R {i}n)) x =
                a  (((pvar R i) [^]Pring R {i}n) x)"
    using  Pring_smult_cfs Pring_var_closed assms cring_def is_cring monoid.nat_pow_closed ring.Pring_is_monoid singletonI
    by (simp add: monoid.nat_pow_closed ring.Pring_is_monoid R.Pring_smult_cfs R.Pring_var_closed R.ring_axioms)
  have 1: "(pvar R i) [^]Pring R {i}n = mset_to_IP R (nat_to_mset i n)"
   using insertI1
   by (simp add: pvar_pow)
  then have 1: "(a Pring R {i}((pvar R i) [^]Pring R {i}n)) x=
              a  (mset_to_IP R (nat_to_mset i n) x)"
    using "0" by presburger
  show ?thesis
    using assms 1 unfolding mset_to_IP_def
    using r_null r_one by simp
qed


lemma UP_to_IP_monom:
  assumes "a  carrier R"
  shows "UP_to_IP R i (up_ring.monom (UP R) a n) = a Pring R {i}((pvar R i)[^]Pring R {i}n)"
proof
  fix x
  show "UP_to_IP R i (up_ring.monom (UP R) a n) x = (a Pring R {i}((pvar R i) [^]Pring R {i}n)) x"
  proof(cases "set_mset x  {i}")
    case True
    then show ?thesis
    proof(cases "count x i = n")
      case T: True
      then have "x = nat_to_mset i n"
        using True
        by (metis count_inI emptyE insert_iff multiset_eqI nat_to_msetE
            nat_to_msetE' subsetD)
      then have LHS: " (a Pring R {i}((pvar R i) [^]Pring R {i}n)) x = a"
        using assms
        by (simp add: one_var_indexed_poly_monom_simp)
      then show ?thesis
        unfolding UP_to_IP_def
        using T True  assms(1)
        by (metis UP_ring.cfs_monom is_UP_ring)
    next
      case False
      then show ?thesis using True
        unfolding UP_to_IP_def
        by (metis INTEG.R.nat_to_msetE P_def assms cfs_monom one_var_indexed_poly_monom_simp)
    qed
  next
    case False
    then show ?thesis
      unfolding UP_to_IP_def
      by (metis (no_types, opaque_lifting) one_var_indexed_poly_monom_simp assms
          count_eq_zero_iff  equalityD2 insert_subset nat_to_msetE' subsetI subset_eq)
  qed
qed

lemma UP_to_IP_monom':
  assumes "a  carrier R"
  shows "UP_to_IP R i (up_ring.monom (UP R) a n) = a Pring R {i}((pvar R i)[^]Pring R {i}n)"
  by (metis R.Pring_smult UP_to_IP_monom assms)

lemma UP_to_IP_closed:
  assumes "p  carrier P"
  shows "(UP_to_IP R i p)  carrier (Pring R {i})"
  apply(rule poly_induct3[of ])
  using assms apply blast
  apply (metis P_def R.Pring_add_closed UP_to_IP_add)
proof-
  fix a fix n::nat
  assume A0: "a  carrier R"
  have "(pvar R i [^]Pring R {i}n)  carrier (Pring R {i})"
    using pvar_closed[of R ] monoid.nat_pow_closed[of "Pring R {i}"]
  proof -
    show ?thesis
      by (meson R.Pring_is_monoid R.Pring_var_closed monoid.nat_pow_closed singleton_iff)
  qed
  then show "a  carrier R 
           UP_to_IP R i (up_ring.monom P a n)  carrier (Pring R {i})"
    using A0 assms(1) UP_to_IP_monom[of a i n] cring.poly_scalar_mult_closed [of R a _ "{i}"]
    by (metis P_def R.Pring_smult_closed)
qed

lemma IP_to_UP_inv:
  assumes "p  carrier P"
  shows "IP_to_UP i (UP_to_IP R i p) = p"
  apply(rule poly_induct3[of ])
  using assms apply linarith
proof-
  show "p q. q  carrier P 
           p  carrier P 
           IP_to_UP i (UP_to_IP R i p) = p 
           IP_to_UP i (UP_to_IP R i q) = q 
           IP_to_UP i (UP_to_IP R i (p Pq)) = p Pq"
  proof-
    fix p q assume A:
           "q  carrier P"
           "p  carrier P"
           "IP_to_UP i (UP_to_IP R i p) = p"
           "IP_to_UP i (UP_to_IP R i q) = q"
    show "IP_to_UP i (UP_to_IP R i (p Pq)) = p Pq"
      using A UP_to_IP_add[of p q i]
                  UP_to_IP_closed
                  IP_to_UP_add
      unfolding P_def
      by metis
  qed
  show "a n. a  carrier R 
           IP_to_UP i (UP_to_IP R i (up_ring.monom P a n)) =
           up_ring.monom P a n"
  proof-
    fix a fix n::nat
    assume A0: "a  carrier R"
    have A1: "pvar R i [^]Pring R {i}n  carrier (Pring R {i})"
      using pvar_closed monoid.nat_pow_closed
      by (metis R.Pring_is_monoid R_cring singletonI)
    have "UP_to_IP R i (up_ring.monom (UP R) a n) = a Pring R {i}(pvar R i [^]Pring R {i}n)"
      by (meson A0 UP_to_IP_monom')
    then have A2: "IP_to_UP i (UP_to_IP R i (up_ring.monom (UP R) a n)) =
            IP_to_UP i (a Pring R {i}(pvar R i [^]Pring R {i}n))"
      by presburger
    have A3: "IP_to_UP i (pvar R i [^]Pring R {i}n) =  (up_ring.monom P 𝟭 n)"
    proof(induction n)
      case 0
      then show ?case
        by (metis Group.nat_pow_0 IP_to_UP_one P_def monom_one)
    next
      case (Suc n)

      then show ?case
      using IP_to_UP_ring_hom[of i]
              ring_hom_mult[of "IP_to_UP i" "Pring R {i}" "UP R" "pvar R i" "pvar R i [^]Pring R {i}n"]
                ring_hom_ring.homh[of "Pring R {i}" "UP R"  "IP_to_UP i"]
      by (metis IP_to_UP_monom P.l_one P.nat_pow_closed P_def R.one_closed UP_cring.ctrm_degree UP_cring.monom_rep_X_pow  UP_one_closed X_closed cfs_monom is_UP_cring monom_one pvar_pow singletonI to_poly_inverse to_poly_mult_simp(1))
    qed
    then show "IP_to_UP i (UP_to_IP R i (up_ring.monom P a n)) =
           up_ring.monom P a n"
      using A2 IP_to_UP_scalar_mult[of a "pvar R i [^]Pring R {i}n" i]
            A0 A1 P_def monic_monom_smult by presburger
  qed
qed

lemma UP_to_IP_mult:
  assumes "p  carrier (UP R)"
  assumes "Q  carrier (UP R)"
  shows "UP_to_IP R i (p UP RQ) =
         UP_to_IP R i p Pring R {i}UP_to_IP R i Q"
proof-
  have 0: "IP_to_UP i (UP_to_IP R i (p UP RQ)) = (p UP RQ)"
    by (meson UP_cring.IP_to_UP_inv UP_ring.UP_mult_closed assms(1) assms(2) is_UP_cring is_UP_ring)
  have 1: "IP_to_UP i (UP_to_IP R i p Pring R {i}UP_to_IP R i Q) =
        IP_to_UP i (UP_to_IP R i p) UP RIP_to_UP i ( UP_to_IP R i Q)"
    using IP_to_UP_ring_hom[of  i]
          ring_hom_mult[of "IP_to_UP i"]
          UP_to_IP_closed assms
    by (smt P_def ring_hom_ring.homh)
  have 2: "IP_to_UP i (UP_to_IP R i (p UP RQ)) =
        IP_to_UP i (UP_to_IP R i p Pring R {i}UP_to_IP R i Q)"
    using 0 1 assms
    by (metis UP_cring.IP_to_UP_inv is_UP_cring)
  then show ?thesis
    by (metis "0" P_def R.Pring_mult_closed R.ring_axioms assms(1) assms(2) ring.Pring_car UP_to_IP_closed UP_to_IP_inv)
qed

lemma UP_to_IP_ring_hom:
shows "ring_hom_ring (UP R) (Pring R {i}) (UP_to_IP R i)"
  apply(rule ring_hom_ringI)
  using P_def UP_ring apply force
  apply (simp add: R.Pring_is_ring; fail)
  apply (metis P_def UP_to_IP_closed)
  apply (meson UP_to_IP_mult)
  apply (meson UP_to_IP_add)
  by (metis IP_to_UP_one R.Pring_car R.Pring_one_closed UP_to_IP_inv)

end

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹The isomorphism $R[I\cup J] \sim R[I][J]$, where $I$ and $J$ are disjoint variable sets›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  Given a ring $R$ and variable sets $I$ and $J$, we'd like to construct the canonical
  (iso)morphism $R[I\cup J] \to R[I][J]$. This can be done with the univeral property of the
  previous section. Let $\phi: R \to R[J]$ be the inclusion of constants, and $f: J \to R[I]$ be
  the map which sends the variable $i$ to the polynomial variable $i$ over the ring $R[I][J]$.
  Then these are the two basic pieces of input required to give us a canonical homomoprhism
  $R[I \cup J] \to R[I][J]$ with the universal property. The first map $\phi$ will be
  "\texttt{dist\_varset\_morpshim}" below, and the second map will be
  "\texttt{dist\_varset\_var\_ass}". The desired induced isomorphism will be
  called "\texttt{var\_factor}".   ›

definition(in ring) dist_varset_morphism
  :: "'d set  'd set 
              ('a, (('a, 'd) mvar_poly, 'd) mvar_poly) ring_hom" where
  "dist_varset_morphism (I:: 'd set) (J:: 'd set) =
        (ring.indexed_const (Pring R J) :: ('d multiset  'a)  'd multiset  ('d multiset  'a)) (ring.indexed_const R ::'a  'd multiset  'a)"

definition(in ring) dist_varset_var_ass
    :: "'d set  'd set  'd  (('a, 'd) mvar_poly, 'd) mvar_poly"
    where
"dist_varset_var_ass (I:: 'd set) (J:: 'd set)  = (λi. if i  J then ring.indexed_const (Pring R J) (pvar R i) else
                                                pvar (Pring R J) i )"

context cring
begin

lemma dist_varset_morphism_is_morphism:
  assumes "(I:: 'd set)  J0  J1"
  assumes "J1  I"
  assumes "φ = dist_varset_morphism I J0"
  shows "ring_hom_ring R (Pring (Pring R J0) J1) φ"
proof-
  have 0:"ring_hom_ring R (Pring R J0) indexed_const"
    by (simp add: indexed_const_ring_hom)
  have 1:"ring_hom_ring (Pring R J0) (Pring (Pring R J0) J1) (ring.indexed_const (Pring R J0))"
    by (simp add: cring.indexed_const_ring_hom is_cring local.ring_axioms ring.Pring_is_cring)
  show ?thesis using 0 1 assms ring_hom_trans[of "indexed_const" R "Pring R J0" "ring.indexed_const (Pring R J0)"
              "(Pring (Pring R J0) J1) "]
    unfolding dist_varset_morphism_def
  by (meson ring_hom_ring.axioms(1) ring_hom_ring.axioms(2) ring_hom_ring.homh ring_hom_ringI2)
qed

definition var_factor ::
  "'d set  'd set  'd set 
         (('a, 'd) mvar_poly, (('a, 'd) mvar_poly, 'd) mvar_poly) ring_hom "where
"var_factor (I:: 'd set) (J0:: 'd set) (J1:: 'd set) = indexed_poly_induced_morphism  I (Pring (Pring R J0) J1)
                                                        (dist_varset_morphism I J0) (dist_varset_var_ass I J0)"

lemma indexed_const_closed:
  assumes "x  carrier R"
  shows "indexed_const x  carrier (Pring R I)"
  using Pring_car assms indexed_pset.indexed_const by blast

lemma var_factor_morphism:
  assumes "(I:: 'd set)  J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "g = dist_varset_var_ass I J0"
  assumes "φ = dist_varset_morphism I J0"
  assumes "ψ = (var_factor I J0 J1)"
  shows "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) ψ "
        "i. i  J0  I  ψ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)"
        "i. i  J1  ψ (pvar R i) = pvar (Pring R J0) i"
        "a. a  carrier (Pring R (J0  I))  ψ a = ring.indexed_const (Pring R J0) a"
proof-
  have 0: "g  I  carrier (Pring (Pring R J0) J1)"
  proof
    fix x
    assume A0: "x  I"
    then have A1: "x  J0  x  J1"
      by (meson UnE assms(1) subsetD)
    have A2: "x  J0  x  J1"
      using A1 by blast
    show "g x  carrier (Pring (Pring R J0) J1)"
      apply(cases "x  J0")
      using assms A0 A1 A2 pvar_closed[of R x J0]
        pvar_closed[of "Pring R J0" x J1]
        cring.indexed_const_closed[of "Pring R J0" ]
      unfolding dist_varset_var_ass_def
       apply (smt Pring_is_cring is_cring)
      using assms A0 A1 A2 pvar_closed[of R x J0]
        pvar_closed[of "Pring R J0" x J1]
        cring.indexed_const_closed[of "Pring R J0" ]
      unfolding dist_varset_var_ass_def
      by (smt Pring_is_cring is_cring)
  qed
  have 1: " cring (Pring R J0)"
    by (simp add: Pring_is_cring is_cring)
  have 2: " cring (Pring (Pring R J0) J1)"
    by (simp add: "1" Pring_is_ring ring.Pring_is_cring)
  show C0: "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) ψ "
    using 0 assms Pring_universal_prop[of "Pring (Pring R J0) J1" g I φ ψ]
        dist_varset_morphism_is_morphism[of I J0 J1]
    unfolding var_factor_def
    by (meson Pring_is_cring Pring_is_ring is_cring ring.Pring_is_cring)
  show C1: "i. i  J0  I  ψ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)"
    using  0 1 2 assms Pring_universal_prop(2)[of "Pring (Pring R J0) J1" g I φ ψ]
        dist_varset_morphism_is_morphism[of I J0 J1 φ] dist_varset_var_ass_def
        dist_varset_morphism_def unfolding var_factor_def  var_to_IP_def
    by (smt IntE dist_varset_var_ass_def inf_commute inf_le2 ring.Pring_is_cring subsetD var_to_IP_def)
  have 3: "i. i  J1  g i = mset_to_IP (Pring R J0) {#i#} "
    using assms unfolding dist_varset_var_ass_def var_to_IP_def
    by (meson disjoint_iff_not_equal)
  show C2: "i. i  J1  ψ (pvar R i) = pvar (Pring R J0) i"
      using  0 1 2 3 assms Pring_universal_prop(2)[of "Pring (Pring R J0) J1" g I φ ψ]
        dist_varset_morphism_is_morphism[of I J0 J1 φ]
         unfolding   var_factor_def  var_to_IP_def
         by (metis subsetD)
  have 4: "k. k  carrier R  ψ (indexed_const k) = ring.indexed_const (Pring R J0) (indexed_const k)"
      using  0 1 2 3 assms Pring_universal_prop(3)[of "Pring (Pring R J0) J1" g I φ ψ]
        dist_varset_morphism_is_morphism[of I J0 J1 φ] comp_apply
      unfolding   var_factor_def  var_to_IP_def dist_varset_morphism_def
      by metis
  show C3:  "a. a  carrier (Pring R (J0  I))  ψ a = ring.indexed_const (Pring R J0) a"
  proof- fix a assume A: "a  carrier (Pring R (J0  I))"
    show " ψ a = ring.indexed_const (Pring R J0) a"
      apply(rule indexed_pset.induct[of a "J0  I" "carrier R"])
      using A Pring_car apply blast
      using 4
      apply blast
    proof-
      show "P Q. P  Pring_set R (J0  I) 
           ψ P = ring.indexed_const (Pring R J0) P 
           Q  Pring_set R (J0  I)  ψ Q = ring.indexed_const (Pring R J0) Q  ψ (P  Q) = ring.indexed_const (Pring R J0) (P  Q)"
      proof- fix P Q
        assume A0: "P  Pring_set R (J0  I)"
        assume A1: "ψ P = ring.indexed_const (Pring R J0) P"
        assume A2: "Q  Pring_set R (J0  I)"
        assume A3: "ψ Q = ring.indexed_const (Pring R J0) Q"
        have A0': "P  Pring_set R I"
          using A0
          by (meson Int_lower2 Pring_carrier_subset subsetD)
        have A1': "Q  Pring_set R I"
          by (meson A2 Int_lower2 Pring_carrier_subset subsetD)
        have B: "ψ (P  Q) = ψ P Pring (Pring R J0) J1ψ Q"
          using A0' A1' A2 A3 C0 assms ring_hom_add
          by (metis (no_types, lifting) Pring_add local.ring_axioms ring.Pring_car ring_hom_ring.homh)
        have " ring.indexed_const (Pring R J0) (P  Q) =  ring.indexed_const (Pring R J0) P
                                 Pring (Pring R J0) J1ring.indexed_const (Pring R J0) Q"
          by (simp add: Pring_add Pring_is_ring ring.Pring_add ring.indexed_padd_const)
        then show " ψ (P  Q) = ring.indexed_const (Pring R J0) (P  Q)"
          using B
          by (simp add: A1 A3)
      qed
      show "P i. P  Pring_set R (J0  I)  ψ P = ring.indexed_const (Pring R J0) P 
                    i  J0  I  ψ (P  i) = ring.indexed_const (Pring R J0) (P  i)"
      proof- fix P i assume A0: "P  Pring_set R (J0  I)"
        assume A1: "ψ P = ring.indexed_const (Pring R J0) P"
        assume A2: "i  J0  I"
        have A0': "P  carrier (Pring R I)"
          using A0  Pring_carrier_subset
          by (metis (no_types, opaque_lifting) Pring_car in_mono inf_commute le_inf_iff subset_refl)
        have A0'': "P  carrier (Pring R J0)"
          using A0  Pring_carrier_subset
          by (metis (no_types, opaque_lifting) Pring_car in_mono inf_commute le_inf_iff subset_refl)
        have " ψ (P  i) = ψ P Pring (Pring R J0) J1ring.indexed_const (Pring R J0) (pvar R i)"
        proof-
          have "(P  i) = P Pring R Ipvar R i"
            using A0' A2 unfolding var_to_IP_def
            by (metis A0 Pring_mult poly_index_mult)
          then show ?thesis
            using A0' C0 A2 C1[of i] ring_hom_ring.homh
                  ring_hom_mult[of ψ "Pring R I" "Pring (Pring R J0) J1" P "pvar R i"]
          by (metis IntE Pring_var_closed)
        qed
        then have" ψ (P  i) = ring.indexed_const (Pring R J0) P Pring (Pring R J0) J1ring.indexed_const (Pring R J0) (pvar R i)"
          by (simp add: A1)
        then have " ψ (P  i) = ring.indexed_const (Pring R J0) (P Pring R J0(pvar R i))"
          using A0'' A2  cring.indexed_const_ring_hom[of "Pring R J0" J1] ring_hom_ring.homh
                ring_hom_mult[of "ring.indexed_const (Pring R J0)" "Pring R J0" _  P "pvar R i"]
          by (smt "1" IntE Pring_var_closed)
        then show "ψ (P  i) = ring.indexed_const (Pring R J0) (P  i)"
          using poly_index_mult[of P J0 i] unfolding var_to_IP_def
          by (metis A0'' A2 IntE Pring_car Pring_mult)
      qed
    qed
  qed
qed

lemma var_factor_morphism':
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "ψ = (var_factor I J0 J1)"
  shows "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) ψ "
        "i. i  J1  ψ (pvar R i) = pvar (Pring R J0) i"
        "a. a  carrier (Pring R (J0  I))  ψ a = ring.indexed_const (Pring R J0) a"
  using assms var_factor_morphism
  apply blast
   using assms var_factor_morphism(3)
  apply (metis subset_refl)
     using assms var_factor_morphism(4)
  by (metis Un_subset_iff Un_upper1)

text‹Constructing the inverse morphism for \texttt{var\_factor\_morphism} ›


lemma pvar_ass_closed:
  assumes "J1  I"
  shows "pvar R  J1  carrier (Pring R I)"
  by (meson Pi_I Pring_var_closed assms subsetD)

text‹The following function gives us the inverse morphism $R[I][J] \to R[I \cup J]$:›
definition var_factor_inv :: "'d set  'd set  'd set 
  ((('a, 'd) mvar_poly, 'd) mvar_poly, ('a, 'd) mvar_poly) ring_hom" where
"var_factor_inv (I:: 'd set) (J0:: 'd set) (J1:: 'd set) = indexed_poly_induced_morphism J1 (Pring R I)
                                                        (id:: ('d multiset  'a)  'd multiset  'a)
                                                       (pvar R)"

lemma var_factor_inv_morphism:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "ψ = (var_factor_inv I J0 J1)"
  shows "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I)  ψ "
        "i. i  J1  ψ (pvar (Pring R J0) i) = pvar R i"
        "a. a  carrier (Pring R J0)  ψ (ring.indexed_const (Pring R J0) a) = a"
proof-
  have 0:  "ring_hom_ring (Pring R J0) (Pring R I) id"
    apply(rule ring_hom_ringI)
      apply (simp add: Pring_is_ring; fail)
        apply (simp add: Pring_is_ring;fail )
       apply (metis Pring_car Pring_carrier_subset Un_upper1 assms(1) id_apply subsetD)
    apply (metis Pring_mult_eq id_apply)
     apply (metis Pring_add_eq id_apply)
  by (simp add: Pring_one_eq)
  then show "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I)  ψ "
    using cring.Pring_universal_prop(1)[of "Pring R J0" "Pring R I" "pvar R" J1 id ψ]
        pvar_ass_closed[of J0 I]
    by (metis Pring_is_cring assms(2) assms(4) is_cring pvar_ass_closed var_factor_inv_def)
  show "i. i  J1  ψ (pvar (Pring R J0) i) = pvar R i"
    using 0 assms pvar_ass_closed[of J0 I]
        cring.Pring_universal_prop(2)[of "Pring R J0" "Pring R I" "pvar R" J1 id ψ]
    by (metis Pring_is_cring is_cring pvar_ass_closed var_factor_inv_def var_to_IP_def)
  show "a. a  carrier (Pring R J0)  ψ (ring.indexed_const (Pring R J0) a) = a"
    using 0 assms pvar_ass_closed[of J0 I]
        cring.Pring_universal_prop(3)[of "Pring R J0" "Pring R I" "pvar R" J1 id ψ]
    by (smt Pi_I Pring_is_cring Pring_var_closed id_def is_cring subsetD var_factor_inv_def)
qed

lemma var_factor_inv_inverse:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "ψ1 = (var_factor_inv I J0 J1)"
  assumes "ψ0 = (var_factor I J0 J1)"
  assumes "P  carrier (Pring R I)"
  shows "ψ1 (ψ0 P) = P"
  apply(rule indexed_pset.induct[of P I "carrier R"])
  using Pring_car assms(6) apply blast
    using var_factor_inv_morphism(3)[of I J0 J1 "ψ1"] var_factor_morphism'(3)[of I J0 J1 ψ0] assms
    apply (metis indexed_const_closed inf_sup_absorb)
proof-
  have 0: "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I) ψ1"
    by (simp add: assms(1) assms(3) assms(4) var_factor_inv_morphism(1))
  have 1: "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) ψ0"
    by (simp add: assms(1) assms(3) assms(5) var_factor_morphism'(1))
  have 2: "ψ1  ψ0  ring_hom (Pring R I) (Pring R I)"
    using 0 1 ring_hom_trans[of ψ0 "Pring R I" "Pring (Pring R J0) J1" ψ1 "Pring R I"]
          ring_hom_ring.homh[of "Pring R I" "Pring (Pring R J0) J1" ψ0]
          ring_hom_ring.homh[of "Pring (Pring R J0) J1" "Pring R I" ψ1]
    by blast
  show "P Q. P  Pring_set R I 
           ψ1 (ψ0 P) = P  Q  Pring_set R I  ψ1 (ψ0 Q) = Q  ψ1 (ψ0 (P  Q)) = P  Q"
  proof- fix P Q assume A0: "P  Pring_set R I" "ψ1 (ψ0 P) = P"
    assume A1: "Q  Pring_set R I" "ψ1 (ψ0 Q) = Q"
    show "ψ1 (ψ0 (P  Q)) = P  Q"
      using  A0 A1 2 ring_hom_add[of "ψ1  ψ0" "Pring R I" "Pring R I" P Q] comp_apply[of ψ1 ψ0]
      by (simp add: "2" P  Pring_set R I Q  Pring_set R I Pring_add Pring_car)
  qed
  show "P i. P  Pring_set R I  ψ1 (ψ0 P) = P  i  I  ψ1 (ψ0 (P  i)) = P  i"
  proof-
    fix P i assume A: "P  Pring_set R I" "ψ1 (ψ0 P) = P" "i  I"
    show "ψ1 (ψ0 (P  i)) = P  i"
    proof-
      have A0: "P  i = P Pring R Ipvar R i"
        by (metis A(1) A(3) Pring_mult local.ring_axioms ring.poly_index_mult var_to_IP_def)
      have A1: "ψ1 (ψ0 (pvar R i)) = pvar R i"
       by (metis A(3) Int_iff Pring_var_closed UnE Un_subset_iff Un_upper1 assms(1) assms(2)
           assms(3) assms(4) assms(5) cring.var_factor_morphism(2) is_cring
           var_factor_inv_morphism(2) var_factor_inv_morphism(3) var_factor_morphism'(2))
     then show ?thesis
       using 2 A0 A ring_hom_mult[of "ψ1  ψ0" "Pring R I" _  P "pvar R i" ]
            Pring_car Pring_var_closed comp_apply[of ψ1 ψ0]
       by smt
   qed
 qed
qed

lemma var_factor_total_eval:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "ψ = (var_factor I J0 J1)"
  assumes "closed_fun R g"
  assumes "P  carrier (Pring R I)"
  shows "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P))"
  apply(rule indexed_pset.induct[of P I "carrier R"])
  using Pring_car assms apply blast
    apply (metis Pring_is_cring assms(1) assms(2) assms(3) assms(4) cring.total_eval_const
      indexed_const_closed is_cring var_factor_morphism'(3))
proof-
  show "P Q. P  Pring_set R I 
           total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P)) 
           Q  Pring_set R I 
           total_eval R g Q = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ Q)) 
           total_eval R g (P  Q) = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  Q)))"
  proof- fix P Q
    assume A: " P  Pring_set R I"
           "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P))"
    assume B: " Q  Pring_set R I"
           "total_eval R g Q = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ Q))"
    have 0: "ψ (P  Q) = ψ P Pring (Pring R J0) J1ψ Q"
      using A B assms var_factor_morphism'(1)[of I J0 J1 ψ]
            Pring_add[of I P Q] Pring_car[of I]
            ring_hom_ring.homh[of "Pring R I" "Pring (Pring R J0) J1" ψ]
            ring_hom_add[of ψ "Pring R I" "Pring (Pring R J0) J1" P Q]
      by metis
    have 1: "closed_fun (Pring R J0) (indexed_const  g)"
      using assms comp_apply
      by (smt Pi_I closed_fun_simp indexed_const_closed)
    have 2: "ψ P  carrier (Pring (Pring R J0) J1)"
      using assms A var_factor_morphism'(1)[of I J0 J1 ψ]
            ring_hom_ring.homh ring_hom_closed Pring_car
      by metis
    have 3: "ψ Q  carrier (Pring (Pring R J0) J1)"
      using assms B var_factor_morphism'(1)[of I J0 J1 ψ]
          ring_hom_ring.homh ring_hom_closed Pring_car
      by metis
    have 4: "(total_eval (Pring R J0) (indexed_const  g) (ψ (P  Q))) =
            (total_eval (Pring R J0) (indexed_const  g) (ψ P)) Pring R J0(total_eval (Pring R J0) (indexed_const  g) (ψ Q))"
      using 0 1 2 3 A B assms cring.total_eval_add[of "Pring R J0" "ψ P" J1 "ψ Q" "indexed_const  g"]
      by (metis Pring_car Pring_is_cring is_cring)
    have 5: "(total_eval (Pring R J0) (indexed_const  g) (ψ P))  carrier (Pring R J0)"
      using 3 assms cring.total_eval_closed "1" "2" Pring_is_cring is_cring
      by blast
    have 6: "(total_eval (Pring R J0) (indexed_const  g) (ψ Q))  carrier (Pring R J0)"
      using 4 assms  "1" "3" Pring_is_cring cring.total_eval_closed is_cring
      by blast
    show "total_eval R g (P  Q) = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  Q)))"
      using 5 6 4 assms
            total_eval_add[of  "(total_eval (Pring R J0) (indexed_const  g) (ψ p))" J0
                              "(total_eval (Pring R J0) (indexed_const  g) (ψ Q))" ]
      by (smt A(1) A(2) B(1) B(2) Pring_add Pring_car in_mono indexed_pset_mono order_refl
          subsetD subset_iff total_eval_add)
  qed
  show "P i. P  Pring_set R I 
           total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P)) 
           i  I  total_eval R g (P  i) = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  i)))"
  proof- fix P i
    assume A: "P  Pring_set R I"
           "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P))"
           "i  I"
    have 0: "(P  i) = P Pring R I(pvar R i)"
      using A poly_index_mult
      by (metis Pring_mult var_to_IP_def)
    have 1: "ψ (P  i) = ψ P Pring (Pring R J0) J1ψ (pvar R i)"
      using 0 A assms var_factor_morphism'(1)[of I J0 J1 ψ]
            pvar_closed[of R i] ring_hom_mult ring_hom_ring.homh
      by (smt Pring_car Pring_var_closed)
    have 2: "ψ P  carrier (Pring (Pring R J0) J1)"
      using assms A var_factor_morphism'(1)[of I J0 J1 ψ]
            ring_hom_ring.homh ring_hom_closed Pring_car
      by metis
    have 3: "ψ (pvar R i)  carrier (Pring (Pring R J0) J1)"
      using assms A var_factor_morphism'(1)[of I J0 J1 ψ]
            ring_hom_ring.homh ring_hom_closed Pring_car pvar_closed[of R i I]
      by (metis is_cring)
    have 4: "closed_fun (Pring R J0) (indexed_const  g)"
      apply(rule cring.closed_funI)
      using Pring_is_cring is_cring apply blast
            using assms indexed_const_closed closed_fun_simp[of g] comp_apply[of indexed_const g]
          proof -
            fix x :: 'c
            show "(indexed_const  g) x  carrier (Pring R J0)"
              by (metis (no_types) n. closed_fun R g  g n  carrier R x. (indexed_const  g) x = indexed_const (g x) closed_fun R g indexed_const_closed)
          qed
    have 5: "(total_eval (Pring R J0) (indexed_const  g) (ψ (P  i))) =
              (total_eval (Pring R J0) (indexed_const  g) (ψ P)) Pring R J0(total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i)))"
      using 2 3 4 cring.total_eval_ring_hom[of "Pring R J0" "indexed_const  g" J1]
      by (metis "1" Pring_is_cring cring.total_eval_mult is_cring)
    have 6: "(total_eval (Pring R J0) (indexed_const  g) (ψ P))  carrier (Pring R J0)"
      using total_eval_closed "2" "4" Pring_is_cring cring.total_eval_closed is_cring
      by blast
    have 7: " (total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i)))  carrier (Pring R J0)"
      using "3" "4" Pring_is_cring cring.total_eval_closed is_cring
      by blast
    have 8: " total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  i))) =
              total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ P)) Rtotal_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i)))"
      using 6 7
      by (metis "5" assms(5) cring.total_eval_mult is_cring)
    have 9: " total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  i))) =
              total_eval R g P Rtotal_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i)))"
      using "8" A(2)
      by presburger
    have 10: "total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i))) =
              g i"
    proof(cases "i  J0")
      case True
      then have "ψ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)"
        using assms
        by (metis Pring_var_closed inf_sup_absorb var_factor_morphism'(3))
      then have "(total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i))) =  (pvar R i)"
        by (metis Pring_is_cring Pring_var_closed True cring.total_eval_const is_cring)
      then show ?thesis
        using total_eval_var[of g i] assms var_to_IP_def
        by metis
    next
      case False
      then have "ψ (pvar R i) = (pvar  (Pring R J0)  i)"
        by (metis A(3) UnE assms(1) assms(2) assms(3) assms(4) cring.var_factor_morphism'(2) is_cring)
      then have "total_eval (Pring R J0) (indexed_const  g) (ψ (pvar R i)) =
                  indexed_const (g i)"
        using cring.total_eval_var[of "Pring R J0" "indexed_const  g" i] comp_apply
        by (metis "4" Pring_is_cring is_cring var_to_IP_def)
      then show ?thesis
        using total_eval_const[of "g i" g] assms closed_fun_simp[of g i]
        by metis
    qed
    show "total_eval R g (P  i) = total_eval R g (total_eval (Pring R J0) (indexed_const  g) (ψ (P  i)))"
      using 9 10
      by (metis A(1) A(3) Pring_car assms(5) total_eval_indexed_pmult)
  qed
qed

lemma var_factor_closed:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "P  carrier (Pring R I)"
  shows "var_factor I J0 J1 P  carrier (Pring (Pring R J0 ) J1)"
  using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh
  by (metis ring_hom_closed)

lemma var_factor_add:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  shows "var_factor I J0 J1 (P Pring R IQ) = var_factor I J0 J1 P Pring (Pring R J0) J1var_factor I J0 J1 Q"
  using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh
  by (metis (no_types, lifting) ring_hom_add)

lemma var_factor_mult:
  assumes "I = J0  J1"
  assumes "J1  I"
  assumes "J1  J0 = {}"
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  shows "var_factor I J0 J1 (P Pring R IQ) = var_factor I J0 J1 P Pring (Pring R J0) J1var_factor I J0 J1 Q"
  using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh
  by (metis (no_types, lifting) ring_hom_mult)

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Viewing a Mulitvariable Polynomial as a Univariate Polynomial over a Multivariate Polynomial Base Ring›
(**************************************************************************************************)
(**************************************************************************************************)

definition multivar_poly_to_univ_poly ::
 "'c set  'c  ('a,'c) mvar_poly 
       (('a,'c) mvar_poly) u_poly  " where
"multivar_poly_to_univ_poly I i P = ((IP_to_UP i)  (var_factor I (I - {i}) {i})) P"

definition univ_poly_to_multivar_poly ::
  "'c set  'c  (('a,'c) mvar_poly) u_poly 
        ('a,'c) mvar_poly"  where
"univ_poly_to_multivar_poly I i P =((var_factor_inv I (I - {i}) {i})  (UP_to_IP (Pring R (I - {i})) i)) P"

lemma multivar_poly_to_univ_poly_is_hom:
  assumes "i  I"
  shows "multivar_poly_to_univ_poly I i  ring_hom (Pring R I) (UP (Pring R (I - {i})))"
  unfolding multivar_poly_to_univ_poly_def comp_apply
  apply(rule ring_hom_compose[of _ "Pring (Pring R (I - {i})) {i}" _ "var_factor I (I - {i}) {i}" "IP_to_UP i"])
       apply (simp add: Pring_is_ring; fail)
      apply (simp add: local.ring_axioms ring.Pring_is_ring; fail)
     apply(rule UP_ring.UP_ring) unfolding UP_ring_def
     apply (simp add: Pring_is_ring; fail)
    using assms var_factor_morphism'[of I "I - {i}" "{i}"] unfolding ring_hom_ring_def ring_hom_ring_axioms_def
    apply blast
    using UP_cring.IP_to_UP_ring_hom[of "Pring R (I - {i})" i]
    unfolding ring_hom_ring_def ring_hom_ring_axioms_def UP_cring_def
    using Pring_is_cring is_cring apply blast
    by blast

lemma multivar_poly_to_univ_poly_inverse:
  assumes "i  I"
  assumes "ψ0 = multivar_poly_to_univ_poly I i"
  assumes "ψ1 = univ_poly_to_multivar_poly I i"
  assumes "P  carrier (Pring R I)"
  shows "ψ1 (ψ0 P) = P"
proof-
  have closed: "var_factor I (I - {i}) {i} P  carrier (Pring (Pring R (I - {i})) {i})"
    using assms var_factor_closed[of I "I - {i}" "{i}" P]
    by blast
  have cancel_1: "UP_to_IP (Pring R (I - {i})) i
                  (IP_to_UP i (var_factor I (I - {i}) {i} P)) =
                        var_factor I (I - {i}) {i} P"
    using closed assms  ring.Pring_car
          UP_cring.UP_to_IP_inv[of "Pring R (I - {i})" "var_factor I (I - {i}) {i} P" i]
          Pring_is_ring unfolding UP_cring_def
    using is_cring local.ring_axioms ring.Pring_is_cring
    by blast
  have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) =
        univ_poly_to_multivar_poly I i ((IP_to_UP i) (var_factor I (I - {i}) {i} P))"
    by (metis comp_apply multivar_poly_to_univ_poly_def)
  then have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) =
        ((var_factor_inv I (I - {i}) {i})  ((UP_to_IP (Pring R (I - {i})) i)
                 ((IP_to_UP i) (var_factor I (I - {i}) {i} P))))"
    using comp_apply univ_poly_to_multivar_poly_def
    by metis
  then have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) =
        ((var_factor_inv I (I - {i}) {i})  (var_factor I (I - {i}) {i} P))"
    using cancel_1
    by presburger
  then show ?thesis using assms var_factor_inv_inverse[of I "I - {i}" "{i}" _ _ P]
    by (metis Diff_cancel Diff_disjoint Diff_eq_empty_iff Diff_partition Un_Diff_cancel
        Un_Diff_cancel2 Un_insert_right empty_Diff empty_subsetI insert_Diff_if insert_absorb )
qed

lemma multivar_poly_to_univ_poly_total_eval:
  assumes "i  I"
  assumes "ψ = multivar_poly_to_univ_poly I i"
  assumes "P  carrier (Pring R I)"
  assumes "closed_fun R g"
  shows "total_eval R g P = total_eval R g (to_function (Pring R (I - {i})) (ψ P) (indexed_const (g i)))"
proof-
  have 0: "var_factor I (I - {i}) {i} P  Pring_set (Pring R (I - {i})) {i}"
  proof-
    have 00: " ring (Pring R (I - {i}))"
      using Pring_is_ring
      by blast
    then show ?thesis
      using assms(1) assms(3) var_factor_closed[of I "I - {i}" "{i}" P] ring.Pring_car[of "Pring R (I - {i})" "{i}"]
      by blast
  qed
  have 1: "closed_fun (Pring R (I - {i})) (indexed_const  g)"
    using assms comp_apply  indexed_const_closed cring.closed_funI[of "Pring R (I - {i})"]
    by (metis Pring_is_cring closed_fun_simp is_cring)
  have 2: "cring (Pring R (I - {i}))"
    using Pring_is_cring is_cring by blast
  have "(to_function (Pring R (I - {i})) (ψ P) (indexed_const (g i)))
        =  total_eval (Pring R (I - {i})) (indexed_const  g) ((var_factor I (I - {i}) {i}) P) "
    using assms 0 1 2 comp_apply UP_cring.IP_to_UP_poly_eval[of "Pring R (I - {i})"
                                             "(var_factor I (I - {i}) {i}) P" i "indexed_const  g"]
    unfolding UP_cring_def multivar_poly_to_univ_poly_def
    by metis
  then have 3: "total_eval R g ((to_function (Pring R (I - {i})) (ψ P) (indexed_const (g i))) )
        = total_eval R g ( total_eval (Pring R (I - {i})) (indexed_const  g) ((var_factor I (I - {i}) {i}) P)) "
    by presburger
  then have "total_eval R g P = total_eval R g (total_eval (Pring R (I - {i})) (indexed_const  g) ((var_factor I (I - {i}) {i}) P)) "
    using assms var_factor_total_eval[of I "I - {i}" "{i}" "var_factor I (I - {i}) {i}" g P]
    by blast
  then show ?thesis
    using 3
    by presburger
qed

text‹
  Induction for polynomial rings. Basically just \texttt{indexed\_pset.induct} with some
  boilerplate translations removed
›

lemma(in ring) Pring_car_induct'':
  assumes "Q  carrier (Pring R I)"
  assumes "c. c  carrier R  P (indexed_const c)"
  assumes "p q. p  carrier (Pring R I)  q  carrier (Pring R I)  P p  P q  P (p Pring R Iq)"
  assumes "p i. p  carrier (Pring R I)  i  I  P p  P (p Pring R Ipvar R i)"
  shows "P Q"
  apply(rule indexed_pset.induct[of Q I "carrier R"])
  using Pring_car assms(1) apply blast
  using assms(2) apply blast
  apply (metis (full_types) Pring_add Pring_car assms(3))
proof-
  fix Pa i assume A: "Pa  Pring_set R I" "P Pa" "i  I"
  then have 0: "Pa  carrier (Pring R I)"
    using  assms A Pring_car
    by blast
  have "Pa  i = Pa Pring R Ipvar R i"
    using 0 poly_index_mult assms A
    by (metis Pring_mult var_to_IP_def)
  then show "P (Pa  i)"
    by (simp add: "0" A(2) A(3) assms)
qed

subsubsection‹Application: A Polynomial Ring over a Domain is a Domain›

text ‹
  In this section, we use the fact the UP R› is a domain when R› is a domain to show the analogous
  result for indexed polynomial rings. We first prove it inductively for rings with a finite
  variable set, and then generalize to infinite variable sets using the fact that any two
  multivariable polynomials over an indexed polynomial ring must also lie in a finitely indexed
  polynomial subring.
›

lemma indexed_const_mult:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "indexed_const a Pring R Iindexed_const b  = indexed_const (a   b)"
  by (metis Pring_mult assms(1) assms(2) indexed_const_P_mult_eq)

lemma(in domain) Pring_fin_vars_is_domain:
  assumes "finite (I ::'c set)"
  shows "domain (Pring R I)"
proof(rule finite_induct, rule assms)
  show "domain (Pring R ({}::'c set))"
  proof(rule domainI)
    show " cring (Pring R {})"
      by (simp add: Pring_is_cring is_cring)
    show "𝟭Pring R {} 𝟬Pring R {}⇙"
    proof-
      have "𝟭Pring R {}{#}  𝟬Pring R {}{#}"
        by (metis Pring_one Pring_zero local.ring_axioms ring.indexed_const_def zero_not_one)
      thus ?thesis
        by metis
    qed
    show "a b. a Pring R ({}::'c set)b = 𝟬Pring R {}
           a  carrier (Pring R {})  b  carrier (Pring R {})  a = 𝟬Pring R {} b = 𝟬Pring R {}⇙"
    proof-
      fix a b assume A: "a Pring R ({}::'c set)b = 𝟬Pring R {}⇙"
           "a  carrier (Pring R {})" " b  carrier (Pring R {})"
      have 0: "monomials_of R a  {{#}}"
        using A Pring_set_zero unfolding monomials_of_def Pring_car
        by blast
      have 1: "monomials_of R b  {{#}}"
        using A Pring_set_zero unfolding monomials_of_def Pring_car
        by blast
      obtain A where A_def: "A = a {#}"
        by blast
      obtain B where B_def: "B = b {#}"
        by blast
      have 2: "a = indexed_const A"
        unfolding A_def
        apply(rule ext)
        by (metis 0 complement_of_monomials_of empty_iff in_mono indexed_const_def insert_iff)
      have 3: "b = indexed_const B"
        unfolding B_def
        apply(rule ext)
        by (metis 1 complement_of_monomials_of empty_iff in_mono indexed_const_def insert_iff)
      have 4: "a Pring R {}b = indexed_const (A  B)"
        using A unfolding 2 3
        by (metis "2" "3" A_def B_def Pring_carrier_coeff' indexed_const_mult)
      have 5: "A  B = 𝟬"
        using A unfolding 4 by (metis Pring_zero indexed_const_def)
      have 6: "A = 𝟬  B = 𝟬"
        using 5 A_def B_def A
        by (simp add: domain.integral_iff domain_axioms)
      show "a = 𝟬Pring R {} b = 𝟬Pring R {}⇙"
        unfolding 2 3 using 6
        by (metis Pring_zero)
    qed
  qed
  show "x F. finite (F:: 'c set)  x  F  domain (Pring R F)  domain (Pring R (insert x F))"
  proof- fix S:: "'c set"  fix s assume A: "finite S" "s  S" "domain (Pring R S)"
    show "domain (Pring R (insert s S))"
    proof-
      have ring_hom: "multivar_poly_to_univ_poly (insert s S) s  ring_hom (Pring R (insert s S)) (UP (Pring R S))"
        using multivar_poly_to_univ_poly_is_hom
        by (metis A(2) Diff_insert_absorb insertI1)
      have domain: "domain (UP (Pring R S))"
        apply(rule UP_domain.UP_domain)
        unfolding UP_domain_def by(rule A)
      show "domain (Pring R (insert s S))"
        apply(rule ring_hom_ring.inj_on_domain[of _ "UP (Pring R S)" "multivar_poly_to_univ_poly (insert s S) s"])
          apply(rule ring_hom_ring.intro)
            apply (simp add: Pring_is_ring; fail)
           apply(rule UP_ring.UP_ring)
        unfolding UP_ring_def
        apply (simp add: Pring_is_ring; fail)
        unfolding ring_hom_ring_axioms_def apply(rule ring_hom)
      proof(rule inj_onI)
        fix x y assume A: "x  carrier (Pring R (insert s S))"
           "y  carrier (Pring R (insert s S))"
           "multivar_poly_to_univ_poly (insert s S) s x = multivar_poly_to_univ_poly (insert s S) s y"
        then have 0: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s x)
       = univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s y)"
          by auto
        have 1: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s x) = x"
          using A by (meson cring.multivar_poly_to_univ_poly_inverse insertI1 is_cring)
        have 2: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s y) = y"
          using A  by (meson insertI1 multivar_poly_to_univ_poly_inverse)
        show "x = y"
          using 0 unfolding 1 2 by auto
      next
        show "domain (UP (Pring R S))"
          apply(rule UP_domain.UP_domain)
          unfolding UP_domain_def by(rule A)
      qed
    qed
  qed
qed

lemma locally_finite:
  assumes "a  carrier (Pring R I)"
  shows "J. J  I  finite J  a  carrier (Pring R J)"
proof(rule Pring_car_induct[of _ I], rule assms)
  have 0: "𝟬Pring R I= 𝟬Pring R {}⇙"
    by (simp add: Pring_zero_eq)
  show "JI. finite J  𝟬Pring R I carrier (Pring R J)"
    unfolding 0
    by (meson Pring_zero_closed empty_subsetI finite.emptyI)
next
  fix m k assume A: "set_mset m  I  k  carrier R"
  obtain J where J_def: "J = set_mset m"
    by blast
  have 0: "k Pring R Imset_to_IP R m = k Pring R Jmset_to_IP R m"
    unfolding J_def by (simp add: Pring_smult)
  have 1: "k Pring R Imset_to_IP R m  carrier (Pring R J)"
    unfolding 0 J_def using A
    by (simp add: Pring_car Pring_smult mset_to_IP_closed poly_scalar_mult_closed)
  show "JI. finite J  k Pring R Imset_to_IP R m  carrier (Pring R J)"
    using J_def 1 A by blast
next
  fix Q m k
  assume A: "Q  carrier (Pring R I)(JI. finite J  Q  carrier (Pring R J))  set_mset m  I  k  carrier R"
  then obtain J where J_def: "JI  finite J  Q  carrier (Pring R J)"
    by blast
  obtain J' where J'_def: "J' = J  set_mset m"
    by blast
  have 0: "finite J'"
    unfolding J'_def using J_def by blast
  have 1: "k Pring R Imset_to_IP R m = k Pring R J'mset_to_IP R m"
    unfolding J'_def using A by (simp add: Pring_smult)
  have 2: "Q  carrier (Pring R J')"
    using J_def unfolding J'_def
    by (metis Pring_car Pring_carrier_subset Un_upper1 subsetD)
  have 3: "Q  k Pring R Imset_to_IP R m  carrier (Pring R J')"
    using 0 1 2 J'_def A
    by (metis Pring_car Pring_smult_closed indexed_pset.indexed_padd mset_to_IP_closed sup.cobounded2)
  show "J'I. finite J'  Q  k Pring R Imset_to_IP R m  carrier (Pring R J')"
    using 3 0
    by (metis A J'_def J_def Un_subset_iff)
qed

lemma(in domain) Pring_is_domain:
  "domain (Pring R I)"
proof(rule domainI, simp add: Pring_is_cring is_cring)
  have 0: "𝟭Pring R I{#} = 𝟭"
    by (simp add: Pring_one indexed_const_def)
  have 1: "𝟬Pring R I{#} = 𝟬"
    by (simp add: Pring_zero indexed_const_def)
  have 2: "𝟭  𝟬"
    by simp
  show "𝟭Pring R I 𝟬Pring R I⇙"
    using 0 1 2 by auto
next
  fix a b assume A: "a Pring R Ib = 𝟬Pring R I⇙"
                    "a  carrier (Pring R I)"
                    "b  carrier (Pring R I)"
  obtain J0 where J0_def: "J0  I  finite J0  a  carrier (Pring R J0)"
    using A locally_finite by blast
  obtain J1 where J1_def: "J1  I  finite J1  b  carrier (Pring R J1)"
    using A locally_finite by blast
  obtain J where J_def: "J = J0  J1"
    by blast
  have J_finite: "finite J"
    using J_def J0_def J1_def by blast
  have 0: "a  carrier (Pring R J)"
    using J0_def unfolding J_def
    by (metis (no_types, lifting) Pring_car Pring_carrier_subset Un_upper1 subset_eq)
  have 1: "b  carrier (Pring R J)"
    using J1_def unfolding J_def
    by (metis Pring_car Pring_carrier_subset in_mono sup.cobounded2)
  have 2: "a Pring R Jb = 𝟬Pring R J⇙"
    using A J_def 0 1 by (metis Pring_mult_eq Pring_zero_eq)
  have 3: "domain (Pring R J)"
    using J_finite Pring_fin_vars_is_domain[of J] by blast
  have 4: "a = 𝟬Pring R J b = 𝟬Pring R J⇙"
    using 3 2 0 1  by (simp add: domain.integral)
  thus "a = 𝟬Pring R I b = 𝟬Pring R I⇙"
    using Pring_zero_eq by blast
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Relabelling of Variables for Indexed Polynomial Rings›
(**************************************************************************************************)
(**************************************************************************************************)

definition relabel_vars :: "'d set  'e set  ('d  'e) 
         ('a, 'd) mvar_poly  ('a, 'e) mvar_poly" where
"relabel_vars I J g = indexed_poly_induced_morphism I (Pring R J) indexed_const (λi. pvar R (g i))"

lemma relabel_vars_is_morphism:
  assumes "g  I  J"
  shows "ring_hom_ring (Pring R I) (Pring R J) (relabel_vars I J g)"
        "i. i  I  relabel_vars I J g (pvar R i) = pvar R (g i)"
        "c. c  carrier R  relabel_vars I J g (indexed_const c) = indexed_const c"
  using Pring_universal_prop(1)[of "Pring R J" "λi. pvar R (g i)" I indexed_const]
    assms unfolding relabel_vars_def
    apply (meson Pi_iff Pring_is_cring Pring_var_closed indexed_const_ring_hom is_cring)
proof-
  have 0: "cring (Pring R J)" " (λi. mset_to_IP R {#g i#})  I  carrier (Pring R J)"
    "ring_hom_ring R (Pring R J) indexed_const"
    using assms Pring_is_cring is_cring apply blast
    apply (smt Pi_iff Pring_var_closed assms var_to_IP_def)
    by (simp add: indexed_const_ring_hom)
  show "i. i  I  indexed_poly_induced_morphism I (Pring R J)
                        indexed_const (λi. pvar R (g i)) (pvar R i) = pvar R (g i)"
    using 0 assms Pring_universal_prop(2)[of "Pring R J" "λi. pvar R (g i)" I indexed_const]
    unfolding relabel_vars_def var_to_IP_def
    by blast
  show "c. c  carrier R  indexed_poly_induced_morphism I (Pring R J)
                          indexed_const (λi. pvar R (g i)) (indexed_const c) = indexed_const c"
    using 0 assms Pring_universal_prop(3)[of "Pring R J" "λi. pvar R (g i)" I indexed_const]
    unfolding  var_to_IP_def
    by blast
qed

lemma relabel_vars_add:
  assumes "g  I  J"
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  shows "relabel_vars I J g (P Pring R IQ) = relabel_vars I J g P Pring R Jrelabel_vars I J g Q"
  using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh ring_hom_add
  by (metis (no_types, lifting))

lemma relabel_vars_mult:
  assumes "g  I  J"
  assumes "P  carrier (Pring R I)"
  assumes "Q  carrier (Pring R I)"
  shows "relabel_vars I J g (P Pring R IQ) = relabel_vars I J g P Pring R Jrelabel_vars I J g Q"
  using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh ring_hom_mult
  by (metis (no_types, lifting))

lemma relabel_vars_closed:
  assumes "g  I  J"
  assumes "P  carrier (Pring R I)"
  shows "relabel_vars I J g P  carrier (Pring R J)"
  using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh
  by (metis ring_hom_closed)

lemma relabel_vars_smult:
  assumes "g  I  J"
  assumes "P  carrier (Pring R I)"
  assumes "a  carrier R"
  shows "relabel_vars I J g (a Pring R IP) = a Pring R Jrelabel_vars I J g P"
proof-
  have 0: "a Pring R IP = indexed_const a Pring R IP"
    by (metis Pring_car Pring_mult Pring_smult assms(2) assms(3) poly_scalar_mult_eq)
  have 1: "a Pring R Jrelabel_vars I J g P = indexed_const a Pring R Jrelabel_vars I J g P"
    by (metis Pring_car Pring_mult Pring_smult assms(1) assms(2) assms(3) poly_scalar_mult_eq relabel_vars_closed)
  show ?thesis using 0 1 assms relabel_vars_mult relabel_vars_is_morphism(3)[of g I J a]
    by (metis indexed_const_closed)
qed

lemma relabel_vars_inverse:
  assumes "g  I  J"
  assumes "h  J  I"
  assumes "i. i  I  h (g i) = i"
  assumes "P  carrier (Pring R I)"
  shows  "relabel_vars J I h (relabel_vars I J g P) = P"
  apply(rule Pring_car_induct''[of _ I])
  using assms(4) apply blast
  using assms
  apply (metis relabel_vars_is_morphism(3))
proof-
  show "p q. p  carrier (Pring R I) 
           q  carrier (Pring R I) 
           relabel_vars J I h (relabel_vars I J g p) = p 
           relabel_vars J I h (relabel_vars I J g q) = q 
           relabel_vars J I h (relabel_vars I J g (p Pring R Iq)) = p Pring R Iq"
  proof- fix p q assume A: " p  carrier (Pring R I)" "q  carrier (Pring R I)"
                "relabel_vars J I h (relabel_vars I J g p) = p"
           "relabel_vars J I h (relabel_vars I J g q) = q"
    show "relabel_vars J I h (relabel_vars I J g (p Pring R Iq)) = p Pring R Iq"
    proof-
      have 0: "relabel_vars I J g (p Pring R Iq) = relabel_vars I J g p Pring R Jrelabel_vars I J g q"
        using A(1) A(2) assms(1) relabel_vars_add by blast
      then show ?thesis
        using assms A relabel_vars_is_morphism(3)[of g I J] relabel_vars_is_morphism(3)[of h J I]
              relabel_vars_closed[of g I J p] relabel_vars_closed[of g I J q]
              relabel_vars_add[of h J I "relabel_vars I J g p" "relabel_vars I J g q"]
        by presburger
    qed
  qed
  show "p i. p  carrier (Pring R I) 
           i  I 
           relabel_vars J I h (relabel_vars I J g p) = p 
        relabel_vars J I h (relabel_vars I J g (p Pring R Ipvar R i)) = p Pring R Ipvar R i"
  proof- fix p i assume A: " p  carrier (Pring R I)"
                         "relabel_vars J I h (relabel_vars I J g p) = p"
                         "i  I"
    show " relabel_vars J I h (relabel_vars I J g (p Pring R Ipvar R i)) = p Pring R Ipvar R i"
    proof-
      have "relabel_vars J I h (relabel_vars I J g (pvar R i)) = pvar R i"
        using assms relabel_vars_is_morphism[of g I J] relabel_vars_is_morphism[of h J I]
        by (metis A(3) funcset_mem )
      then show ?thesis
        using assms  relabel_vars_is_morphism[of g I J] relabel_vars_is_morphism[of h J I]
              relabel_vars_closed[of g I J p] relabel_vars_closed[of g I J "pvar R i"]
              relabel_vars_mult[of g I J "p" "pvar R i"]
              relabel_vars_mult[of h J I "relabel_vars I J g p" "relabel_vars I J g (pvar R i)"]
        by (metis A(1) A(2) A(3) Pring_var_closed)
    qed
  qed
qed

lemma relabel_vars_total_eval:
  assumes "g  I  J"
  assumes "P  carrier (Pring R I)"
  assumes "closed_fun R f"
  shows "total_eval R (f  g) P = total_eval R f (relabel_vars I J g P)"
proof(rule Pring_car_induct''[of P I])
  show "P  carrier (Pring R I)"
    using assms(2) by blast
  show "c. c  carrier R  total_eval R (f  g) (indexed_const c) = total_eval R f (relabel_vars I J g (indexed_const c))"
    by (metis assms(1) relabel_vars_is_morphism(3) total_eval_const)
  show " p q. p  carrier (Pring R I) 
           q  carrier (Pring R I) 
           total_eval R (f  g) p = total_eval R f (relabel_vars I J g p) 
           total_eval R (f  g) q = total_eval R f (relabel_vars I J g q) 
           total_eval R (f  g) (p Pring R Iq) = total_eval R f (relabel_vars I J g (p Pring R Iq))"

  proof- fix p q assume A: "p  carrier (Pring R I)" "q  carrier (Pring R I)"
                           "total_eval R (f  g) p = total_eval R f (relabel_vars I J g p)"
                           "total_eval R (f  g) q = total_eval R f (relabel_vars I J g q)"
    have 0: "closed_fun R (f  g)"
      apply(rule closed_funI)
      using comp_apply[of f g] closed_fun_simp[of f] assms(3) by presburger
    have 1: " (relabel_vars I J g (p Pring R Iq)) =
               (relabel_vars I J g p) Pring R J(relabel_vars I J g q)"
      using A(1) A(2) assms(1) relabel_vars_add by blast
    have 2: "total_eval R f (relabel_vars I J g (p Pring R Iq)) =
      (total_eval R f (relabel_vars I J g p)) R(total_eval R f (relabel_vars I J g q))"
      using total_eval_add[of _ J _ f]
      by (metis "1" A(1) A(2) A(3) A(4) assms(1) assms(3) relabel_vars_closed)
    show "total_eval R (f  g) (p Pring R Iq) = total_eval R f (relabel_vars I J g (p Pring R Iq))"
      using A 0 1 2
      by (metis total_eval_add)
  qed
  show "p i. p  carrier (Pring R I) 
           i  I 
           total_eval R (f  g) p = total_eval R f (relabel_vars I J g p) 
           total_eval R (f  g) (p Pring R Ipvar R i) = total_eval R f (relabel_vars I J g (p Pring R Ipvar R i))"
  proof- fix p i assume A: "p  carrier (Pring R I)" " i  I "
           "total_eval R (f  g) p = total_eval R f (relabel_vars I J g p)"
    have 0: "closed_fun R (f  g)"
            apply(rule closed_funI)
      using comp_apply[of f g] closed_fun_simp[of f] assms(3) by presburger
    have 1: " (relabel_vars I J g (p Pring R I(pvar R i))) =
               (relabel_vars I J g p) Pring R J(relabel_vars I J g (pvar R i))"
      by (meson A(1) A(2) assms(1) local.ring_axioms relabel_vars_mult ring.Pring_var_closed)
    have 2: "total_eval R f (relabel_vars I J g (p Pring R I(pvar R i))) =
      (total_eval R f (relabel_vars I J g p)) R(total_eval R f (relabel_vars I J g (pvar R i)))"
      using total_eval_mult[of "relabel_vars I J g p" J "relabel_vars I J g (pvar R i)"]
      by (metis "1" A(1) A(2) A(3) assms(1) assms(3) is_cring pvar_closed relabel_vars_closed)
    have 3: " total_eval R (f  g) (p Pring R Ipvar R i) =
                total_eval R (f  g) p Rtotal_eval R (f  g)( pvar R i)"
      by (meson "0" A(1) A(2) Pring_var_closed total_eval_mult)
    have 4: "total_eval R (f  g)( pvar R i) = (total_eval R f (relabel_vars I J g (pvar R i)))"
    proof-
      have 40: "total_eval R (f  g)( pvar R i) = (f  g) i"
        using total_eval_var[of "f g" i]
        by (metis "0" var_to_IP_def)
      have 41: "relabel_vars I J g (pvar R i) = pvar R (g i)"
        by (simp add: A(2) assms(1) relabel_vars_is_morphism(2))
      have 42: "total_eval R f (relabel_vars I J g (pvar R i)) = f (g i)"
        using total_eval_var
        by (metis "41" assms(3) var_to_IP_def)
      show ?thesis
        by (metis "40" "42" comp_apply)
    qed
    show " total_eval R (f  g) (p Pring R Ipvar R i) =
            total_eval R f (relabel_vars I J g (p Pring R Ipvar R i))"
      using A 0 1 2 3 4
      by presburger
  qed
qed

end


end