Theory Cring_Poly

theory Cring_Poly
  imports "HOL-Algebra.UnivPoly" "HOL-Algebra.Subrings" Function_Ring
begin

text‹
  This theory extends the material in theoryHOL-Algebra.UnivPoly. The main additions are 
  material on Taylor expansions of polynomials and polynomial derivatives, and various applications
  of the universal property of polynomial evaluation. These include construing polynomials as 
  functions from the base ring to itself, composing one polynomial with another, and extending
  homomorphisms between rings to homomoprhisms of their polynomial rings. These formalizations 
  are necessary components of the proof of Hensel's lemma for $p$-adic integers, and for the 
  proof of $p$-adic quantifier elimination. ›

lemma(in ring) ring_hom_finsum:
  assumes "h  ring_hom R S"
  assumes "ring S"
  assumes "finite I"
  assumes "F  I  carrier R"
  shows "h (finsum R F I) = finsum S (h  F) I"
proof- 
  have I: "(h  ring_hom R S  F  I  carrier R)  h (finsum R F I) = finsum S (h  F) I"
    apply(rule finite_induct, rule assms)
    using assms ring_hom_zero[of h R S] 
    apply (metis abelian_group_def abelian_monoid.finsum_empty ring_axioms ring_def)
  proof(rule)
    fix A a 
    assume A: "finite A" "a  A" "h  ring_hom R S  F  A  carrier R 
           h (finsum R F A) = finsum S (h  F) A" "h  ring_hom R S  F  insert a A  carrier R"
    have 0: "h  ring_hom R S  F  A  carrier R "
      using A by auto 
    have 1: "h (finsum R F A) = finsum S (h  F) A"
      using A 0 by auto 
    have 2: "abelian_monoid S"
      using assms ring_def abelian_group_def by auto
    have 3: "h (F a  finsum R F A) = h (F a) S(finsum S (h  F) A) "
      using ring_hom_add assms finsum_closed 1 A(4) by fastforce
    have 4: "finsum R F (insert a A) = F a  finsum R F A"
      using finsum_insert[of A a F] A assms by auto 
    have 5: "finsum S (h  F) (insert a A) = (h  F) a Sfinsum S (h  F) A"
      apply(rule abelian_monoid.finsum_insert[of S A a "h  F"])
          apply (simp add: "2")
         apply(rule A)
        apply(rule A)
      using ring_hom_closed A  "0" apply fastforce
      using A ring_hom_closed by auto 
    show "h (finsum R F (insert a A)) =
           finsum S (h  F) (insert a A)"
      unfolding 4 5 3 by auto 
  qed
  thus ?thesis using assms by blast 
qed

lemma(in ring) ring_hom_a_inv:
  assumes "ring S"
  assumes "h  ring_hom R S"
  assumes "b  carrier R"
  shows "h ( b) =  Sh b"
proof-
  have "h b Sh ( b) = 𝟬S⇙"
    by (metis (no_types, opaque_lifting) abelian_group.a_inv_closed assms(1) assms(2) assms(3) 
        is_abelian_group local.ring_axioms r_neg ring_hom_add ring_hom_zero)        
  then show ?thesis 
    by (metis (no_types, lifting) abelian_group.minus_equality add.inv_closed assms(1) 
        assms(2) assms(3) ring.is_abelian_group ring.ring_simprules(10) ring_hom_closed)
qed

lemma(in ring) ring_hom_minus:
  assumes "ring S"
  assumes "h  ring_hom R S"
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "h (a  b) = h a Sh b" 
  using assms ring_hom_add[of h R S a "Rb"]
  unfolding a_minus_def 
  using ring_hom_a_inv[of S h b] by auto 

lemma ring_hom_nat_pow:
  assumes "ring R"
  assumes "ring S"
  assumes "h  ring_hom R S"
  assumes "a  carrier R"
  shows "h (a[^]R(n::nat)) = (h a)[^]S(n::nat)"
  using assms by (simp add: ring_hom_ring.hom_nat_pow ring_hom_ringI2)

lemma (in ring) Units_not_right_zero_divisor:
  assumes "a  Units R"
  assumes "b  carrier R"
  assumes "a  b = 𝟬"
  shows  "b = 𝟬"
proof- 
  have "inv a  a  b = 𝟬 "
    using assms Units_closed Units_inv_closed r_null m_assoc[of "inv a" a b] by presburger
  thus ?thesis using assms 
    by (metis Units_l_inv l_one)    
qed

lemma (in ring) Units_not_left_zero_divisor:
  assumes "a  Units R"
  assumes "b  carrier R"
  assumes "b  a = 𝟬"
  shows  "b = 𝟬"
proof- 
  have "b  (a  inv a) = 𝟬 "
    using assms Units_closed Units_inv_closed l_null m_assoc[of b a"inv a"] by presburger
  thus ?thesis using assms 
    by (metis Units_r_inv r_one)    
qed

lemma (in cring) finsum_remove:
  assumes "i. i  Y  f i  carrier R"
  assumes "finite Y"
  assumes "i  Y"
  shows "finsum R f Y = f i  finsum R f (Y - {i})"
proof- 
  have "finsum R f (insert i (Y - {i})) = f i  finsum R f (Y - {i})"
    apply(rule finsum_insert)
    using assms apply blast apply blast using assms apply blast
    using assms by blast 
  thus ?thesis using assms 
    by (metis insert_Diff)
qed


type_synonym degree = nat
text‹The composition of two ring homomorphisms is a ring homomorphism›
lemma ring_hom_compose:                                                                                                           
  assumes "ring R"
  assumes "ring S"
  assumes "ring T"
  assumes "h  ring_hom R S"
  assumes "g  ring_hom S T"
  assumes "c. c  carrier R  f c = g (h c)"
  shows "f   ring_hom R T"
proof(rule ring_hom_memI)
  show "x. x  carrier R  f x  carrier T"
    using assms  by (metis ring_hom_closed)
  show " x y. x  carrier R  y  carrier R  f (x Ry) = f x Tf y"
  proof-
    fix x y
    assume A: "x   carrier R" "y   carrier R"
    show "f (x Ry) = f x Tf y"
    proof-
      have  "f (x Ry) = g (h (x Ry))"
        by (simp add: A(1) A(2) assms(1) assms(6) ring.ring_simprules(5))
      then have  "f (x Ry) = g ((h x) S(h y))"
        using A(1) A(2) assms(4) ring_hom_mult by fastforce
      then have  "f (x Ry) = g (h x) Tg (h y)"
        using A(1) A(2) assms(4) assms(5) ring_hom_closed ring_hom_mult by fastforce
      then show ?thesis 
        by (simp add: A(1) A(2) assms(6))
    qed
  qed
  show "x y. x  carrier R  y  carrier R  f (x Ry) = f x Tf y"
  proof-
    fix x y
    assume A: "x   carrier R" "y   carrier R"
    show "f (x Ry) = f x Tf y"
    proof-
      have  "f (x Ry) = g (h (x Ry))"
        by (simp add: A(1) A(2) assms(1) assms(6) ring.ring_simprules(1))
      then have  "f (x Ry) = g ((h x) S(h y))"
        using A(1) A(2) assms(4) ring_hom_add by fastforce
      then have  "f (x Ry) = g (h x) Tg (h y)"
        by (metis (no_types, opaque_lifting) A(1) A(2) assms(4) assms(5) ring_hom_add ring_hom_closed)
      then show ?thesis
        by (simp add: A(1) A(2) assms(6))
    qed
  qed
  show "f 𝟭R= 𝟭T⇙" 
    by (metis assms(1) assms(4) assms(5) assms(6) ring.ring_simprules(6) ring_hom_one)
qed


(**************************************************************************************************)
(**************************************************************************************************)
section‹Basic Notions about Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)
context UP_ring
begin


text‹rings are closed under monomial terms›
lemma monom_term_car:
  assumes "c  carrier R"
  assumes "x  carrier R"
  shows "c  x[^](n::nat)  carrier R"
  using assms monoid.nat_pow_closed 
  by blast

text‹Univariate polynomial ring over R›

lemma P_is_UP_ring:
"UP_ring R"
  by (simp add: UP_ring_axioms)

text‹Degree function›
abbreviation(input) degree  where
"degree f   deg R f"

lemma UP_car_memI:
  assumes "n. n > k  p n = 𝟬"
  assumes "n. p n  carrier R"
  shows "p  carrier P"
proof-
  have "bound 𝟬 k p"
    by (simp add: assms(1) bound.intro)
  then show ?thesis 
    by (metis (no_types, lifting) P_def UP_def assms(2) mem_upI partial_object.select_convs(1))
qed

lemma(in UP_cring) UP_car_memI':
  assumes "x. g x  carrier R"
  assumes "x. x > k  g x = 𝟬"
  shows "g  carrier (UP R)"
proof-
  have "bound 𝟬 k g"
    using assms unfolding bound_def by blast 
  then show ?thesis 
    using P_def UP_car_memI assms(1) by blast
qed

lemma(in UP_cring) UP_car_memE:
  assumes "g  carrier (UP R)"
  shows "x. g x  carrier R"
        "x. x > (deg R g)  g x = 𝟬"
  using P_def assms UP_def[of R] apply (simp add: mem_upD)
  using assms UP_def[of R] up_def[of R] 
  by (smt R.ring_axioms UP_ring.deg_aboveD UP_ring.intro partial_object.select_convs(1) restrict_apply up_ring.simps(2))
  
end

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Lemmas About Coefficients›
  (**************************************************************************************************)
  (**************************************************************************************************)

context UP_ring
begin
text‹The goal here is to reduce dependence on the function coeff from Univ\_Poly, in favour of using
a polynomial itself as its coefficient function.›

lemma coeff_simp:
  assumes "f  carrier P"
  shows "coeff (UP R)  f = f "
proof fix x show "coeff (UP R) f x = f x" 
    using assms P_def UP_def[of R] by auto 
qed

text‹Coefficients are in R›

lemma cfs_closed:
  assumes "f  carrier P"
  shows "f n  carrier R"
  using assms coeff_simp[of f]  P_def coeff_closed 
  by fastforce

lemma cfs_monom:
  "a  carrier R ==> (monom P a m) n = (if m=n then a else 𝟬)"
using coeff_simp  P_def coeff_monom monom_closed by auto

lemma cfs_zero [simp]: "𝟬Pn = 𝟬" 
  using P_def UP_zero_closed coeff_simp coeff_zero by auto

lemma cfs_one [simp]: "𝟭Pn = (if n=0 then 𝟭 else 𝟬)"
  by (metis P_def R.one_closed UP_ring.cfs_monom UP_ring_axioms monom_one)

lemma cfs_smult [simp]:
  "[| a  carrier R; p  carrier P |] ==> (a Pp) n = a  p n"
  using P_def UP_ring.coeff_simp UP_ring_axioms UP_smult_closed coeff_smult by fastforce
  
lemma cfs_add [simp]:
  "[| p  carrier P; q  carrier P |] ==> (p Pq) n = p n  q n"
  by (metis P.add.m_closed P_def UP_ring.coeff_add UP_ring.coeff_simp UP_ring_axioms)

lemma cfs_a_inv [simp]:
  assumes R: "p  carrier P"
  shows "(Pp) n =  (p n)"
  using P.add.inv_closed P_def UP_ring.coeff_a_inv UP_ring.coeff_simp UP_ring_axioms assms 
  by fastforce

lemma cfs_minus [simp]:
  "[| p  carrier P; q  carrier P |] ==> (p Pq) n = p n  q n"
  using P.minus_closed P_def coeff_minus coeff_simp by auto

lemma cfs_monom_mult_r:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "(monom P a n Pp) (k + n) = a  p k"
  using coeff_monom_mult assms P.m_closed P_def coeff_simp monom_closed by auto

lemma(in UP_cring) cfs_monom_mult_l:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "(p Pmonom P a n) (k + n) = a  p k"
  using UP_m_comm assms(1) assms(2) cfs_monom_mult_r by auto

lemma(in UP_cring) cfs_monom_mult_l':
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "m  n"
  shows "(f P(monom P a n)) m = a  (f (m - n))"
  using cfs_monom_mult_l[of f a n "m-n"] assms 
  by simp

lemma(in UP_cring) cfs_monom_mult_r':
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "m  n"
  shows "((monom P a n) Pf) m = a  (f (m - n))"
  using cfs_monom_mult_r[of f a n "m-n"] assms 
  by simp 
end
  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Degree Bound Lemmas›
  (**************************************************************************************************)
  (**************************************************************************************************)

context UP_ring
begin

lemma bound_deg_sum:
  assumes " f  carrier P"
  assumes "g  carrier P"
  assumes "degree f   n"
  assumes "degree g  n"
  shows "degree (f Pg)  n"
  using P_def UP_ring_axioms assms(1) assms(2) assms(3) assms(4) 
  by (meson deg_add max.boundedI order_trans)

lemma bound_deg_sum':
  assumes " f  carrier P"
  assumes "g  carrier P"
  assumes "degree f < n"
  assumes "degree g < n"
  shows "degree (f Pg) < n"
  using P_def UP_ring_axioms assms(1) assms(2) 
        assms(3) assms(4) 
  by (metis bound_deg_sum le_neq_implies_less less_imp_le_nat not_less)        

lemma equal_deg_sum:
  assumes " f  carrier P"
  assumes "g  carrier P"
  assumes "degree f < n"
  assumes "degree g = n"
  shows "degree (f Pg) = n"
proof-
  have 0: "degree (f Pg) n"
    using assms bound_deg_sum 
          P_def UP_ring_axioms by auto
  show "degree (f Pg) = n"
  proof(rule ccontr)
    assume "degree (f Pg)  n "
    then have 1: "degree (f Pg) < n"
      using 0 by auto 
    have 2: "degree (Pf) < n"
      using assms by simp
    have 3: "g = (f Pg) P(Pf)"
      using assms 
      by (simp add: P.add.m_comm P.r_neg1)
    then show False using 1 2 3 assms 
      by (metis UP_a_closed UP_a_inv_closed deg_add leD le_max_iff_disj)
  qed
qed

lemma equal_deg_sum':
  assumes "f  carrier P"
  assumes "g  carrier P"
  assumes "degree g < n"
  assumes "degree f = n"
  shows "degree (f Pg) = n"
  using P_def UP_a_comm UP_ring.equal_deg_sum UP_ring_axioms assms(1) assms(2) assms(3) assms(4)
  by fastforce

lemma degree_of_sum_diff_degree:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree q < degree p"
  shows "degree (p Pq) = degree p" 
  by(rule equal_deg_sum', auto simp: assms)

lemma degree_of_difference_diff_degree:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree q < degree p"
  shows "degree (p Pq) = degree p" 
proof-
  have A: "(p Pq) = p P(Pq)" 
    by (simp add: P.minus_eq)
  have "degree (Pq) = degree q " 
    by (simp add: assms(2))
  then show ?thesis 
    using assms A 
    by (simp add: degree_of_sum_diff_degree)
qed

lemma (in UP_ring) deg_diff_by_const:
  assumes "g  carrier (UP R)"
  assumes "a  carrier R"
  assumes "h = g UP Rup_ring.monom (UP R) a 0"
  shows "deg R g = deg R h"
  unfolding assms using assms 
  by (metis P_def UP_ring.bound_deg_sum UP_ring.deg_monom_le UP_ring.monom_closed UP_ring_axioms degree_of_sum_diff_degree gr_zeroI not_less)

lemma (in UP_ring) deg_diff_by_const':
  assumes "g  carrier (UP R)"
  assumes "a  carrier R"
  assumes "h = g UP Rup_ring.monom (UP R) a 0"
  shows "deg R g = deg R h"
  apply(rule deg_diff_by_const[of _  " a"])
  using assms apply blast 
  using assms apply blast
  by (metis P.minus_eq P_def assms(2) assms(3) monom_a_inv)

lemma(in UP_ring) deg_gtE:
  assumes "p  carrier P"
  assumes "i > deg R p"
  shows "p i = 𝟬"
  using assms  P_def coeff_simp deg_aboveD by metis
end

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Leading Term Function›
  (**************************************************************************************************)
  (**************************************************************************************************)

definition leading_term where 
"leading_term R f = monom (UP R) (f (deg R f)) (deg R f)"

context UP_ring
begin

abbreviation(input) ltrm  where
"ltrm f  monom P (f (deg R f)) (deg R f)"

text‹leading term is a polynomial›
lemma ltrm_closed:
  assumes "f  carrier P"
  shows "ltrm f  carrier P"
  using assms 
  by (simp add: cfs_closed)  
  
text‹Simplified coefficient function description for leading term›
lemma ltrm_coeff:
  assumes "f  carrier P"
  shows "coeff P (ltrm f) n  = (if (n = degree f) then (f (degree f)) else 𝟬)"
    using assms 
    by (simp add: cfs_closed)

lemma ltrm_cfs:
  assumes "f  carrier P"
  shows "(ltrm f) n  = (if (n = degree f) then (f (degree f)) else 𝟬)"
  using assms 
  by (simp add: cfs_closed cfs_monom)

lemma ltrm_cfs_above_deg:
  assumes "f  carrier P"
  assumes "n > degree f"
  shows "ltrm f n = 𝟬"
  using assms 
  by (simp add: ltrm_cfs)

text‹The leading term of f has the same degree as f›

lemma deg_ltrm:
  assumes "f  carrier P"
  shows "degree (ltrm f) = degree f"
  using assms 
  by (metis P_def UP_ring.lcoeff_nonzero_deg UP_ring_axioms cfs_closed coeff_simp deg_const deg_monom)

text‹Subtracting the leading term yields a drop in degree›

lemma minus_ltrm_degree_drop:
  assumes "f  carrier P"
  assumes "degree f = Suc n"
  shows "degree (f P(ltrm f))  n"
proof(rule UP_ring.deg_aboveI)
  show C0: "UP_ring R" 
    by (simp add: UP_ring_axioms)    
  show C1: "f Pltrm f  carrier (UP R)"
    using assms ltrm_closed P.minus_closed P_def 
    by blast     
  show C2: "m. n < m  coeff (UP R) (f Pltrm f) m = 𝟬"
  proof-
    fix m
    assume A: "n<m"
    show "coeff (UP R) (f Pltrm f) m = 𝟬"
    proof(cases " m = Suc n")
      case True
      have B: "f m  carrier R" 
        using UP.coeff_closed P_def assms(1) cfs_closed by blast       
      have "m = degree f"
        using True by (simp add: assms(2))
      then have "f m = (ltrm f) m" 
        using ltrm_cfs assms(1) by auto 
      then have "(f m) R( ltrm f) m = 𝟬"
        using B UP_ring_def P_is_UP_ring 
          B R.add.r_inv R.is_abelian_group abelian_group.minus_eq by fastforce
      then have "(f UP Rltrm f) m = 𝟬"
        by (metis C1 ltrm_closed P_def assms(1) coeff_minus coeff_simp)        
      then show ?thesis 
        using C1 P_def UP_ring.coeff_simp UP_ring_axioms by fastforce
    next
      case False
      have D0: "m > degree f" using False 
        using A assms(2) by linarith
       have B: "f m  carrier R" 
         using UP.coeff_closed P_def assms(1) cfs_closed 
         by blast
       have "f m = (ltrm f) m" 
         using D0 ltrm_cfs_above_deg P_def assms(1) coeff_simp deg_aboveD 
         by auto 
      then show ?thesis 
        by (metis B ltrm_closed P_def R.r_neg UP_ring.coeff_simp UP_ring_axioms a_minus_def assms(1) coeff_minus)       
    qed
  qed
qed

lemma ltrm_decomp:
  assumes "f  carrier P"
  assumes "degree f >(0::nat)"
  obtains g where "g  carrier P  f = g P(ltrm f)  degree g < degree f"
proof-
  have 0: "f P(ltrm f)  carrier P"
    using ltrm_closed assms(1) by blast    
  have 1: "f = (f P(ltrm f)) P(ltrm f)"
    using assms 
    by (metis "0" ltrm_closed P.add.inv_solve_right P.minus_eq)
  show ?thesis using assms 0 1 minus_ltrm_degree_drop[of f]
    by (metis ltrm_closed Suc_diff_1 Suc_n_not_le_n deg_ltrm equal_deg_sum' linorder_neqE_nat that)
qed

text‹leading term of a sum›
lemma coeff_of_sum_diff_degree0:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree q < n"
  shows "(p Pq) n = p n"
  using assms P_def UP_ring.deg_aboveD UP_ring_axioms cfs_add coeff_simp cfs_closed deg_aboveD 
  by auto

lemma coeff_of_sum_diff_degree1:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree q < degree p"
  shows "(p Pq) (degree p) = p (degree p)"
  using assms(1) assms(2) assms(3) coeff_of_sum_diff_degree0 by blast



lemma ltrm_of_sum_diff_degree:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree p > degree q"
  shows "ltrm (p Pq) = ltrm p" 
  unfolding leading_term_def 
  using assms(1) assms(2) assms(3) coeff_of_sum_diff_degree1 degree_of_sum_diff_degree 
  by presburger

text‹leading term of a monomial›

lemma ltrm_monom:
  assumes "a  carrier R"
  assumes "f = monom P a n"
  shows "ltrm f = f"
  unfolding leading_term_def
  by (metis P_def UP_ring.cfs_monom UP_ring.monom_zero UP_ring_axioms assms(1) assms(2) deg_monom)  

lemma ltrm_monom_simp:
  assumes "a  carrier R"
  shows "ltrm (monom P a n) = monom P a n"
  using assms ltrm_monom by auto

lemma ltrm_inv_simp[simp]:
  assumes "f  carrier P"
  shows "ltrm (ltrm f) = ltrm f"
  by (metis assms deg_ltrm ltrm_cfs)

lemma ltrm_deg_0:
  assumes "p  carrier P"
  assumes "degree p = 0"
  shows "ltrm p = p"
  using ltrm_monom assms P_def UP_ring.deg_zero_impl_monom UP_ring_axioms coeff_simp
  by fastforce

lemma ltrm_prod_ltrm:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "ltrm ((ltrm p) P(ltrm q)) = (ltrm p) P(ltrm q)"
  using ltrm_monom R.m_closed assms(1) assms(2) cfs_closed monom_mult
  by metis 

text‹lead coefficient function›

abbreviation(input) lcf where
"lcf p  p (deg R p)"

lemma(in UP_ring) lcf_ltrm:
"ltrm p = monom P (lcf p) (degree p)"
  by auto

lemma lcf_closed:
  assumes "f  carrier P"
  shows "lcf f  carrier R"
  by (simp add: assms cfs_closed)

lemma(in UP_cring) lcf_monom:
  assumes "a  carrier R"
  shows "lcf (monom P a n) = a" "lcf (monom (UP R) a n) = a"
  using assms deg_monom cfs_monom apply fastforce
  by (metis UP_ring.cfs_monom UP_ring.deg_monom UP_ring_axioms assms)


end 

text‹Function which truncates a polynomial by removing the leading term›

definition truncate where
"truncate R f = f (UP R)(leading_term R f)"

context UP_ring
begin 

abbreviation(input) trunc where
"trunc  truncate R"

lemma trunc_closed:
  assumes "f  carrier P"
  shows "trunc f  carrier P"
  using assms unfolding truncate_def 
  by (metis ltrm_closed P_def UP_ring.UP_ring UP_ring_axioms leading_term_def ring.ring_simprules(4))

lemma trunc_simps:
  assumes "f  carrier P"
  shows "f = (trunc f) P(ltrm f)"
        "f P(trunc f) = ltrm f"   
  apply (metis ltrm_closed P.add.inv_solve_right P.minus_closed P_def a_minus_def assms Cring_Poly.truncate_def leading_term_def)
  using trunc_closed[of f] ltrm_closed[of f] P_def P.add.inv_solve_right[of "ltrm f" f "trunc f"]
        assms  unfolding UP_cring_def
  by (metis P.add.inv_closed P.add.m_lcomm P.add.r_inv_ex P.minus_eq P.minus_minus
      P.r_neg2 P.r_zero Cring_Poly.truncate_def leading_term_def)

lemma trunc_zero:
  assumes "f  carrier P"
  assumes "degree f = 0"
  shows "trunc f = 𝟬P⇙"
  unfolding truncate_def 
  using assms ltrm_deg_0[of f] 
  by (metis P.r_neg P_def a_minus_def leading_term_def)

lemma trunc_degree:
  assumes "f  carrier P"
  assumes "degree f > 0"
  shows "degree (trunc f) < degree f"
  unfolding truncate_def using assms 
  by (metis ltrm_closed ltrm_decomp P.add.right_cancel Cring_Poly.truncate_def trunc_closed trunc_simps(1))

text‹The coefficients of trunc agree with f for small degree›

lemma trunc_cfs:
  assumes "p  carrier P"
  assumes "n < degree p"
  shows " (trunc p) n = p n"
  using P_def assms(1) assms(2) unfolding truncate_def 
  by (smt ltrm_closed ltrm_cfs R.minus_zero R.ring_axioms UP_ring.cfs_minus
      UP_ring_axioms a_minus_def cfs_closed leading_term_def nat_neq_iff ring.ring_simprules(15))

text‹monomial predicate›

definition is_UP_monom where
"is_UP_monom = (λf. f  carrier (UP R)  f = ltrm f)"

lemma is_UP_monomI:
  assumes "a  carrier R"
  assumes "p = monom P a n"
  shows "is_UP_monom p"
  using assms(1) assms(2) is_UP_monom_def ltrm_monom P_def monom_closed 
  by auto
  
lemma is_UP_monomI':
  assumes "f  carrier (UP R)"
  assumes "f = ltrm f"
  shows "is_UP_monom f"
  using assms P_def unfolding is_UP_monom_def by blast 

lemma monom_is_UP_monom:
  assumes "a  carrier R"
  shows "is_UP_monom (monom P a n)" "is_UP_monom (monom (UP R) a n)"
  using assms P_def ltrm_monom_simp monom_closed
  unfolding is_UP_monom_def  
  by auto

lemma is_UP_monomE:
  assumes "is_UP_monom f"
  shows "f  carrier P" "f = monom P (lcf f) (degree f)"  "f = monom (UP R) (lcf f) (degree f)"
  using assms unfolding is_UP_monom_def 
  by(auto simp: P_def ) 

lemma ltrm_is_UP_monom:
  assumes "p  carrier P"
  shows "is_UP_monom (ltrm p)"
  using assms 
  by (simp add: cfs_closed monom_is_UP_monom(1))

lemma is_UP_monom_mult:
  assumes "is_UP_monom p"
  assumes "is_UP_monom q"
  shows "is_UP_monom (p Pq)"
  apply(rule is_UP_monomI')
  using assms is_UP_monomE P_def UP_mult_closed 
  apply simp  
  using assms is_UP_monomE[of p] is_UP_monomE[of q] 
        P_def monom_mult 
  by (metis lcf_closed ltrm_monom R.m_closed)
end 

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Properties of Leading Terms and Leading Coefficients in Commutative Rings and Domains›
  (**************************************************************************************************)
  (**************************************************************************************************)

context UP_cring
begin

lemma cring_deg_mult:
  assumes "q  carrier P"
  assumes "p  carrier P"
  assumes "lcf q  lcf p 𝟬"
  shows "degree (q Pp) = degree p + degree q"
proof-
  have "q Pp = (trunc q Pltrm q) P(trunc p Pltrm p)"
    using assms(1) assms(2) trunc_simps(1) by auto
  then have "q Pp = (trunc q Pltrm q) P(trunc p Pltrm p)"
    by linarith
  then have 0: "q Pp = (trunc q  P(trunc p Pltrm p)) P( ltrm q P(trunc p Pltrm p))"
    by (simp add: P.l_distr assms(1) assms(2) ltrm_closed trunc_closed)
  have 1: "(trunc q  P(trunc p Pltrm p)) (degree p + degree q) = 𝟬"
  proof(cases "degree q = 0")
    case True
    then show ?thesis 
      using assms(1) assms(2) trunc_simps(1) trunc_zero by auto           
  next
    case False
    have "degree ((trunc q)  Pp)  degree (trunc q) + degree p"
      using assms trunc_simps[of q]  deg_mult_ring[of "trunc q" p]  trunc_closed 
      by blast      
    then have "degree (trunc q  P(trunc p Pltrm p)) < degree q + degree p"
      using False assms(1) assms(2) trunc_degree trunc_simps(1) by fastforce
    then show ?thesis 
      by (metis P_def UP_mult_closed UP_ring.coeff_simp UP_ring_axioms
          add.commute assms(1) assms(2) deg_belowI not_less trunc_closed trunc_simps(1))      
  qed
  have 2: "(q Pp) (degree p + degree q) = 
                   ( ltrm q P(trunc p Pltrm p)) (degree p + degree q)"
    using 0 1 assms cfs_closed trunc_closed by auto       
  have 3: "(q Pp) (degree p + degree q) = 
                   ( ltrm q Ptrunc p) (degree p + degree q)  ( ltrm q Pltrm p) (degree p + degree q)"
    by (simp add: "2" ltrm_closed UP_r_distr assms(1) assms(2) trunc_closed)   
  have 4: "( ltrm q Ptrunc p) (degree p + degree q) = 𝟬"
  proof(cases "degree p = 0")
    case True
    then show ?thesis 
      using "2" "3" assms(1) assms(2) cfs_closed ltrm_closed trunc_zero by auto            
  next
    case False
    have "degree ( ltrm q Ptrunc p)  degree (ltrm q) + degree (trunc p)"
      using assms trunc_simps deg_mult_ring ltrm_closed trunc_closed by presburger              
    then have "degree ( ltrm q Ptrunc p) < degree q + degree p"
      using False assms(1) assms(2) trunc_degree trunc_simps(1) deg_ltrm by fastforce 
    then show ?thesis 
      by (metis ltrm_closed P_def UP_mult_closed UP_ring.coeff_simp UP_ring_axioms add.commute assms(1) assms(2) deg_belowI not_less trunc_closed)          
  qed
  have 5: "(q Pp) (degree p + degree q) = ( ltrm q Pltrm p) (degree p + degree q)"
    by (simp add: "3" "4" assms(1) assms(2) cfs_closed)   
  have 6: "ltrm q Pltrm p = monom P (lcf q  lcf p) (degree p + degree q)"
    unfolding  leading_term_def 
    by (metis P_def UP_ring.monom_mult UP_ring_axioms add.commute assms(1) assms(2) cfs_closed)
  have 7: "( ltrm q Pltrm p) (degree p + degree q) 𝟬"
    using 5  6 assms 
    by (metis R.m_closed cfs_closed cfs_monom)   
  have 8: "degree (q Pp) degree p + degree q"
    using 5 6 7 P_def UP_mult_closed assms(1) assms(2) 
    by (simp add: UP_ring.coeff_simp UP_ring_axioms deg_belowI)    
  then show ?thesis 
    using assms(1) assms(2) deg_mult_ring by fastforce
qed

text‹leading term is multiplicative›

lemma ltrm_of_sum_diff_deg:
  assumes "q  carrier P"
  assumes "a  carrier R"
  assumes "a 𝟬"
  assumes "degree q < n"
  assumes "p = q P(monom P a n)"
  shows "ltrm p =  (monom P a n)"
proof-
  have 0: "degree  (monom P a n) = n" 
    by (simp add: assms(2) assms(3))
  have 1: "(monom P a n)  carrier P"
    using assms(2) by auto
  have 2: "ltrm ((monom P a n) Pq) = ltrm (monom P a n)"
    using assms ltrm_of_sum_diff_degree[of "(monom P a n)" q] 1  "0" by linarith
  then show ?thesis 
    using UP_a_comm assms(1) assms(2) assms(5) ltrm_monom by auto
qed

lemma(in UP_cring) ltrm_smult_cring:
  assumes "p  carrier P"
  assumes "a  carrier R"
  assumes "lcf p  a  𝟬"
  shows "ltrm (a Pp) = aP(ltrm p)"
  using assms 
  by (smt lcf_monom(1) P_def R.m_closed R.m_comm cfs_closed cfs_smult coeff_simp
      cring_deg_mult deg_monom deg_ltrm monom_closed monom_mult_is_smult monom_mult_smult) 

lemma(in UP_cring) deg_zero_ltrm_smult_cring:
  assumes "p  carrier P"
  assumes "a  carrier R"
  assumes "degree p = 0"
  shows "ltrm (a Pp) = aP(ltrm p)"
  by (metis ltrm_deg_0 assms(1) assms(2) assms(3) deg_smult_decr le_0_eq module.smult_closed module_axioms)

lemma(in UP_domain) ltrm_smult:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "ltrm (a Pp) = aP(ltrm p)"
  by (metis lcf_closed ltrm_closed ltrm_smult_cring P_def R.integral_iff UP_ring.deg_ltrm 
      UP_ring_axioms UP_smult_zero assms(1) assms(2) cfs_zero deg_nzero_nzero deg_zero_ltrm_smult_cring monom_zero)

lemma(in UP_cring) cring_ltrm_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "lcf p  lcf q  𝟬"
  shows "ltrm (p Pq) = (ltrm p) P(ltrm q)"
proof(cases "degree p = 0  degree q = 0")
  case True
  then show ?thesis 
    by (smt ltrm_closed ltrm_deg_0 ltrm_smult_cring R.m_comm UP_m_comm assms(1) assms(2) assms(3) cfs_closed monom_mult_is_smult)
next
  case False
    obtain q0 where q0_def: "q0 = trunc q" 
      by simp
    obtain p0 where p0_def: "p0 = trunc p" 
      by simp
    have Pq: "degree q0 < degree q"
      using False P_def assms(2) q0_def trunc_degree by blast          
    have Pp: "degree p0 < degree p"
      using False P_def assms(1) p0_def trunc_degree by blast
    have "p Pq = (p0 Pltrm(p)) P(q0 Pltrm(q))"
      using assms(1) assms(2) p0_def q0_def trunc_simps(1) by auto
    then have P0: "p Pq = ((p0 Pltrm(p)) Pq0) P((p0 Pltrm(p))Pltrm(q))"
      by (simp add: P.r_distr assms(1) assms(2) ltrm_closed p0_def q0_def trunc_closed)
    have P1: "degree ((p0 Pltrm(p)) Pq0) < degree ((p0 Pltrm(p))Pltrm(q))"
    proof-
      have LHS: "degree ((p0 Pltrm(p)) Pq0)  degree p + degree q0 "
      proof(cases "q0 = 𝟬P⇙")
        case True
        then show ?thesis 
          using assms(1) p0_def trunc_simps(1) by auto
      next
        case False
        then show ?thesis 
          using assms(1) assms(2) deg_mult_ring  p0_def 
            q0_def trunc_simps(1) trunc_closed by auto
      qed
      have RHS: "degree ((p0 Pltrm(p))Pltrm(q)) = degree p + degree q"
        using assms(1) assms(2) deg_mult_ring ltrm_closed p0_def trunc_simps(1) 
        by (smt P_def UP_cring.lcf_monom(1) UP_cring.cring_deg_mult UP_cring_axioms add.commute assms(3) cfs_closed deg_ltrm)                    
        then show ?thesis 
          using RHS LHS  Pq 
          by linarith    
    qed
    then have P2: "ltrm (p Pq) = ltrm ((p0 Pltrm(p))Pltrm(q))"
      using P0 P1  
      by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm 
          UP_mult_closed assms(1) assms(2) p0_def q0_def trunc_closed trunc_simps(1))      
    have P3: " ltrm ((p0 Pltrm(p))Pltrm(q)) = ltrm p Pltrm q"
    proof-
      have Q0: "((p0 Pltrm(p))Pltrm(q)) = (p0 Pltrm(q)) P(ltrm(p))Pltrm(q)"
        by (simp add: P.l_distr assms(1) assms(2) ltrm_closed p0_def trunc_closed)
      have Q1: "degree ((p0 Pltrm(q)) ) < degree ((ltrm(p))Pltrm(q))"
      proof(cases "p0 = 𝟬P⇙")
        case True
        then show ?thesis 
          using P1 assms(1) assms(2)  ltrm_closed by auto
      next
        case F: False
        then show ?thesis
          proof-
            have LHS: "degree ((p0 Pltrm(q))) < degree p + degree q"
              using False F  Pp assms(1) assms(2) deg_nzero_nzero 
                 deg_ltrm ltrm_closed p0_def trunc_closed 
              by (smt add_le_cancel_right deg_mult_ring le_trans not_less)   
            have RHS: "degree ((ltrm(p))Pltrm(q)) = degree p + degree q" 
              using cring_deg_mult[of "ltrm p" "ltrm q"] assms 
              by (simp add: ltrm_closed ltrm_cfs deg_ltrm)                      
            then show ?thesis using LHS RHS by auto 
        qed
      qed
      have Q2: "ltrm ((p0 Pltrm(p))Pltrm(q)) = ltrm ((ltrm(p))Pltrm(q))" 
        using Q0 Q1 
        by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm 
            UP_mult_closed assms(1) assms(2) p0_def trunc_closed)        
      show ?thesis using ltrm_prod_ltrm Q0 Q1 Q2 
        by (simp add: assms(1) assms(2))
    qed
    then show ?thesis 
      by (simp add: P2)
qed

lemma(in UP_domain) ltrm_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "ltrm (p Pq) = (ltrm p) P(ltrm q)"
  using cring_ltrm_mult assms 
  by (smt ltrm_closed ltrm_deg_0 cfs_closed deg_nzero_nzero deg_ltrm local.integral_iff monom_mult monom_zero)

lemma lcf_deg_0:
  assumes "degree p = 0"
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "(p Pq) = (lcf p)Pq" 
  using P_def assms(1) assms(2) assms(3) 
  by (metis ltrm_deg_0 cfs_closed monom_mult_is_smult)
  
text‹leading term powers›

lemma (in domain) nonzero_pow_nonzero:
  assumes "a  carrier R"
  assumes "a 𝟬"
  shows "a[^](n::nat)  𝟬"  
proof(induction n)
  case 0
  then show ?case 
    by auto
next
  case (Suc n)
  fix n::nat
  assume IH: "a[^] n  𝟬" 
  show "a[^] (Suc n)  𝟬" 
  proof-
    have "a[^] (Suc n) = a[^] n  a"
      by simp
    then show ?thesis using assms IH 
      using IH assms(1) assms(2) local.integral by auto
  qed
qed

lemma (in UP_cring) cring_monom_degree:
  assumes "a  (carrier R)"
  assumes "p = monom P a m"
  assumes "a[^]n  𝟬"
  shows "degree (p[^]Pn) = n*m"
  by (simp add: assms(1) assms(2) assms(3) monom_pow)
  
lemma (in UP_domain) monom_degree:
  assumes "a 𝟬"
  assumes "a  (carrier R)"
  assumes "p = monom P a m"
  shows "degree (p[^]Pn) = n*m"
  by (simp add: R.domain_axioms assms(1) assms(2) assms(3) domain.nonzero_pow_nonzero monom_pow)
  
lemma(in UP_cring)  cring_pow_ltrm:
  assumes "p  carrier P" 
  assumes "lcf p [^]n  𝟬"
  shows "ltrm (p[^]P(n::nat)) = (ltrm p)[^]Pn"
proof-
  have "lcf p [^]n  𝟬  ltrm (p[^]P(n::nat)) = (ltrm p)[^]Pn"
  proof(induction n)
    case 0
    then show ?case 
      using P.ring_simprules(6) P.nat_pow_0 cfs_one deg_one monom_one by presburger
  next
    case (Suc n) fix n::nat
    assume IH : "(lcf p [^] n  𝟬  ltrm (p [^]Pn) = ltrm p [^]Pn)"
    assume A: "lcf p [^] Suc n  𝟬"
    have a: "ltrm (p [^]Pn) = ltrm p [^]Pn"
      apply(cases "lcf p [^] n = 𝟬")
      using A lcf_closed assms(1) apply auto[1]
      by(rule IH)
    have 0: "lcf (ltrm (p [^]Pn)) = lcf p [^] n"
      unfolding a 
      by (simp add: lcf_monom(1) assms(1) cfs_closed monom_pow)
    then have 1: "lcf (ltrm (p [^]Pn))  lcf p  𝟬"
      using assms A R.nat_pow_Suc IH  by metis   
    then show "ltrm (p [^]PSuc n) = ltrm p [^]PSuc n"
      using IH 0 assms(1) cring_ltrm_mult cfs_closed 
      by (smt A lcf_monom(1) ltrm_closed P.nat_pow_Suc2 P.nat_pow_closed R.nat_pow_Suc2 a)    
  qed
  then show ?thesis 
    using assms(2) by blast
qed

lemma(in UP_cring)  cring_pow_deg:
  assumes "p  carrier P" 
  assumes "lcf p [^]n  𝟬"
  shows "degree (p[^]P(n::nat)) = n*degree p"
proof-
  have "degree ( (ltrm p)[^]Pn) = n*degree p"
    using assms(1) assms(2) cring_monom_degree lcf_closed lcf_ltrm by auto
  then show ?thesis 
    using assms cring_pow_ltrm 
    by (metis P.nat_pow_closed P_def UP_ring.deg_ltrm UP_ring_axioms)
qed

lemma(in UP_cring)  cring_pow_deg_bound:
  assumes "p  carrier P" 
  shows "degree (p[^]P(n::nat))  n*degree p"
  apply(induction n)
   apply (metis Group.nat_pow_0 deg_one le_zero_eq mult_is_0)  
  using deg_mult_ring[of _ p]
  by (smt P.nat_pow_Suc2 P.nat_pow_closed ab_semigroup_add_class.add_ac(1) assms deg_mult_ring le_iff_add mult_Suc)  

lemma(in UP_cring) deg_smult:
  assumes "a  carrier R"
  assumes "f  carrier (UP R)"
  assumes "a  lcf f  𝟬"
  shows "deg R (a UP Rf) = deg R  f"
  using assms  P_def cfs_smult deg_eqI deg_smult_decr smult_closed 
  by (metis deg_gtE le_neq_implies_less)
 
lemma(in UP_cring) deg_smult':
  assumes "a  Units R"
  assumes "f  carrier (UP R)"
  shows "deg R (a UP Rf) = deg R  f"
  apply(cases "deg R f = 0")
  apply (metis P_def R.Units_closed assms(1) assms(2) deg_smult_decr le_zero_eq)
  apply(rule deg_smult) 
  using assms apply blast
  using assms apply blast
proof
  assume A: "deg R f  0" "a  f (deg R f) = 𝟬"
  have 0: "f (deg R f) = 𝟬"
    using A assms R.Units_not_right_zero_divisor[of a "f (deg R f)"] UP_car_memE(1) by blast
  then show False  using assms A 
    by (metis P_def deg_zero deg_ltrm monom_zero)
qed

lemma(in UP_domain)  pow_sum0:
" p q. p  carrier P  q  carrier P  degree q < degree p  degree ((p Pq )[^]Pn) = (degree p)*n"
proof(induction n)
  case 0
  then show ?case 
    by (metis Group.nat_pow_0 deg_one mult_is_0)
next
    case (Suc n)
    fix n
    assume IH: " p q. p  carrier P  q  carrier P  
              degree q < degree p  degree ((p Pq )[^]Pn) = (degree p)*n"
    then show " p q. p  carrier P  q  carrier P  
             degree q < degree p  degree ((p Pq )[^]P(Suc n)) = (degree p)*(Suc n)"
    proof-
      fix p q
      assume A0: "p  carrier P" and 
           A1: "q  carrier P" and 
           A2:  "degree q < degree p"
      show "degree ((p Pq )[^]P(Suc n)) = (degree p)*(Suc n)"
      proof(cases "q = 𝟬P⇙")
        case True
        then show ?thesis 
        by (metis A0 A1 A2 IH P.nat_pow_Suc2 P.nat_pow_closed P.r_zero deg_mult 
            domain.nonzero_pow_nonzero local.domain_axioms mult_Suc_right nat_neq_iff)
      next
        case False
        then show ?thesis 
        proof-
          have P0: "degree ((p Pq )[^]Pn) = (degree p)*n" 
          using A0 A1 A2 IH by auto 
          have P1: "(p Pq )[^]P(Suc n) = ((p Pq )[^]Pn) P(p Pq )"
            by simp
          then have P2: "(p Pq )[^]P(Suc n) = (((p Pq )[^]Pn) Pp) P(((p Pq )[^]Pn) Pq)"
            by (simp add: A0 A1 UP_r_distr)
          have P3: "degree (((p Pq )[^]Pn) Pp) = (degree p)*n + (degree p)" 
            using P0 A0 A1 A2 deg_nzero_nzero  degree_of_sum_diff_degree local.nonzero_pow_nonzero by auto
          have P4: "degree (((p Pq )[^]Pn) Pq) = (degree p)*n + (degree q)" 
            using P0 A0 A1 A2 deg_nzero_nzero  degree_of_sum_diff_degree local.nonzero_pow_nonzero False deg_mult 
            by simp
          have P5: "degree (((p Pq )[^]Pn) Pp) > degree (((p Pq )[^]Pn) Pq)"
            using P3 P4 A2 by auto 
          then show ?thesis using P5 P3 P2 
            by (simp add: A0 A1 degree_of_sum_diff_degree)
        qed
      qed
    qed
qed

lemma(in UP_domain)  pow_sum:
  assumes "p  carrier P" 
  assumes "q  carrier P"
  assumes "degree q < degree p"
  shows "degree ((p Pq )[^]Pn) = (degree p)*n"
  using assms(1) assms(2) assms(3) pow_sum0 by blast

lemma(in UP_domain)  deg_pow0:
 " p. p  carrier P  n  degree p  degree (p [^]Pm) = m*(degree p)"
proof(induction n)
  case 0
  show "p  carrier P  0  degree p  degree (p [^]Pm) = m*(degree p)"
  proof-
    assume B0:"p  carrier P"
    assume B1: "0  degree p"
    then obtain a where a_def: "a  carrier R  p = monom P a 0"
      using B0 deg_zero_impl_monom  by fastforce
    show "degree (p [^]Pm) = m*(degree p)"  using UP_cring.monom_pow 
      by (metis P_def R.nat_pow_closed UP_cring_axioms a_def deg_const  
        mult_0_right mult_zero_left)
  qed
next
  case (Suc n)
  fix n
  assume IH: "p. (p  carrier P  n degree p  degree (p [^]Pm) = m * (degree p))"
  show "p  carrier P  Suc n  degree p  degree (p [^]Pm) = m * (degree p)"
  proof-
    assume A0: "p  carrier P"
    assume A1: "Suc n  degree p"
    show "degree (p [^]Pm) = m * (degree p)"
    proof(cases "Suc n > degree p")
      case True
      then show ?thesis using IH A0 by simp
    next
      case False
      then show ?thesis 
      proof-
        obtain q where q_def: "q = trunc p"
          by simp
        obtain k where k_def: "k = degree q"
          by simp
        have q_is_poly: "q  carrier P" 
          by (simp add: A0 q_def trunc_closed)
        have k_bound0: "k <degree p" 
          using k_def q_def trunc_degree[of p] A0 False by auto
        have k_bound1: "k  n" 
          using k_bound0 A0 A1 by auto  
        have P_q:"degree (q [^]Pm) = m * k" 
          using IH[of "q"] k_bound1 k_def q_is_poly by auto  
        have P_ltrm: "degree ((ltrm p) [^]Pm) = m*(degree p)"
        proof-
          have "degree p = degree (ltrm p)" 
            by (simp add: A0 deg_ltrm)            
          then show ?thesis using monom_degree 
            by (metis A0 P.r_zero P_def cfs_closed coeff_simp equal_deg_sum k_bound0 k_def lcoeff_nonzero2 nat_neq_iff q_is_poly)            
        qed
        have "p = q P(ltrm p)" 
          by (simp add: A0 q_def trunc_simps(1))
        then show ?thesis 
          using P_q pow_sum[of "(ltrm p)" q m] A0 UP_a_comm 
            deg_ltrm k_bound0 k_def ltrm_closed q_is_poly by auto
      qed
    qed
  qed
qed

lemma(in UP_domain)  deg_pow:
  assumes "p  carrier P"
  shows "degree (p [^]Pm) = m*(degree p)"
  using deg_pow0 assms by blast

lemma(in UP_domain)  ltrm_pow0:
"f. f  carrier P  ltrm (f [^]P(n::nat)) = (ltrm f) [^]Pn"
proof(induction n)
  case 0
  then show ?case 
    using ltrm_deg_0 P.nat_pow_0 P.ring_simprules(6) deg_one by presburger  
next
  case (Suc n)
  fix n::nat
  assume IH: "f. f  carrier P  ltrm (f [^]Pn) = (ltrm f) [^]Pn"
  then show "f. f  carrier P  ltrm (f [^]P(Suc n)) = (ltrm f) [^]P(Suc n)"
  proof-
    fix f
    assume A: "f  carrier P"
    show " ltrm (f [^]P(Suc n)) = (ltrm f) [^]P(Suc n)"
    proof-
      have 0: "ltrm (f [^]Pn) = (ltrm f) [^]Pn" 
        using A IH  by blast
      have 1: "ltrm (f [^]P(Suc n)) = ltrm ((f [^]Pn)Pf)" 
        by auto then 
      show ?thesis using ltrm_mult 0 1 
        by (simp add: A)
    qed
  qed
qed

lemma(in UP_domain)  ltrm_pow:
  assumes "f  carrier P"
  shows " ltrm (f [^]P(n::nat)) = (ltrm f) [^]Pn"
  using assms ltrm_pow0 by blast

text‹lemma on the leading coefficient›

lemma lcf_eq:
  assumes "f  carrier P"
  shows "lcf f = lcf (ltrm f)"
  using ltrm_deg_0 
  by (simp add: ltrm_cfs assms deg_ltrm)
  
lemma lcf_eq_deg_eq_imp_ltrm_eq:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree p > 0"
  assumes "degree p = degree q"
  assumes "lcf p = lcf q"
  shows "ltrm p = ltrm q"
  using assms(4) assms(5) 
  by (simp add: leading_term_def)
  
lemma ltrm_eq_imp_lcf_eq:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "ltrm p = ltrm q"
  shows "lcf p = lcf q"
  using assms 
  by (metis lcf_eq)

lemma ltrm_eq_imp_deg_drop:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "ltrm p = ltrm q"
  assumes "degree p >0"
  shows "degree (p Pq) < degree p"
proof-
  have P0: "degree p = degree q"
    by (metis assms(1) assms(2) assms(3) deg_ltrm)
  then have P1: "degree (p Pq)  degree p"
    by (metis P.add.inv_solve_right P.minus_closed P.minus_eq assms(1)
        assms(2) degree_of_sum_diff_degree neqE order.strict_implies_order order_refl)
  have "degree (p Pq)  degree p"
  proof
    assume A: "degree (p Pq) = degree p"
    have Q0: "p Pq = ((trunc p) P(ltrm p)) P((trunc q) P(ltrm p))"
      using assms(1) assms(2) assms(3) trunc_simps(1) by force
    have Q1: "p Pq = (trunc p)  P(trunc q)" 
    proof-
      have "p Pq = ((trunc p) P(ltrm p)) P(trunc q) P(ltrm p)"
        using Q0 
        by (simp add: P.minus_add P.minus_eq UP_a_assoc assms(1) assms(2) ltrm_closed trunc_closed)
      then show ?thesis 
        by (metis (no_types, lifting) ltrm_closed P.add.inv_mult_group P.minus_eq 
            P.r_neg2 UP_a_assoc assms(1) assms(2) assms(3) carrier_is_submodule submoduleE(6) trunc_closed trunc_simps(1))        
    qed
    have Q2: "degree (trunc p) < degree p" 
      by (simp add: assms(1) assms(4) trunc_degree)
    have Q3: "degree (trunc q) < degree q" 
      using P0 assms(2) assms(4) trunc_degree by auto
    then show False  using A Q1 Q2 Q3 by (simp add: P.add.inv_solve_right
          P.minus_eq P0 assms(1) assms(2) degree_of_sum_diff_degree trunc_closed)
  qed
  then show ?thesis 
    using P1 by auto
qed

lemma(in UP_cring) cring_lcf_scalar_mult:
  assumes "p  carrier P"
  assumes "a  carrier R"
  assumes "a  (lcf p) 𝟬"
  shows "lcf (a Pp) = a  (lcf p)"
proof-
  have 0: "lcf (a Pp) = lcf (ltrm (a Pp))"
    using assms  lcf_eq smult_closed by blast
  have 1: "degree (a Pp) = degree p"
    by (smt lcf_monom(1) P_def R.one_closed R.r_null UP_ring.coeff_smult UP_ring_axioms  
        assms(1) assms(2) assms(3) coeff_simp cring_deg_mult deg_const monom_closed monom_mult_is_smult smult_one)
  then have "lcf (a Pp) = lcf (a P(ltrm p))"
    using lcf_eq[of "a Pp"] smult_closed  assms 0 
    by (metis cfs_closed cfs_smult monom_mult_smult)     
  then show ?thesis 
    unfolding leading_term_def 
    by (metis P_def R.m_closed UP_cring.lcf_monom UP_cring_axioms assms(1) assms(2) cfs_closed monom_mult_smult)   
qed

lemma(in UP_domain) lcf_scalar_mult:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "lcf (a Pp) = a  (lcf p)"
proof-
  have "lcf (a Pp) = lcf (ltrm (a Pp))"
    using lcf_eq UP_smult_closed assms(1) assms(2) by blast 
  then have "lcf (a Pp) = lcf (a P(ltrm p))"
    using ltrm_smult assms(1) assms(2) by metis   
  then show ?thesis 
    by (metis (full_types) UP_smult_zero assms(1) assms(2) cfs_smult cfs_zero deg_smult)    
qed

lemma(in UP_cring)  cring_lcf_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "(lcf p)  (lcf q) 𝟬"
  shows "lcf (p Pq) = (lcf p)  (lcf q)"
  using assms cring_ltrm_mult 
  by (smt lcf_monom(1) P.m_closed R.m_closed cfs_closed monom_mult)

lemma(in UP_domain)  lcf_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "lcf (p Pq) = (lcf p)  (lcf q)"
  by (smt ltrm_deg_0 R.integral_iff assms(1) assms(2) cfs_closed cring_lcf_mult deg_zero deg_ltrm local.integral_iff monom_zero)

lemma(in UP_cring)  cring_lcf_pow:
  assumes "p  carrier P"
  assumes "(lcf p)[^]n 𝟬"
  shows "lcf (p[^]P(n::nat)) = (lcf p)[^]n"
  by (smt P.nat_pow_closed R.nat_pow_closed assms(1) assms(2) cring_pow_ltrm lcf_closed lcf_ltrm lcf_monom monom_pow)

lemma(in UP_domain)  lcf_pow:
  assumes "p  carrier P"
  shows "lcf (p[^]P(n::nat)) = (lcf p)[^]n"
proof-
  show ?thesis 
  proof(induction n)
    case 0
    then show ?case 
      by (metis Group.nat_pow_0 P_def R.one_closed UP_cring.lcf_monom UP_cring_axioms monom_one)    
  next
    case (Suc n)
    fix n
    assume IH: "lcf (p[^]P(n::nat)) = (lcf p)[^]n"
    show "lcf (p[^]P(Suc n)) = (lcf p)[^](Suc n)"
    proof-
      have "lcf (p[^]P(Suc n)) = lcf ((p[^]Pn) Pp)"
        by simp
      then have "lcf (p[^]P(Suc n)) = (lcf p)[^]n  (lcf p)"
        by (simp add: IH assms lcf_mult)
      then show ?thesis by auto 
    qed
  qed
qed
end

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Constant Terms and Constant Coefficients›
  (**************************************************************************************************)
  (**************************************************************************************************)

text‹Constant term  and coefficient function›

definition zcf where
"zcf f = (f 0)"

abbreviation(in UP_cring)(input) ctrm where
"ctrm f  monom P (f 0) 0"

context UP_cring
begin

lemma ctrm_is_poly:
  assumes "p  carrier P"
  shows "ctrm p  carrier P"
  by (simp add: assms cfs_closed)

lemma ctrm_degree:
  assumes "p  carrier P"
  shows "degree (ctrm p) = 0"
  by (simp add: assms cfs_closed) 
  
lemma ctrm_zcf:
assumes "f  carrier P"
assumes "zcf f = 𝟬"
shows "ctrm f = 𝟬P⇙"
  by (metis P_def UP_ring.monom_zero UP_ring_axioms zcf_def assms(2))

lemma zcf_degree_zero:
  assumes "f  carrier P"
  assumes "degree f = 0"
  shows "lcf f = zcf f"
  by (simp add: zcf_def assms(2))
  
lemma zcf_zero_degree_zero:
  assumes "f  carrier P"
  assumes "degree f = 0"
  assumes "zcf f = 𝟬"
  shows "f = 𝟬P⇙"
  using zcf_degree_zero[of f] assms ltrm_deg_0[of f] 
  by simp

lemma zcf_ctrm:
  assumes "p  carrier P"
  shows "zcf (ctrm p) = zcf p"
  unfolding zcf_def 
  using P_def UP_ring.cfs_monom UP_ring_axioms assms cfs_closed by fastforce
 
lemma ctrm_trunc:
  assumes "p  carrier P"
  assumes "degree p >0"
  shows "zcf(trunc p) = zcf p"
  by (simp add: zcf_def assms(1) assms(2) trunc_cfs)

text‹Constant coefficient function is a ring homomorphism›


lemma zcf_add:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "zcf(p Pq) = (zcf p)  (zcf q)"
  by (simp add: zcf_def assms(1) assms(2))

lemma coeff_ltrm[simp]:
  assumes "p  carrier P"
  assumes "degree p > 0"
  shows "zcf(ltrm p) = 𝟬"
  by (metis ltrm_cfs_above_deg ltrm_cfs zcf_def assms(1) assms(2))

lemma zcf_zero[simp]:
"zcf 𝟬P= 𝟬"
  using zcf_degree_zero by auto

lemma zcf_one[simp]:
"zcf 𝟭P= 𝟭"
  by (simp add: zcf_def)

lemma ctrm_smult:
  assumes "f  carrier P"
  assumes "a   carrier R"
  shows "ctrm (a Pf) = a P(ctrm f)"
  using P_def UP_ring.monom_mult_smult UP_ring_axioms assms(1) assms(2) cfs_smult coeff_simp 
  by (simp add: UP_ring.monom_mult_smult cfs_closed)

lemma ctrm_monom[simp]:
  assumes "a  carrier R"
  shows "ctrm (monom P a (Suc k)) = 𝟬P⇙"
  by (simp add: assms cfs_monom)
end
  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹Polynomial Induction Rules›
  (**************************************************************************************************)
  (**************************************************************************************************)

context UP_ring
begin

text‹Rule for strong induction on polynomial degree›

lemma poly_induct:
  assumes "p  carrier P"
  assumes Deg_0: "p. p  carrier P  degree p = 0  Q p"
  assumes IH: "p. (q. q  carrier P  degree q < degree p  Q q)  p  carrier P  degree p > 0  Q p"
  shows "Q p"
proof-
  have "n. p. p  carrier P  degree p  n  Q p"
  proof-
    fix n
    show "p. p  carrier P   degree p  n  Q p"
    proof(induction n)
      case 0
      then show ?case 
        using Deg_0  by simp
    next
      case (Suc n)
      fix n 
      assume I:  "p. p  carrier P   degree p  n  Q p"
      show  "p. p  carrier P   degree p  (Suc n)  Q p"
      proof-
        fix p
        assume A0: " p  carrier P "
        assume A1: "degree p Suc n"
        show "Q p"
        proof(cases "degree p < Suc n")
          case True
          then show ?thesis 
            using I  A0 by auto
        next
          case False
          then have D: "degree p = Suc n" 
            by (simp add: A1 nat_less_le)
          then  have "(q. q  carrier P  degree q < degree p  Q q)" 
              using I   by simp
            then show "Q p" 
                  using IH D A0 A1 Deg_0 by blast
        qed
      qed
    qed
  qed
  then show ?thesis using assms by blast 
qed

text‹Variant on induction on degree›

lemma poly_induct2:
  assumes "p  carrier P"
  assumes Deg_0: "p. p  carrier P  degree p = 0  Q p"
  assumes IH: "p. degree p > 0   p  carrier P  Q (trunc p)   Q p"
  shows "Q p"
proof(rule poly_induct)
  show "p  carrier P" 
    by (simp add: assms(1))
  show "p. p  carrier P  degree p = 0  Q p" 
      by (simp add: Deg_0)
    show "p. (q. q  carrier P  degree q < degree p  Q q)  p  carrier P  0 < degree p  Q p"
    proof-
      fix p
      assume A0: "(q. q  carrier P  degree q < degree p  Q q)"
      assume A1: " p  carrier P"
      assume A2: "0 < degree p" 
      show "Q p"
      proof-
        have "degree (trunc p) < degree p"
          by (simp add: A1 A2 trunc_degree)
        have "Q (trunc p)"
          by (simp add: A0 A1 degree (trunc p) < degree p trunc_closed)
        then show ?thesis 
          by (simp add: A1 A2 IH)
      qed
    qed
qed

text‹Additive properties which are true for all monomials are true for all polynomials ›

lemma poly_induct3:
  assumes "p  carrier P"
  assumes add: "p q. q  carrier P  p  carrier P  Q p  Q q   Q (p Pq)"
  assumes monom: "a n. a  carrier R  Q (monom P a n)"
  shows "Q p"
  apply(rule poly_induct2)
  apply (simp add: assms(1))
  apply (metis lcf_closed P_def coeff_simp deg_zero_impl_monom monom)
  by (metis lcf_closed ltrm_closed add monom trunc_closed trunc_simps(1))

lemma poly_induct4:
  assumes "p  carrier P"
  assumes add: "p q. q  carrier P  p  carrier P  Q p  Q q   Q (p Pq)"
  assumes monom_zero: "a. a  carrier R  Q (monom P a 0)"
  assumes monom_Suc: "a n. a  carrier R  Q (monom P a (Suc n))"
  shows "Q p"
  apply(rule poly_induct3)
  using assms(1) apply auto[1]
  using add apply blast
  using monom_zero monom_Suc 
  by (metis P_def UP_ring.monom_zero UP_ring_axioms deg_monom deg_monom_le le_0_eq le_SucE zero_induct)

lemma monic_monom_smult:
  assumes "a  carrier R"
  shows "a Pmonom P 𝟭 n = monom P a n"
  using assms 
  by (metis R.one_closed R.r_one monom_mult_smult)

lemma poly_induct5:
  assumes "p  carrier P"
  assumes add: "p q. q  carrier P  p  carrier P  Q p  Q q   Q (p Pq)"
  assumes monic_monom: "n. Q (monom P 𝟭 n)"
  assumes smult: "p a . a  carrier R  p  carrier P  Q p  Q (a Pp)"
  shows "Q p"
  apply(rule poly_induct3)
  apply (simp add: assms(1))
  using add apply blast
proof-
  fix a n assume A: "a  carrier R" show "Q (monom P a n)"
    using monic_monom[of n] smult[of a "monom P 𝟭 n"] monom_mult_smult[of a 𝟭 n]
  by (simp add: A)
qed

lemma poly_induct6:
  assumes "p  carrier P"
  assumes monom: "a n. a  carrier R  Q (monom P a 0)"
  assumes plus_monom: "a n p. a  carrier R  a  𝟬  p  carrier P  degree p < n  Q p 
                                Q(p Pmonom P a n)"
  shows "Q p"
  apply(rule poly_induct2)
  using assms(1) apply auto[1]
  apply (metis lcf_closed P_def coeff_simp deg_zero_impl_monom monom)
  using plus_monom 
  by (metis lcf_closed P_def coeff_simp lcoeff_nonzero_deg nat_less_le trunc_closed trunc_degree trunc_simps(1))
  

end

(**************************************************************************************************)
(**************************************************************************************************)
section‹Mapping a Polynomial to its Associated Ring Function›
(**************************************************************************************************)
(**************************************************************************************************)


text‹Turning a polynomial into a function on R:›
definition to_function  where
"to_function S f  = (λs  carrier S. eval S S (λx. x) s f)"

context UP_cring
begin

definition to_fun where
"to_fun f  to_function R f"

text‹Explicit formula for evaluating a polynomial function:›

lemma to_fun_eval:
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun f x  = eval R R (λx. x) x f"
  using assms unfolding to_function_def to_fun_def 
  by auto 

lemma to_fun_formula:
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun f x = (i  {..degree f}. (f i)  x [^] i)"
proof-
  have "f  carrier (UP R)"
    using assms P_def by auto 
  then  have "eval R R (λx. x) x f =  (Ri{..deg R f}. (λx. x) (coeff (UP R) f i) Rx [^]Ri)"
    apply(simp add:UnivPoly.eval_def) done
  then have "to_fun f x = (Ri{..deg R f}. (λx. x) (coeff (UP R) f i) Rx [^]Ri)"
    using to_function_def assms unfolding to_fun_def  
    by (simp add: to_function_def)
  then show ?thesis  
    by(simp add: assms coeff_simp) 
qed

lemma eval_ring_hom:
  assumes "a  carrier R"
  shows "eval R R (λx. x) a  ring_hom P R"
proof-
  have "(λx. x)  ring_hom R R"
    apply(rule ring_hom_memI)
    apply auto done
  then have "UP_pre_univ_prop R R (λx. x)"
    using R_cring UP_pre_univ_propI by blast
  then show ?thesis 
    by (simp add: P_def UP_pre_univ_prop.eval_ring_hom assms)
qed

lemma to_fun_closed:
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun f x   carrier R"
  using assms to_fun_eval[of f x] eval_ring_hom[of x]
  ring_hom_closed 
  by fastforce
  
lemma to_fun_plus:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun (f Pg) x = (to_fun f x)  (to_fun g x)"
  using assms to_fun_eval[of ] eval_ring_hom[of x]
  by (simp add: ring_hom_add)
  
lemma to_fun_mult:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun (f Pg) x = (to_fun f x)  (to_fun g x)"
  using assms to_fun_eval[of ] eval_ring_hom[of x]
  by (simp add: ring_hom_mult)

lemma to_fun_ring_hom:
  assumes "a  carrier R"
  shows "(λp. to_fun p a)  ring_hom P R"
  apply(rule ring_hom_memI)
  apply (simp add: assms to_fun_closed)
    apply (simp add: assms to_fun_mult)
      apply (simp add: assms to_fun_plus)
        using to_fun_eval[of "𝟭P⇙" a] eval_ring_hom[of a] 
              ring_hom_closed 
        by (simp add: assms ring_hom_one)

lemma ring_hom_uminus:
  assumes "ring S"
  assumes "f  (ring_hom S R)"
  assumes "a  carrier S"
  shows "f (Sa) =  (f a)"
proof-
  have "f (a Sa) = (f a)  f (Sa)"
    unfolding a_minus_def 
    by (simp add: assms(1) assms(2) assms(3) ring.ring_simprules(3) ring_hom_add)
  then have "(f a)  f (Sa) = 𝟬 "
    by (metis R.ring_axioms a_minus_def assms(1) assms(2) assms(3) 
        ring.ring_simprules(16) ring_hom_zero)
  then show ?thesis 
    by (metis (no_types, lifting) R.add.m_comm R.minus_equality assms(1)
        assms(2) assms(3) ring.ring_simprules(3) ring_hom_closed)
qed

lemma to_fun_minus:
  assumes "f  carrier P"
  assumes "x  carrier R"
  shows "to_fun (Pf) x =  (to_fun f x)"
  unfolding to_function_def to_fun_def 
  using  eval_ring_hom[of x] assms  
  by (simp add: UP_ring ring_hom_uminus)

lemma id_is_hom:
"ring_hom_cring R R (λx. x)"
proof(rule ring_hom_cringI)
  show "cring R" 
    by (simp add: R_cring )
  show "cring R" 
    by (simp add: R_cring ) 
  show "(λx. x)  ring_hom R R"
    unfolding ring_hom_def
    apply(auto)
    done
qed

lemma UP_pre_univ_prop_fact:
"UP_pre_univ_prop R R (λx. x)"
  unfolding UP_pre_univ_prop_def
  by (simp add: UP_cring_def R_cring  id_is_hom)

end

  (**************************************************************************************************)
  (**************************************************************************************************)
  subsection‹to-fun is a Ring Homomorphism from Polynomials to Functions›
  (**************************************************************************************************)
  (**************************************************************************************************)

context UP_cring
begin

lemma to_fun_is_Fun:
  assumes "x  carrier P"
  shows "to_fun x  carrier (Fun R)"
  apply(rule ring_functions.function_ring_car_memI)
  unfolding ring_functions_def apply(simp add: R.ring_axioms)
  using to_fun_closed assms apply auto[1]
  unfolding to_function_def to_fun_def  by auto 

lemma to_fun_Fun_mult:
  assumes "x  carrier P"
  assumes "y  carrier P"
  shows "to_fun (x Py) = to_fun x function_ring (carrier R) Rto_fun y"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  apply (simp add: R.ring_axioms ring_functions_def)
    apply (simp add: assms(1) assms(2) to_fun_is_Fun)
      apply (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_mult_closed ring_functions.intro to_fun_is_Fun)
        by (simp add: R.ring_axioms assms(1) assms(2) ring_functions.function_mult_eval_car ring_functions.intro to_fun_is_Fun to_fun_mult)

lemma to_fun_Fun_add:
  assumes "x  carrier P"
  assumes "y  carrier P"
  shows "to_fun (x Py) = to_fun x function_ring (carrier R) Rto_fun y"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  apply (simp add: R.ring_axioms ring_functions_def)
    apply (simp add: assms(1) assms(2) to_fun_is_Fun)
      apply (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_add_closed ring_functions.intro to_fun_is_Fun)
        by (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_add_eval_car ring_functions.intro to_fun_is_Fun to_fun_plus)

lemma to_fun_Fun_one:
"to_fun 𝟭P= 𝟭Fun R⇙"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  apply (simp add: R.ring_axioms ring_functions_def)
    apply (simp add: to_fun_is_Fun)
      apply (simp add: R.ring_axioms ring_functions.function_one_closed ring_functions_def)
        using P_def R.ring_axioms UP_cring.eval_ring_hom UP_cring.to_fun_eval UP_cring_axioms UP_one_closed ring_functions.function_one_eval ring_functions.intro ring_hom_one 
          by fastforce

lemma to_fun_Fun_zero:
"to_fun 𝟬P= 𝟬Fun R⇙"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  apply (simp add: R.ring_axioms ring_functions_def)
    apply (simp add: to_fun_is_Fun)
      apply (simp add: R.ring_axioms ring_functions.function_zero_closed ring_functions_def)
        using P_def R.ring_axioms UP_cring.eval_ring_hom UP_cring.to_fun_eval UP_cring_axioms UP_zero_closed ring_functions.function_zero_eval ring_functions.intro ring_hom_zero 
        by (metis UP_ring eval_ring_hom)

lemma to_fun_function_ring_hom:
"to_fun  ring_hom P (Fun R)"
  apply(rule ring_hom_memI)
    using to_fun_is_Fun apply auto[1]
      apply (simp add: to_fun_Fun_mult)
        apply (simp add: to_fun_Fun_add)
          by (simp add: to_fun_Fun_one)

lemma(in UP_cring) to_fun_one:
  assumes "a  carrier R"
  shows "to_fun 𝟭Pa = 𝟭"
  using assms to_fun_Fun_one 
  by (metis P_def UP_cring.to_fun_eval UP_cring_axioms UP_one_closed eval_ring_hom ring_hom_one) 

lemma(in UP_cring) to_fun_zero:
  assumes "a  carrier R"
  shows "to_fun 𝟬Pa = 𝟬"
  by (simp add: assms R.ring_axioms ring_functions.function_zero_eval ring_functions.intro to_fun_Fun_zero) 

lemma(in UP_cring) to_fun_nat_pow:
  assumes "h  carrier (UP R)"
  assumes "a  carrier R"
  shows "to_fun (h[^]UP R(n::nat)) a = (to_fun h a)[^]n"
  apply(induction n)
  using assms to_fun_one 
  apply (metis P.nat_pow_0 P_def R.nat_pow_0)
  using assms to_fun_mult  P.nat_pow_closed P_def by auto

lemma(in UP_cring) to_fun_finsum: 
  assumes "finite (Y::'d set)"
  assumes "f  UNIV  carrier (UP R)"
  assumes "t  carrier R"
  shows "to_fun (finsum (UP R) f Y) t = finsum R (λi. (to_fun (f i) t)) Y"
proof(rule finite.induct[of Y])
  show "finite Y"
    using assms by blast 
  show "to_fun (finsum (UP R) f {}) t = (i{}. to_fun (f i) t)"
    using P.finsum_empty[of f] assms unfolding P_def R.finsum_empty 
    using P_def to_fun_zero by presburger
  show  "A a. finite A 
           to_fun (finsum (UP R) f A) t = (iA. to_fun (f i) t)  to_fun (finsum (UP R) f (insert a A)) t = (iinsert a A. to_fun (f i) t)"
  proof-     
    fix A :: "'d set" fix a
    assume A: "finite A" "to_fun (finsum (UP R) f A) t = (iA. to_fun (f i) t)"
    show "to_fun (finsum (UP R) f (insert a A)) t = (iinsert a A. to_fun (f i) t)"
    proof(cases "a  A")
      case True
      then show ?thesis using A 
        by (metis insert_absorb)
    next
      case False
      have 0: "finsum (UP R) f (insert a A) = f a UP Rfinsum (UP R) f A"
        using A False finsum_insert[of A a f] assms unfolding P_def  by blast
      have 1: "to_fun (f a Pfinsum (UP R) f A ) t =  to_fun (f a) t  to_fun (finsum (UP R) f A) t"
        apply(rule to_fun_plus[of  "finsum (UP R) f A" "f a" t])
        using assms(2) finsum_closed[of f A] A unfolding P_def apply blast
        using P_def assms apply blast
        using assms by blast
      have 2: "to_fun (f a Pfinsum (UP R) f A ) t =  to_fun (f a) t  (iA. to_fun (f i) t)"
        unfolding  1 A by blast 
      have 3: "(iinsert a A. to_fun (f i) t) = to_fun (f a) t  (iA. to_fun (f i) t)"
        apply(rule  R.finsum_insert, rule A, rule False)
        using to_fun_closed assms unfolding P_def apply blast
        apply(rule to_fun_closed) using assms unfolding P_def apply blast using assms by blast 
      show ?thesis 
        unfolding 0 unfolding 3 using 2 unfolding P_def by blast 
    qed
  qed
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Inclusion of a Ring into its Polynomials Ring via Constants›
(**************************************************************************************************)
(**************************************************************************************************)

definition to_polynomial where
"to_polynomial R = (λa. monom (UP R) a 0)"

context UP_cring
begin

abbreviation(input) to_poly where
"to_poly   to_polynomial R"

lemma to_poly_mult_simp:
  assumes "b  carrier R"
  assumes "f  carrier (UP R)"
  shows "(to_polynomial R b) UP Rf = b UP Rf"
        "f  UP R(to_polynomial R b) = b UP Rf"   
  unfolding to_polynomial_def
  using assms  P_def monom_mult_is_smult apply auto[1]
  using UP_cring.UP_m_comm UP_cring_axioms UP_ring.monom_closed 
        UP_ring.monom_mult_is_smult UP_ring_axioms assms(1) assms(2) 
  by fastforce  

lemma to_fun_to_poly:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_fun (to_poly a) b = a"
  unfolding to_function_def to_fun_def to_polynomial_def 
  by (simp add: UP_pre_univ_prop.eval_const UP_pre_univ_prop_fact assms(1) assms(2))
  
lemma to_poly_inverse:
  assumes "f  carrier P"
  assumes "degree f = 0"
  shows "f = to_poly (f 0)"
  using P_def assms(1) assms(2) 
  by (metis ltrm_deg_0 to_polynomial_def) 

lemma to_poly_closed:
  assumes "a  carrier R"
  shows "to_poly a  carrier P"
  by (metis P_def assms monom_closed to_polynomial_def)

lemma degree_to_poly[simp]:
  assumes "a  carrier R"
  shows "degree (to_poly a) = 0"
  by (metis P_def assms deg_const to_polynomial_def)

lemma to_poly_is_ring_hom:
"to_poly  ring_hom R P"
  unfolding to_polynomial_def
  unfolding P_def
  using UP_ring.const_ring_hom[of R]
  UP_ring_axioms by simp 

lemma to_poly_add:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_poly (a  b) = to_poly a Pto_poly b"
  by (simp add: assms(1) assms(2) ring_hom_add to_poly_is_ring_hom)

lemma to_poly_mult:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_poly (a  b) = to_poly a Pto_poly b"
  by (simp add: assms(1) assms(2) ring_hom_mult to_poly_is_ring_hom)

lemma to_poly_minus:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_poly (a  b) = to_poly a Pto_poly b"
  by (metis P.minus_eq P_def R.add.inv_closed R.ring_axioms UP_ring.monom_add 
      UP_ring_axioms assms(1) assms(2) monom_a_inv ring.ring_simprules(14) to_polynomial_def)

lemma to_poly_a_inv:
  assumes "a  carrier R"
  shows "to_poly ( a) =  Pto_poly a"
  by (metis P_def assms monom_a_inv to_polynomial_def)

lemma to_poly_nat_pow:
  assumes "a  carrier R"
  shows "(to_poly a) [^]P(n::nat)= to_poly (a[^]n)"
  using assms UP_cring UP_cring_axioms UP_cring_def UnivPoly.ring_hom_cringI ring_hom_cring.hom_pow to_poly_is_ring_hom 
  by fastforce


end

(**************************************************************************************************)
(**************************************************************************************************)
section‹Polynomial Substitution›
(**************************************************************************************************)
(**************************************************************************************************)

definition compose where
"compose R f g = eval R (UP R) (to_polynomial R) g f"

abbreviation(in UP_cring)(input) sub  (infixl "of" 70) where
"sub f g  compose R f g"

definition rev_compose  where
"rev_compose R = eval R (UP R) (to_polynomial R)"

abbreviation(in UP_cring)(input) rev_sub  where
"rev_sub  rev_compose R"

context UP_cring
begin

lemma  sub_rev_sub:
"sub f g = rev_sub g f"
  unfolding compose_def rev_compose_def 
  by simp  

lemma(in UP_cring) to_poly_UP_pre_univ_prop:
"UP_pre_univ_prop R P to_poly"
proof 
  show "to_poly  ring_hom R P" 
    by (simp add: to_poly_is_ring_hom)
qed

lemma rev_sub_is_hom:
  assumes "g  carrier P"
  shows "rev_sub g  ring_hom P P"
  unfolding rev_compose_def
  using to_poly_UP_pre_univ_prop assms(1) UP_pre_univ_prop.eval_ring_hom[of R P to_poly g] 
  unfolding P_def apply auto 
  done

lemma rev_sub_closed:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "rev_sub q p  carrier P"
  using rev_sub_is_hom[of q] assms ring_hom_closed[of "rev_sub q" P P p] by auto  

lemma sub_closed:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "sub q p  carrier P"
  by (simp add: assms(1) assms(2) rev_sub_closed sub_rev_sub)

lemma rev_sub_add:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "h carrier P"
  shows "rev_sub g (f Ph) = (rev_sub g f) P(rev_sub g h)"
  using rev_sub_is_hom assms ring_hom_add by fastforce

lemma sub_add: 
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "h carrier P"
  shows "((f Ph) of g) = ((f of g) P(h of g))"
  by (simp add: assms(1) assms(2) assms(3) rev_sub_add sub_rev_sub)

lemma rev_sub_mult:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "h carrier P"
  shows "rev_sub g (f Ph) = (rev_sub g f) P(rev_sub g h)"
  using rev_sub_is_hom assms ring_hom_mult  by fastforce

lemma sub_mult: 
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "h carrier P"
  shows "((f Ph) of g) = ((f of g) P(h of g))"
  by (simp add: assms(1) assms(2) assms(3) rev_sub_mult sub_rev_sub)

lemma sub_monom:
  assumes "g  carrier (UP R)"
  assumes "a  carrier R"
  shows "sub (monom (UP R) a n) g = to_poly a  UP R(g[^]UP R(n::nat))" 
        "sub (monom (UP R) a n) g = a UP R(g[^]UP R(n::nat))" 
   apply (simp add: UP_cring.to_poly_UP_pre_univ_prop UP_cring_axioms 
          UP_pre_univ_prop.eval_monom assms(1) assms(2) Cring_Poly.compose_def)
  by (metis P_def UP_cring.to_poly_mult_simp(1) UP_cring_axioms UP_pre_univ_prop.eval_monom 
      UP_ring assms(1) assms(2) Cring_Poly.compose_def monoid.nat_pow_closed ring_def to_poly_UP_pre_univ_prop) 
 
text‹Subbing into a constant does nothing›

lemma rev_sub_to_poly:
  assumes "g  carrier P"
  assumes "a  carrier R"
  shows "rev_sub g (to_poly a) = to_poly a"
  unfolding to_polynomial_def rev_compose_def
  using to_poly_UP_pre_univ_prop 
  unfolding to_polynomial_def 
     using P_def UP_pre_univ_prop.eval_const assms(1) assms(2) by fastforce

lemma sub_to_poly:
  assumes "g  carrier P"
  assumes "a  carrier R"
  shows "(to_poly a) of g  = to_poly a"
  by (simp add: assms(1) assms(2) rev_sub_to_poly sub_rev_sub)

lemma sub_const:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = 0"
  shows "f of g = f"
  by (metis lcf_closed assms(1) assms(2) assms(3) sub_to_poly to_poly_inverse)

text‹Substitution into a monomial›

lemma monom_sub:
  assumes "a  carrier R"
  assumes "g  carrier P"
  shows "(monom P a n) of g = a Pg[^]Pn"
    unfolding compose_def
    using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g n] to_poly_UP_pre_univ_prop 
    unfolding P_def  
    using P.nat_pow_closed P_def  to_poly_mult_simp(1) 
    by (simp add: to_poly_mult_simp(1) UP_cring_axioms)
    
lemma(in UP_cring)  cring_sub_monom_bound:
  assumes "a  carrier R"
  assumes "a 𝟬"
  assumes "f = monom P a n"
  assumes "g  carrier P"
  shows "degree (f of g)  n*(degree g)"
proof-
  have "f of g = (to_poly a) P(g[^]Pn)"
    unfolding compose_def
    using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop 
    unfolding P_def  
    by blast
  then show ?thesis 
    by (smt P.nat_pow_closed assms(1) assms(4) cring_pow_deg_bound deg_mult_ring
        degree_to_poly le_trans plus_nat.add_0 to_poly_closed)
qed

lemma(in UP_cring)  cring_sub_monom:
  assumes "a  carrier R"
  assumes "a 𝟬"
  assumes "f = monom P a n"
  assumes "g  carrier P"
  assumes "a  (lcf g [^] n)  𝟬"
  shows "degree (f of g) = n*(degree g)"
proof-
  have 0: "f of g = (to_poly a) P(g[^]Pn)"
    unfolding compose_def
    using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop 
    unfolding P_def  
    by blast
  have 1: "lcf (to_poly a)  lcf (g [^]Pn)  𝟬" 
    using assms 
    by (smt P.nat_pow_closed P_def R.nat_pow_closed R.r_null cring_pow_ltrm lcf_closed lcf_ltrm lcf_monom monom_pow to_polynomial_def)
  then show ?thesis 
    using 0 1 assms cring_pow_deg[of g n] cring_deg_mult[of "to_poly a" "g[^]Pn"]
    by (metis P.nat_pow_closed R.r_null add.right_neutral degree_to_poly to_poly_closed)
qed
    
lemma(in UP_domain)  sub_monom:
  assumes "a  carrier R"
  assumes "a 𝟬"
  assumes "f = monom P a n"
  assumes "g  carrier P"
  shows "degree (f of g) = n*(degree g)"
proof-
  have "f of g = (to_poly a) P(g[^]Pn)"
    unfolding compose_def
    using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop 
    unfolding P_def  
    by blast
  then show ?thesis using deg_pow deg_mult 
    by (metis P.nat_pow_closed P_def assms(1) assms(2) 
        assms(4) deg_smult monom_mult_is_smult to_polynomial_def)
qed

text‹Subbing a constant into a polynomial yields a constant›
lemma sub_in_const:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree g = 0"
  shows "degree (f of g) = 0"
proof-
  have "n. (p. p  carrier P  degree p  n  degree (p of g) = 0)"
  proof-
    fix n
    show "p. p  carrier P  degree p  n  degree (p of g) = 0"
    proof(induction n)
      case 0
      then show ?case 
        by (simp add: assms(1) sub_const)       
    next
      case (Suc n)
      fix n
      assume IH: "p. p  carrier P  degree p  n  degree (p of g) = 0"
      show  "p. p  carrier P  degree p  (Suc n)  degree (p of g) = 0"
      proof-
        fix p
        assume A0: "p  carrier P"
        assume A1: "degree p  (Suc n)"
        show "degree (p of g) = 0"
        proof(cases "degree p < Suc n")
          case True
          then show ?thesis using IH 
            using A0 by auto
        next
          case False
          then have D: "degree p = Suc n" 
            by (simp add: A1 nat_less_le)
          show ?thesis
          proof-
            have P0: "degree ((trunc p) of g) = 0" using IH 
              by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc)
            have P1: "degree ((ltrm p) of g) = 0" 
            proof-
              obtain a n where an_def: "ltrm p = monom P a n  a  carrier R"
                unfolding leading_term_def 
                using A0 P_def cfs_closed by blast
              obtain b where b_def: "g = monom P b 0  b  carrier R"
                using assms deg_zero_impl_monom  coeff_closed 
                by blast               
              have 0: " monom P b 0 [^]Pn = monom P (b[^]n) 0"
                apply(induction n)
                 apply fastforce[1]
              proof- fix n::nat assume IH: "monom P b 0 [^]Pn = monom P (b [^] n) 0"
                have "monom P b 0 [^]PSuc n = (monom P (b[^]n) 0) Pmonom P b 0"
                  using IH by simp
                then have "monom P b 0 [^]PSuc n = (monom P ((b[^]n)b) 0)"
                  using b_def  
                  by (simp add: monom_mult_is_smult monom_mult_smult)
                then show "monom P b 0 [^]PSuc n = monom P (b [^] Suc n) 0 "
                  by simp 
              qed
                   
              then have 0: "a Pmonom P b 0 [^]Pn = monom P (a  b[^]n) 0"
                by (simp add: an_def b_def monom_mult_smult)


              then show ?thesis using monom_sub[of a "monom P b 0" n] assms an_def 
                by (simp add: a  carrier R; monom P b 0  carrier P  monom P a n of monom P b 0 = a Pmonom P b 0 [^]Pn b_def)
            qed
            have P2: "p of g = (trunc p of g) P((ltrm p) of g)"
              by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed)
            then show ?thesis 
              using P0 P1 P2 deg_add[of "trunc p of g" "ltrm p of g"] 
              by (metis A0 assms(1) le_0_eq ltrm_closed max_0R sub_closed trunc_closed)              
          qed
        qed
      qed
    qed
  qed
  then show ?thesis 
    using assms(2) by blast
qed

lemma (in UP_cring)  cring_sub_deg_bound:
  assumes "g  carrier P"
  assumes "f  carrier P"
  shows "degree (f of g)  degree f * degree g"
proof-
  have "n.  p. p  carrier P  (degree p)  n  degree (p of g)  degree p * degree g"
  proof-
    fix n::nat
    show " p. p  carrier P  (degree p)  n  degree (p of g)  degree p * degree g"
    proof(induction n)
      case 0
      then have B0: "degree p = 0" by auto 
      then show ?case using sub_const[of g p] 
        by (simp add: "0.prems"(1) assms(1))
    next
      case (Suc n)
      fix n
      assume IH: "(p. p  carrier P  degree p  n  degree (p of g)  degree p * degree g)"
      show " p  carrier P  degree p  Suc n  degree (p of g)   degree p * degree g"
      proof-
        assume A0: "p  carrier P"
        assume A1: "degree p  Suc n"
        show ?thesis 
        proof(cases "degree p < Suc n")
          case True
          then show ?thesis using IH 
            by (simp add: A0)
        next
          case False
          then have D: "degree p = Suc n" 
            using A1 by auto  
          have P0: "(p of g) = ((trunc p) of g) P((ltrm p) of g)"
            by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed)
          have P1: "degree ((trunc p) of g)  (degree (trunc p))*(degree g)"
            using IH  by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc)
          have P2: "degree ((ltrm p) of g)  (degree p) * degree g"
            using A0 D P_def  UP_cring_axioms assms(1) 
            by (metis False cfs_closed coeff_simp cring_sub_monom_bound deg_zero lcoeff_nonzero2 less_Suc_eq_0_disj)                              
          then show ?thesis
            proof(cases "degree g = 0")
              case True
              then show ?thesis 
                by (simp add: Suc(2) assms(1) sub_in_const)
            next
              case F: False
              then show ?thesis 
              proof-
                have P3: "degree ((trunc p) of g)  n*degree g"
                  using A0 False D  P1 P2  IH[of "trunc p"] trunc_degree[of p]
                proof -
                  { assume "degree (trunc p) < degree p"
                    then have "degree (trunc p)  n"
                      using D by auto
                    then have ?thesis
                      by (meson P1 le_trans mult_le_cancel2) }
                  then show ?thesis
                    by (metis (full_types) A0 D Suc_mult_le_cancel1 nat_mult_le_cancel_disj trunc_degree)
                qed      
                then have P3': "degree ((trunc p) of g) < (degree p)*degree g"
                  using F D by auto 
                have P4: "degree (ltrm p of g)  (degree p)*degree g"
                  using cring_sub_monom_bound  D P2 
                  by auto
                then show ?thesis 
                  using D  P0 P1 P3 P4 A0 P3' assms(1) bound_deg_sum less_imp_le_nat
                    ltrm_closed sub_closed trunc_closed 
                  by metis                   
              qed
            qed
          qed
        qed
      qed
  qed
  then show ?thesis 
    using assms(2) by blast
qed

lemma (in UP_cring)  cring_sub_deg:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "lcf f  (lcf g [^] (degree f))  𝟬"
  shows "degree (f of g) = degree f * degree g"
proof-
  have 0: "f of g = (trunc f of g) P((ltrm f) of g)"
    by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed)
  have 1: "lcf f  𝟬"
    using assms  cring.cring_simprules(26) lcf_closed 
    by auto
  have 2: "degree ((ltrm f) of g) = degree f * degree g"
    using 0 1 assms cring_sub_monom[of "lcf f" "ltrm f" "degree f" g] lcf_closed lcf_ltrm 
    by blast
  show ?thesis
    apply(cases "degree f = 0")
     apply (simp add: assms(1) assms(2))
    apply(cases "degree g = 0")
     apply (simp add: assms(1) assms(2) sub_in_const)
    using 0 1 assms cring_sub_deg_bound[of g "trunc f"] trunc_degree[of f]   
    using sub_const apply auto[1]
    apply(cases "degree g = 0")
    using 0 1 assms cring_sub_deg_bound[of g "trunc f"] trunc_degree[of f]   
    using sub_in_const apply fastforce
    unfolding 0 using 1 2 
    by (smt "0" ltrm_closed f  carrier P; 0 < deg R f  deg R (Cring_Poly.truncate R f) < deg R f
        assms(1) assms(2) cring_sub_deg_bound degree_of_sum_diff_degree equal_deg_sum
        le_eq_less_or_eq mult_less_cancel2 nat_neq_iff neq0_conv sub_closed trunc_closed)    
qed

lemma (in UP_domain)  sub_deg0:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "g  𝟬P⇙"
  assumes "f  𝟬P⇙"
  shows "degree (f of g) = degree f * degree g"
proof-
  have "n.  p. p  carrier P  (degree p)  n  degree (p of g) = degree p * degree g"
  proof-
    fix n::nat
    show " p. p  carrier P  (degree p)  n  degree (p of g) = degree p * degree g"
    proof(induction n)
      case 0
      then have B0: "degree p = 0" by auto 
      then show ?case using sub_const[of g p] 
        by (simp add: "0.prems"(1) assms(1))
    next
      case (Suc n)
      fix n
      assume IH: "(p. p  carrier P  degree p  n  degree (p of g) = degree p * degree g)"
      show " p  carrier P  degree p  Suc n  degree (p of g) = degree p * degree g"
      proof-
        assume A0: "p  carrier P"
        assume A1: "degree p  Suc n"
        show ?thesis 
        proof(cases "degree p < Suc n")
          case True
          then show ?thesis using IH 
            by (simp add: A0)
        next
          case False
          then have D: "degree p = Suc n" 
            using A1 by auto  
          have P0: "(p of g) = ((trunc p) of g) P((ltrm p) of g)"
            by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed)
          have P1: "degree ((trunc p) of g) = (degree (trunc p))*(degree g)"
            using IH  by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc)
          have P2: "degree ((ltrm p) of g) = (degree p) * degree g"
            using A0 D P_def UP_domain.sub_monom UP_cring_axioms assms(1) 
            by (metis False UP_domain_axioms UP_ring.coeff_simp UP_ring.lcoeff_nonzero2 UP_ring_axioms cfs_closed deg_nzero_nzero less_Suc_eq_0_disj)
                     
          then show ?thesis
            proof(cases "degree g = 0")
              case True
              then show ?thesis 
                by (simp add: Suc(2) assms(1) sub_in_const)
            next
              case False
              then show ?thesis 
              proof-
                have P3: "degree ((trunc p) of g) < degree ((ltrm p) of g)"
                  using False D  P1 P2  
                  by (metis (no_types, lifting) A0 mult.commute mult_right_cancel 
                      nat_less_le nat_mult_le_cancel_disj trunc_degree zero_less_Suc)
                then show ?thesis 
                 by (simp add: A0 ltrm_closed P0 P2 assms(1) equal_deg_sum sub_closed trunc_closed)                 
              qed
            qed
          qed
        qed
      qed
  qed
  then show ?thesis 
      using assms(2) by blast
qed

lemma(in UP_domain)  sub_deg:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "g  𝟬P⇙"
  shows "degree (f of g) = degree f * degree g"
proof(cases "f = 𝟬P⇙")
  case True
  then show ?thesis 
    using assms(1)  sub_const by auto
next
  case False
  then show ?thesis 
    by (simp add: assms(1) assms(2) assms(3) sub_deg0)
qed

lemma(in UP_cring)  cring_ltrm_sub:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree g > 0"
  assumes "lcf f  (lcf g [^] (degree f))  𝟬"
  shows "ltrm (f of g) = ltrm ((ltrm f) of g)"
proof-
  have P0: "degree (f of g) = degree ((ltrm f) of g)"
    using assms(1) assms(2) assms(4) cring_sub_deg lcf_eq ltrm_closed deg_ltrm 
    by auto               
  have P1: "f of g = ((trunc f) of g) P((ltrm f) of g)"
    by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed)
  then show ?thesis
  proof(cases "degree f = 0")
    case True
    then show ?thesis 
      using ltrm_deg_0 assms(2) by auto     
  next
    case False
    have P2: "degree (f of g) = degree f * degree g"
      by (simp add: assms(1) assms(2) assms(4) cring_sub_deg)      
    then have P3: "degree ((trunc f) of g) < degree ((ltrm f) of g)"
      using False P0 P1 P_def UP_cring.sub_closed trunc_closed UP_cring_axioms
          UP_ring.degree_of_sum_diff_degree UP_ring.ltrm_closed UP_ring_axioms assms(1) 
          assms(2) assms(4) cring_sub_deg_bound le_antisym less_imp_le_nat less_nat_zero_code
          mult_right_le_imp_le nat_neq_iff trunc_degree
      by (smt assms(3))
    then show ?thesis using P0 P1 P2 
      by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm assms(1) assms(2) sub_closed trunc_closed)      
  qed
qed

lemma(in UP_domain)  ltrm_sub:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree g > 0"
  shows "ltrm (f of g) = ltrm ((ltrm f) of g)"
proof-
  have P0: "degree (f of g) = degree ((ltrm f) of g)"
    using sub_deg 
    by (metis ltrm_closed assms(1) assms(2) assms(3) deg_zero deg_ltrm nat_neq_iff)       
  have P1: "f of g = ((trunc f) of g) P((ltrm f) of g)"
    by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed)
  then show ?thesis
  proof(cases "degree f = 0")
    case True
    then show ?thesis 
      using ltrm_deg_0 assms(2) by auto    
  next
    case False
    then have P2: "degree ((trunc f) of g) < degree ((ltrm f) of g)"
      using sub_deg 
      by (metis (no_types, lifting) ltrm_closed assms(1) assms(2) assms(3) deg_zero
          deg_ltrm mult_less_cancel2 neq0_conv trunc_closed trunc_degree)      
    then show ?thesis 
      using P0 P1 P2 
      by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm assms(1) assms(2) sub_closed trunc_closed)
  qed
qed

lemma(in UP_cring)  cring_lcf_of_sub_in_ltrm:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  assumes "(lcf f)  ((lcf g)[^]n) 𝟬"
  shows "lcf ((ltrm f) of g) = (lcf f)  ((lcf g)[^]n)"
  by (metis (no_types, lifting) P.nat_pow_closed P_def R.r_null UP_cring.monom_sub UP_cring_axioms 
      assms(1) assms(2) assms(3) assms(5) cfs_closed cring_lcf_pow cring_lcf_scalar_mult)

lemma(in UP_domain)  lcf_of_sub_in_ltrm:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  shows "lcf ((ltrm f) of g) = (lcf f)  ((lcf g)[^]n)"
proof(cases "degree f = 0")
  case True
  then show ?thesis 
    using ltrm_deg_0 assms(1) assms(2) assms(3) cfs_closed 
    by (simp add: sub_const)       
next
  case False
  then show ?thesis 
  proof-
    have P0: "(ltrm f) of g = (to_poly (lcf f)) P(g[^]Pn)"
      unfolding compose_def
      using assms UP_pre_univ_prop.eval_monom[of R P to_poly "(lcf f)" g n] to_poly_UP_pre_univ_prop 
      unfolding P_def  
      using P_def cfs_closed by blast        
    have P1: "(ltrm f) of g = (lcf f) P(g[^]Pn)"
      using P0 P.nat_pow_closed 
      by (simp add: assms(1) assms(2) assms(3) cfs_closed monom_sub)          
    have P2: "ltrm ((ltrm f) of g) = (ltrm (to_poly (lcf f))) P(ltrm (g[^]Pn))"
      using P0 ltrm_mult P.nat_pow_closed P_def assms(1) assms(2) 
         to_poly_closed 
      by (simp add: cfs_closed)                  
    have P3: "ltrm ((ltrm f) of g) =  (to_poly (lcf f)) P(ltrm (g[^]Pn))"
      using P2  ltrm_deg_0 assms(2) to_poly_closed 
      by (simp add: cfs_closed)           
    have P4: "ltrm ((ltrm f) of g) = (lcf f) P((ltrm g)[^]Pn)"
      using P.nat_pow_closed P1 P_def assms(1) assms(2) ltrm_pow0 ltrm_smult 
      by (simp add: cfs_closed)
    have P5: "lcf ((ltrm f) of g) = (lcf f)  (lcf ((ltrm g)[^]Pn))"
      using lcf_scalar_mult P4  by (metis P.nat_pow_closed P1 cfs_closed 
          UP_smult_closed assms(1) assms(2) assms(3) lcf_eq ltrm_closed sub_rev_sub)
    show ?thesis
      using P5 ltrm_pow lcf_pow assms(1) lcf_eq ltrm_closed by presburger
  qed
qed

lemma(in UP_cring)  cring_ltrm_of_sub_in_ltrm:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  assumes "(lcf f)  ((lcf g)[^]n) 𝟬"
  shows "ltrm ((ltrm f) of g) = (lcf f) P((ltrm g)[^]Pn)"
  by (smt lcf_eq ltrm_closed R.nat_pow_closed R.r_null assms(1) assms(2) assms(3) 
      assms(4) assms(5) cfs_closed cring_lcf_of_sub_in_ltrm cring_lcf_pow cring_pow_ltrm
      cring_pow_deg cring_sub_deg deg_zero deg_ltrm monom_mult_smult neq0_conv)
  
lemma(in UP_domain)  ltrm_of_sub_in_ltrm:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  shows "ltrm ((ltrm f) of g) = (lcf f) P((ltrm g)[^]Pn)"
  by (smt Group.nat_pow_0 lcf_of_sub_in_ltrm lcf_pow lcf_scalar_mult ltrm_closed 
      ltrm_pow0 ltrm_smult P.nat_pow_closed P_def UP_ring.monom_one UP_ring_axioms assms(1) 
      assms(2) assms(3) assms(4) cfs_closed coeff_simp deg_const deg_nzero_nzero deg_pow 
      deg_smult deg_ltrm lcoeff_nonzero2 nat_less_le sub_deg)

text‹formula for the leading term of a composition ›

lemma(in UP_domain)  cring_ltrm_of_sub:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  assumes "(lcf f)  ((lcf g)[^]n) 𝟬"
  shows "ltrm (f of g) = (lcf f) P((ltrm g)[^]Pn)"
  using ltrm_of_sub_in_ltrm ltrm_sub assms(1) assms(2) assms(3) assms(4) by presburger
    
lemma(in UP_domain)  ltrm_of_sub:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "degree f = n"
  assumes "degree g > 0"
  shows "ltrm (f of g) = (lcf f) P((ltrm g)[^]Pn)"
  using ltrm_of_sub_in_ltrm ltrm_sub assms(1) assms(2) assms(3) assms(4) by presburger
   
text‹subtitution is associative›

lemma  sub_assoc_monom:
  assumes "f  carrier P"
  assumes "q  carrier P"
  assumes "r  carrier P"
  shows "(ltrm f) of (q of r) = ((ltrm f) of q) of r"
proof-
  obtain n where n_def: "n = degree f"
    by simp
  obtain a where a_def: "a  carrier R  (ltrm f) = monom P a n"
    using assms(1) cfs_closed n_def by blast         
  have LHS: "(ltrm f) of (q of r) = a P(q of r)[^]Pn"
    by (metis P.nat_pow_closed P_def UP_pre_univ_prop.eval_monom a_def assms(2)
        assms(3) compose_def monom_mult_is_smult sub_closed to_poly_UP_pre_univ_prop to_polynomial_def)
  have RHS0: "((ltrm f) of q) of r = (a Pq[^]Pn)of r"
    by (metis P.nat_pow_closed P_def UP_pre_univ_prop.eval_monom a_def 
        assms(2) compose_def monom_mult_is_smult to_poly_UP_pre_univ_prop to_polynomial_def)
  have RHS1: "((ltrm f) of q) of r = ((to_poly a) Pq[^]Pn)of r"
    using RHS0  by (metis P.nat_pow_closed P_def a_def 
        assms(2) monom_mult_is_smult to_polynomial_def)
  have RHS2: "((ltrm f) of q) of r = ((to_poly a) of r) P(q[^]Pn of r)"
    using RHS1 a_def assms(2) assms(3) sub_mult to_poly_closed by auto
  have RHS3: "((ltrm f) of q) of r = (to_poly a) P(q[^]Pn of r)"
    using RHS2 a_def assms(3) sub_to_poly by auto
  have RHS4: "((ltrm f) of q) of r = a P((q[^]Pn)of r)"
    using RHS3 
    by (metis P.nat_pow_closed P_def a_def assms(2) assms(3) 
        monom_mult_is_smult sub_closed to_polynomial_def)
  have "(q of r)[^]Pn = ((q[^]Pn)of r)" 
    apply(induction n) 
    apply (metis Group.nat_pow_0 P.ring_simprules(6) assms(3) deg_one sub_const)
    by (simp add: assms(2) assms(3) sub_mult)    
  then show ?thesis using RHS4 LHS by simp
qed

lemma sub_assoc:
  assumes "f  carrier P"
  assumes "q  carrier P"
  assumes "r  carrier P"
  shows "f of (q of r) = (f of q) of r"
proof-
  have " n.  p. p  carrier P  degree p  n  p of (q of r) = (p of q) of r"
  proof-
    fix n
    show " p. p  carrier P  degree p  n  p of (q of r) = (p of q) of r"
    proof(induction n)
      case 0
      then have deg_p: "degree p = 0"
        by blast
      then have B0: "p of (q of r) = p"
        using sub_const[of "q of r" p] assms  "0.prems"(1) sub_closed by blast
      have B1: "(p of q) of r = p"
      proof-
        have p0: "p of q = p"
          using deg_p 0 assms(2) 
          by (simp add: P_def UP_cring.sub_const UP_cring_axioms)                             
        show ?thesis 
          unfolding p0 using deg_p 0 assms(3) 
          by (simp add: P_def UP_cring.sub_const UP_cring_axioms)                             
      qed
      then show "p of (q of r) = (p of q) of r" using B0 B1 by auto 
    next
      case (Suc n)
      fix n
      assume IH: " p. p  carrier P  degree p  n  p of (q of r) = (p of q) of r"
      then show " p. p  carrier P  degree p  Suc n  p of (q of r) = (p of q) of r"
      proof-
        fix p
        assume A0: " p  carrier P "
        assume A1: "degree p  Suc n"
        show "p of (q of r) = (p of q) of r"
        proof(cases "degree p < Suc n")
          case True
          then show ?thesis using A0 A1 IH by auto 
        next
          case False
          then have "degree p = Suc n"
            using A1 by auto 
          have I0: "p of (q of r) = ((trunc p) P(ltrm p)) of (q of r)"
            using A0 trunc_simps(1) by auto
          have I1: "p of (q of r) = ((trunc p)  of (q of r)) P((ltrm p)  of (q of r))"
            using I0 sub_add 
            by (simp add: A0 assms(2) assms(3) ltrm_closed rev_sub_closed sub_rev_sub trunc_closed)
          have I2: "p of (q of r) = (((trunc p)  of q) of r) P(((ltrm p)  of q) of r)"
            using IH[of "trunc p"] sub_assoc_monom[of p q r] 
            by (metis A0 I1 degree p = Suc n assms(2) assms(3) 
                less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc)
          have I3: "p of (q of r) = (((trunc p)  of q) P((ltrm p)  of q)) of r"
            using sub_add trunc_simps(1) assms   
            by (simp add: A0 I2 ltrm_closed sub_closed trunc_closed)
          have I4: "p of (q of r) = (((trunc p)P(ltrm p))   of q)  of r"
            using sub_add trunc_simps(1) assms   
            by (simp add: trunc_simps(1) A0 I3 ltrm_closed trunc_closed)
          then show ?thesis 
            using A0 trunc_simps(1) by auto
        qed
      qed
    qed
  qed
  then show ?thesis 
    using assms(1) by blast
qed

lemma sub_smult:
  assumes "f  carrier P"
  assumes "q  carrier P"
  assumes "a  carrier R"
  shows "(aPf ) of q = aP(f of q)"
proof-
  have "(aPf ) of q = ((to_poly a) Pf) of q"
    using assms  by (metis P_def monom_mult_is_smult to_polynomial_def)
  then have "(aPf ) of q = ((to_poly a) of q) P(f of q)"
    by (simp add: assms(1) assms(2) assms(3) sub_mult to_poly_closed)
  then have "(aPf ) of q = (to_poly a) P(f of q)"
    by (simp add: assms(2) assms(3) sub_to_poly)
  then show ?thesis 
    by (metis P_def assms(1) assms(2) assms(3) 
        monom_mult_is_smult sub_closed to_polynomial_def)
qed

lemma to_fun_sub_monom:
  assumes "is_UP_monom f"
  assumes "g  carrier P"
  assumes "a  carrier R"
  shows "to_fun (f of g) a = to_fun f (to_fun g a)"   
proof-
  obtain b n where b_def: "b  carrier R  f = monom P b n"
    using assms unfolding is_UP_monom_def 
     using P_def cfs_closed by blast     
  then have P0: "f of g = b P(g[^]Pn)"
    using b_def assms(2) monom_sub by blast
  have P1: "UP_pre_univ_prop R R (λx. x)" 
    by (simp add: UP_pre_univ_prop_fact)
  then have P2: "to_fun f (to_fun g a) = b ((to_fun g a)[^]n)"
    using P1 to_fun_eval[of f "to_fun g a"] P_def UP_pre_univ_prop.eval_monom assms(1)
          assms(2) assms(3) b_def is_UP_monomE(1) to_fun_closed 
    by force
  have P3: "to_fun (monom P b n of g) a =  b ((to_fun g a)[^]n)"
  proof-
    have 0: "to_fun (monom P b n of g) a = eval R R (λx. x) a  (b P(g[^]Pn) )"

      using UP_pre_univ_prop.eval_monom[of R "(UP R)" to_poly b g n]
            P_def assms(2) b_def to_poly_UP_pre_univ_prop to_fun_eval P0
      by (metis assms(3) monom_closed sub_closed)      
    have 1: "to_fun (monom P b n of g) a = (eval R R (λx. x) a  (to_poly b))  ( eval R R (λx. x) a ( g [^]UP Rn ))"
      using 0 eval_ring_hom 
      by (metis P.nat_pow_closed P0 P_def assms(2) assms(3) b_def monom_mult_is_smult to_fun_eval to_fun_mult to_poly_closed to_polynomial_def)                
    have 2: "to_fun (monom P b n of g) a = b  ( eval R R (λx. x) a ( g [^]UP Rn ))"
      using 1 assms(3) b_def to_fun_eval to_fun_to_poly to_poly_closed by auto              
    then show ?thesis 
      unfolding to_function_def to_fun_def 
      using eval_ring_hom P_def UP_pre_univ_prop.ring_homD UP_pre_univ_prop_fact 
            assms(2) assms(3) ring_hom_cring.hom_pow by fastforce
  qed
  then show ?thesis 
    using b_def P2 by auto
qed

lemma to_fun_sub:
  assumes "g  carrier P"
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "to_fun (f of  g) a = (to_fun f) (to_fun g a)"
proof(rule poly_induct2[of f])
  show "f  carrier P" 
    using assms by auto 
  show "p. p  carrier P  degree p = 0  to_fun (p of g) a = to_fun p (to_fun g a)"
  proof-
    fix p
    assume A0: "p  carrier P"
    assume A1: "degree p = 0"
    then have P0: "degree (p of g) = 0"
      by (simp add: A0 assms(1) sub_const)    
    then obtain b where b_def: "p of g = to_poly b  b  carrier R"
      using A0 A1 cfs_closed assms(1) to_poly_inverse 
      by (meson sub_closed)
    then have "to_fun (p of g) a = b" 
      by (simp add: assms(3) to_fun_to_poly)     
    have "p of g = p"
      using A0 A1 P_def sub_const UP_cring_axioms assms(1) by blast
    then have P1: "p = to_poly b"
      using b_def by auto
    have "to_fun g a  carrier R"
      using assms
      by (simp add: to_fun_closed)  
    then show "to_fun (p of g) a =  to_fun p (to_fun g a)"
      using P1 to_fun (p of g) a = b b_def 
      by (simp add: to_fun_to_poly)      
  qed
  show "p. 0 < degree p  p  carrier P  
            to_fun (trunc p of g) a = to_fun (trunc p) (to_fun g a)  
            to_fun (p of g) a = to_fun p (to_fun g a)"
  proof-
    fix p
    assume A0: "0 < degree p"
    assume A1: " p  carrier P"
    assume A2: "to_fun (trunc p of g) a = to_fun (trunc p) (to_fun g a)"
    show "to_fun (p of g) a = to_fun p (to_fun g a)"
    proof-
      have "p of g = (trunc p) of g P(ltrm p) of g"
        by (metis A1 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed)
      then have "to_fun (p of g) a = to_fun ((trunc p) of g) a  (to_fun ((ltrm p) of g) a)"
        by (simp add: A1 assms(1) assms(3) to_fun_plus ltrm_closed sub_closed trunc_closed)        
      then have 0: "to_fun (p of g) a = to_fun (trunc p) (to_fun g a)  (to_fun ((ltrm p) of g) a)"
        by (simp add: A2)
      have "(to_fun ((ltrm p) of g) a) = to_fun (ltrm p) (to_fun g a)"
        using to_fun_sub_monom 
        by (simp add: A1 assms(1) assms(3) ltrm_is_UP_monom)
      then have "to_fun (p of g) a = to_fun (trunc p) (to_fun g a)   to_fun (ltrm p) (to_fun g a)"
        using 0 by auto 
      then show ?thesis 
        by (metis A1 assms(1) assms(3) to_fun_closed to_fun_plus ltrm_closed trunc_simps(1) trunc_closed)
    qed
  qed
qed
end


text‹More material on constant terms and constant coefficients›

context UP_cring
begin

lemma to_fun_ctrm:
  assumes "f  carrier P"
  assumes  "b  carrier R"  
  shows "to_fun (ctrm f) b = (f 0)"
  using assms 
  by (metis ctrm_degree ctrm_is_poly lcf_monom(2) P_def cfs_closed to_fun_to_poly to_poly_inverse)   

lemma to_fun_smult:
  assumes "f  carrier P"
  assumes "b  carrier R"
  assumes "c  carrier R"
  shows "to_fun (c Pf) b = c (to_fun f b)"
proof-
  have "(c Pf) = (to_poly c) Pf"
    by (metis P_def assms(1) assms(3) monom_mult_is_smult to_polynomial_def)
  then have "to_fun (c Pf) b = to_fun (to_poly c) b  to_fun f b"
    by (simp add: assms(1) assms(2) assms(3) to_fun_mult to_poly_closed)    
  then show  ?thesis 
    by (simp add: assms(2) assms(3) to_fun_to_poly)  
qed

lemma to_fun_monom:
  assumes "c  carrier R"
  assumes "x  carrier R"
  shows "to_fun (monom P c n) x = c  x [^] n"
  by (smt P_def R.m_comm R.nat_pow_closed UP_cring.to_poly_nat_pow UP_cring_axioms assms(1) 
      assms(2) monom_is_UP_monom(1) sub_monom(1) to_fun_smult to_fun_sub_monom to_fun_to_poly 
      to_poly_closed to_poly_mult_simp(2))

lemma zcf_monom:
  assumes "a  carrier R"
  shows "zcf (monom P a n) = to_fun (monom P a n) 𝟬"
  using to_fun_monom unfolding zcf_def 
  by (simp add: R.nat_pow_zero assms cfs_monom)

lemma zcf_to_fun: 
  assumes "p  carrier P"
  shows "zcf p = to_fun p 𝟬"
  apply(rule poly_induct3[of p])
  apply (simp add: assms)
  using R.zero_closed zcf_add to_fun_plus apply presburger
  using zcf_monom by blast

lemma zcf_to_poly[simp]:
  assumes "a  carrier R"
  shows "zcf (to_poly a) = a"
  by (metis assms cfs_closed degree_to_poly to_fun_to_poly to_poly_inverse to_poly_closed zcf_def)

lemma zcf_ltrm_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree p > 0"
  shows "zcf((ltrm p) Pq) = 𝟬"
  using zcf_to_fun[of "ltrm p Pq" ]
  by (metis ltrm_closed P.l_null P.m_closed R.zero_closed UP_zero_closed zcf_to_fun
      zcf_zero assms(1) assms(2) assms(3) coeff_ltrm to_fun_mult)

lemma zcf_mult:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "zcf(p Pq) = (zcf p)  (zcf q)"
  using zcf_to_fun[of " p Pq" ] zcf_to_fun[of "p" ] zcf_to_fun[of "q" ] to_fun_mult[of q p 𝟬]
  by (simp add: assms(1) assms(2))

lemma zcf_is_ring_hom:
"zcf ring_hom P R"
  apply(rule ring_hom_memI)
  using zcf_mult zcf_add
  apply (simp add: P_def UP_ring.cfs_closed UP_ring_axioms zcf_def)
    apply (simp add: zcf_mult)
      using zcf_add apply auto[1]
        by simp  

lemma ctrm_is_ring_hom:
"ctrm  ring_hom P P"
  apply(rule ring_hom_memI)
     apply (simp add: ctrm_is_poly)
      apply (metis zcf_def zcf_mult cfs_closed monom_mult zero_eq_add_iff_both_eq_0)     
        using cfs_add[of _ _ 0] 
        apply (simp add: cfs_closed)        
  by auto  


(**************************************************************************************************)
(**************************************************************************************************)
section‹Describing the Image of (UP R) in the Ring of Functions from R to R›
(**************************************************************************************************)
(**************************************************************************************************)

lemma to_fun_diff:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "a  carrier R"
  shows "to_fun (p Pq) a = to_fun p a  to_fun q a"
  using to_fun_plus[of "Pq" p a]
  by (simp add: P.minus_eq R.minus_eq assms(1) assms(2) assms(3) to_fun_minus)

lemma to_fun_const:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_fun (monom P a 0) b = a"
  by (metis lcf_monom(2) P_def UP_cring.to_fun_ctrm UP_cring_axioms assms(1) assms(2) deg_const monom_closed)

lemma to_fun_monic_monom:
  assumes "b  carrier R"
  shows "to_fun (monom P 𝟭 n) b = b[^]n"
  by (simp add: assms to_fun_monom)


text‹Constant polynomials map to constant polynomials›

lemma const_to_constant:
  assumes "a  carrier R"
  shows "to_fun (monom P a 0) = constant_function (carrier R) a"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  unfolding ring_functions_def apply(simp add: R.ring_axioms)
    apply (simp add: assms to_fun_is_Fun)
      using assms ring_functions.constant_function_closed[of R a "carrier R"]
      unfolding ring_functions_def apply (simp add: R.ring_axioms)
        using assms to_fun_const[of a ] unfolding constant_function_def 
        by auto 

text‹Monomial polynomials map to monomial functions›

lemma monom_to_monomial:
  assumes "a  carrier R"
  shows "to_fun (monom P a n) = monomial_function R a n"
  apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"])
  unfolding ring_functions_def apply(simp add: R.ring_axioms)
    apply (simp add: assms to_fun_is_Fun)
      using assms U_function_ring.monomial_functions[of R a n] R.ring_axioms
      unfolding U_function_ring_def  
      apply auto[1]
        unfolding monomial_function_def 
        using assms to_fun_monom[of a _ n]
        by auto 
end 


(**************************************************************************************************)
(**************************************************************************************************)
section‹Taylor Expansions›
(**************************************************************************************************)
(**************************************************************************************************)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Monic Linear Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

text‹The polynomial representing the variable X›

definition X_poly where
"X_poly R = monom (UP R) 𝟭R1"

context UP_cring
begin

abbreviation(input) X where
"X  X_poly R"

lemma X_closed:
"X  carrier P"
  unfolding X_poly_def 
  using P_def monom_closed by blast

lemma degree_X[simp]:
  assumes "𝟭 𝟬"
  shows"degree X = 1"
  unfolding X_poly_def 
  using assms P_def deg_monom[of 𝟭 1]  
  by blast

lemma X_not_zero:
  assumes "𝟭 𝟬"
  shows"X  𝟬P⇙"
  using degree_X assms by force  

lemma sub_X[simp]:
  assumes "p  carrier P"
  shows "X of p = p"
  unfolding X_poly_def
  using P_def UP_pre_univ_prop.eval_monom1 assms compose_def to_poly_UP_pre_univ_prop 
  by metis

lemma sub_monom_deg_one:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "monom P a 1 of p = a Pp"
  using assms sub_smult[of X p a] unfolding X_poly_def
  by (metis P_def R.one_closed R.r_one X_closed X_poly_def monom_mult_smult sub_X)

lemma monom_rep_X_pow:
  assumes "a  carrier R"
  shows "monom P a n = aP(X[^]Pn)"
proof-
  have "monom P a n = aPmonom P 𝟭 n"
    by (metis R.one_closed R.r_one assms monom_mult_smult)
  then show ?thesis 
    unfolding X_poly_def 
    using monom_pow 
    by (simp add: P_def)
qed

lemma X_sub[simp]:
  assumes "p  carrier P"
  shows "p of X = p"
  apply(rule poly_induct3)
  apply (simp add: assms)
  using X_closed sub_add apply presburger
  using sub_monom[of X] P_def monom_rep_X_pow X_closed by auto

text‹representation of monomials as scalar multiples of powers of X›

lemma ltrm_rep_X_pow:
  assumes "p  carrier P"
  shows "ltrm p = (lcf p)P(X[^]P(degree p))"
proof-
  have "ltrm p =  monom P (lcf p) (degree p)"
    using assms unfolding leading_term_def by (simp add: P_def)
  then show ?thesis 
    using monom_rep_X_pow P_def assms 
    by (simp add: cfs_closed)     
qed

lemma to_fun_monom':
  assumes "c  carrier R"
  assumes "c 𝟬"
  assumes "x  carrier R"
  shows "to_fun (c PX[^]P(n::nat)) x = c  x [^] n"
  using P_def to_fun_monom monom_rep_X_pow UP_cring_axioms assms(1) assms(2) assms(3) by fastforce

lemma to_fun_X_pow:
  assumes "x  carrier R"
  shows "to_fun (X[^]P(n::nat)) x = x [^] n"
  using to_fun_monom'[of 𝟭 x n] assms 
  by (metis P.nat_pow_closed R.l_one R.nat_pow_closed R.one_closed R.r_null R.r_one 
      UP_one_closed X_closed to_fun_to_poly ring_hom_one smult_l_null smult_one to_poly_is_ring_hom)  
end

text‹Monic linear polynomials›

definition X_poly_plus where
"X_poly_plus R a = (X_poly R) (UP R)to_polynomial R a"

definition X_poly_minus where
"X_poly_minus R a = (X_poly R) (UP R)to_polynomial R a"

context UP_cring
begin

abbreviation(input) X_plus where
"X_plus  X_poly_plus R"

abbreviation(input) X_minus where
"X_minus  X_poly_minus R"

lemma X_plus_closed:
  assumes "a  carrier R"
  shows "(X_plus a)  carrier P"
  unfolding X_poly_plus_def using X_closed to_poly_closed 
  using P_def UP_a_closed assms by auto

lemma X_minus_closed:
  assumes "a  carrier R"
  shows "(X_minus a)  carrier P"
  unfolding X_poly_minus_def using X_closed to_poly_closed
  by (simp add: P_def UP_cring.UP_cring UP_cring_axioms assms cring.cring_simprules(4))   

lemma X_minus_plus:
  assumes "a  carrier R"
  shows "(X_minus a) = X_plus (a)"
  using P_def UP_ring.UP_ring  UP_ring_axioms
  by (simp add: X_poly_minus_def X_poly_plus_def a_minus_def assms to_poly_a_inv)
  
lemma degree_of_X_plus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "degree (X_plus a) = 1"
proof-
  have 0:"degree (X_plus a)  1"
    using deg_add degree_X  P_def unfolding X_poly_plus_def 
      using UP_cring.to_poly_closed UP_cring_axioms X_closed assms(1) assms(2) by fastforce      
  have 1:"degree (X_plus a) > 0"
    by (metis One_nat_def P_def R.one_closed R.r_zero X_poly_def 
        X_closed X_poly_plus_def X_plus_closed  assms  coeff_add coeff_monom deg_aboveD  
         gr0I  lessI  n_not_Suc_n to_polynomial_def to_poly_closed)
  then show ?thesis 
    using "0" by linarith
qed

lemma degree_of_X_minus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "degree (X_minus a) = 1"
  using degree_of_X_plus[of "a"] X_minus_plus[simp] assms by auto  

lemma ltrm_of_X:
shows"ltrm X = X"
  unfolding leading_term_def
  by (metis P_def R.one_closed X_poly_def is_UP_monom_def is_UP_monomI leading_term_def)

lemma ltrm_of_X_plus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "ltrm (X_plus a) = X"
  unfolding X_poly_plus_def
  using X_closed  assms  ltrm_of_sum_diff_degree[of X "to_poly a"]  
    degree_to_poly[of a]  to_poly_closed[of a] degree_X ltrm_of_X 
    by (simp add: P_def)  

lemma ltrm_of_X_minus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "ltrm (X_minus a) = X"
  using X_minus_plus[of a]  assms 
  by (simp add: ltrm_of_X_plus)

lemma lcf_of_X_minus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "lcf (X_minus a) = 𝟭"
  using ltrm_of_X_minus unfolding X_poly_def 
  using P_def UP_cring.X_minus_closed UP_cring.lcf_eq UP_cring_axioms assms(1) assms(2) lcf_monom 
  by (metis R.one_closed)
  
lemma lcf_of_X_plus:
  assumes "a  carrier R"
  assumes "𝟭 𝟬"
  shows "lcf (X_plus a) = 𝟭"
  using ltrm_of_X_plus unfolding X_poly_def 
  by (metis lcf_of_X_minus P_def UP_cring.lcf_eq UP_cring.X_plus_closed UP_cring_axioms X_minus_closed assms(1) assms(2) degree_of_X_minus)

lemma to_fun_X[simp]:
  assumes "a  carrier R"
  shows "to_fun X a = a"
  using X_closed assms to_fun_sub_monom ltrm_is_UP_monom ltrm_of_X to_poly_closed 
  by (metis sub_X to_fun_to_poly)

lemma to_fun_X_plus[simp]:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_fun (X_plus a) b = b  a"
  unfolding X_poly_plus_def 
  using assms to_fun_X[of b] to_fun_plus[of "to_poly a" X  b] to_fun_to_poly[of a b] 
  using P_def X_closed to_poly_closed by auto

lemma to_fun_X_minus[simp]:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_fun (X_minus a) b = b  a"
  using to_fun_X_plus[of " a" b] X_minus_plus[of a] assms 
  by (simp add: R.minus_eq)

lemma cfs_X_plus:
  assumes "a  carrier R"
  shows "X_plus a n = (if n = 0 then a else (if n = 1 then 𝟭 else 𝟬))"
  using assms cfs_add monom_closed UP_ring_axioms cfs_monom 
  unfolding X_poly_plus_def to_polynomial_def X_poly_def P_def    
  by auto 

lemma cfs_X_minus:
  assumes "a  carrier R"
  shows "X_minus a n = (if n = 0 then  a else (if n = 1 then 𝟭 else 𝟬))"
  using cfs_X_plus[of " a"] assms
  unfolding X_poly_plus_def X_poly_minus_def 
  by (simp add: P_def a_minus_def to_poly_a_inv)

text‹Linear substituions›
 
lemma X_plus_sub_deg:
  assumes "a  carrier R"
  assumes "f  carrier P"
  shows "degree (f of (X_plus a)) = degree f"
  apply(cases "𝟭 = 𝟬")
   apply (metis P_def UP_one_closed X_plus_closed X_poly_def sub_X assms(1) assms(2) deg_one monom_one monom_zero sub_const)
  using cring_sub_deg[of "X_plus a" f] assms X_plus_closed[of a] lcf_of_X_plus[of a]  
        ltrm_of_X_plus degree_of_X_plus[of a] P_def 
  by (metis lcf_eq R.nat_pow_one R.r_one UP_cring.cring_sub_deg UP_cring_axioms X_closed X_sub 
      cfs_closed coeff_simp deg_nzero_nzero degree_X lcoeff_nonzero2 sub_const) 

lemma X_minus_sub_deg:
  assumes "a  carrier R"
  assumes "f  carrier P"
  shows "degree (f of (X_minus a)) = degree f"
  using X_plus_sub_deg[of "a"] assms X_minus_plus[of a]
  by simp

lemma plus_minus_sub:
  assumes " a  carrier R"
  shows "X_plus a of X_minus a = X"
  unfolding X_poly_plus_def
proof-
  have "(X Pto_poly a) of X_minus a = (X  of X_minus a) P(to_poly a) of X_minus a"
    using sub_add
    by (simp add: X_closed X_minus_closed assms to_poly_closed)     
  then have "(X Pto_poly a) of X_minus a = (X_minus a) P(to_poly a)"
    by (simp add: X_minus_closed assms sub_to_poly)
  then show "(X UP Rto_poly a) of X_minus a = X" 
    unfolding to_polynomial_def X_poly_minus_def
    by (metis P.add.inv_solve_right P.minus_eq P_def 
        X_closed X_poly_minus_def X_minus_closed assms monom_closed to_polynomial_def)
qed

lemma minus_plus_sub:
  assumes " a  carrier R"
  shows "X_minus a of X_plus a = X"
  using plus_minus_sub[of "a"]
  unfolding X_poly_minus_def
  unfolding X_poly_plus_def
  using assms  apply simp
  by (metis P_def R.add.inv_closed R.minus_minus a_minus_def to_poly_a_inv)

lemma ltrm_times_X:
  assumes "p  carrier P"
  shows "ltrm (X Pp) = X P(ltrm p)"
  using assms ltrm_of_X cring_ltrm_mult[of X p]
  by (metis ltrm_deg_0 P.r_null R.l_one R.one_closed UP_cring.lcf_monom(1)
      UP_cring_axioms X_closed X_poly_def cfs_closed deg_zero deg_ltrm monom_zero)

lemma times_X_not_zero:
  assumes "p  carrier P"
  assumes "p  𝟬P⇙"
  shows "(X Pp)  𝟬P⇙"
  by (metis (no_types, opaque_lifting) lcf_monom(1) lcf_of_X_minus ltrm_of_X_minus P.inv_unique
      P.r_null R.l_one R.one_closed UP_zero_closed X_closed zcf_def 
      zcf_zero_degree_zero assms(1) assms(2) cfs_closed cfs_zero cring_lcf_mult 
      deg_monom deg_nzero_nzero deg_ltrm degree_X degree_of_X_minus 
      monom_one monom_zero)

lemma degree_times_X:
  assumes "p  carrier P"
  assumes "p  𝟬P⇙"
  shows "degree (X Pp) = degree p + 1"
  using cring_deg_mult[of X p] assms times_X_not_zero[of p]
  by (metis (no_types, lifting) P.r_null P.r_one P_def R.l_one R.one_closed
      UP_cring.lcf_monom(1) UP_cring_axioms UP_zero_closed X_closed X_poly_def cfs_closed 
      deg_zero deg_ltrm degree_X monom_one monom_zero to_poly_inverse)
   
end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Basic Facts About Taylor Expansions›
(**************************************************************************************************)
(**************************************************************************************************)
definition taylor_expansion where
"taylor_expansion R a p = compose R p (X_poly_plus R a)"

definition(in UP_cring) taylor  where
"taylor  taylor_expansion R"

context UP_cring
begin

lemma taylor_expansion_ring_hom:
  assumes "c  carrier R"
  shows "taylor_expansion R c  ring_hom P P"
  unfolding taylor_expansion_def 
  using rev_sub_is_hom[of "X_plus c"]
  unfolding rev_compose_def compose_def 
  using X_plus_closed assms by auto 

notation  taylor ("T⇘_")

lemma(in UP_cring) taylor_closed:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "T⇘af   carrier P"
  unfolding taylor_def
  by (simp add: X_plus_closed assms(1) assms(2) sub_closed taylor_expansion_def)

lemma taylor_deg:
  assumes "a  carrier R"
  assumes "p  carrier P"
  shows "degree (T⇘ap) = degree p"
  unfolding taylor_def taylor_expansion_def 
  using X_plus_sub_deg[of a p] assms 
  by (simp add: taylor_expansion_def)

lemma taylor_id:
  assumes "a  carrier R"
  assumes "p  carrier P"
  shows "p = (T⇘ap) of (X_minus a)"
  unfolding taylor_expansion_def taylor_def
  using assms sub_assoc[of p "X_plus a" "X_minus a"] X_plus_closed[of a]  X_minus_closed[of a]
  by (metis X_sub plus_minus_sub taylor_expansion_def)

lemma taylor_eval:
  assumes "a  carrier R"
  assumes "f  carrier P"
  assumes "b  carrier R"
  shows "to_fun (T⇘af) b = to_fun f (b  a)"
  unfolding  taylor_expansion_def taylor_def
  using to_fun_sub[of "(X_plus a)" f b] to_fun_X_plus[of a b] 
        assms X_plus_closed[of a] by auto

lemma taylor_eval':
  assumes "a  carrier R"
  assumes "f  carrier P"
  assumes "b  carrier R"
  shows "to_fun f (b) = to_fun (T⇘af) (b  a) "
  unfolding taylor_expansion_def taylor_def
  using to_fun_sub[of "(X_minus a)" "T⇘af" b] to_fun_X_minus[of b a] 
        assms X_minus_closed[of a] 
  by (metis taylor_closed taylor_def taylor_id taylor_expansion_def to_fun_X_minus)

lemma(in UP_cring) degree_monom:
  assumes "a  carrier R"
  shows "degree (a UP R(X_poly R)[^]UP Rn) = (if a = 𝟬 then 0 else n)"
  apply(cases "a = 𝟬")
  apply (metis (full_types) P.nat_pow_closed P_def R.one_closed UP_smult_zero X_poly_def deg_zero monom_closed)
  using P_def UP_cring.monom_rep_X_pow UP_cring_axioms assms deg_monom by fastforce  

lemma(in UP_cring) poly_comp_finsum:
  assumes "i::nat. i  n  g i  carrier P"
  assumes "q  carrier P"
  assumes "p = (Pi  {..n}. g i)"
  shows "p of q = (Pi  {..n}. (g i) of q)"  
proof- 
  have 0: "p of q  = rev_sub q p"
    unfolding compose_def rev_compose_def by blast 
  have 1: "p of q = finsum P (rev_compose R q  g) {..n}" 
    unfolding 0 unfolding assms 
    apply(rule ring_hom_finsum[of "rev_compose R q" P "{..n}" g ])
    using assms(2) rev_sub_is_hom apply blast
    apply (simp add: UP_ring)
     apply simp
  by (simp add: assms(1))
  show ?thesis unfolding 1 
    unfolding comp_apply rev_compose_def compose_def
    by auto 
qed

lemma(in UP_cring) poly_comp_expansion:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "degree p  n"
  shows "p of q = (Pi  {..n}. (p i) Pq[^]Pi)"
proof-
  obtain g where g_def: "g = (λi. monom P (p i) i)"
    by blast 
  have 0: "i. (g i) of q = (p i) Pq[^]Pi"
  proof- fix i show "g i of q = p i Pq [^]Pi"
      using assms g_def  P_def coeff_simp monom_sub 
      by (simp add: cfs_closed)           
  qed
  have 1: "(i. i  n  g i  carrier P)"
    using g_def assms 
    by (simp add: cfs_closed)   
  have  "(Pi{..n}. monom P (p i) i) = p"
    using assms up_repr_le[of p n] coeff_simp[of p] unfolding P_def 
    by auto     
  then have  "p = (Pi  {..n}. g i)"
    using g_def by auto       
  then have "p of q = (Pi{..n}. g i of q)" 
    using 0 1 poly_comp_finsum[of n g q p]
    using assms(2) 
    by blast
  then show ?thesis 
    by(simp add: 0)
qed

lemma(in UP_cring) taylor_sum:
  assumes "p  carrier P"
  assumes "degree p  n"
  assumes "a  carrier R"
  shows "p = (Pi  {..n}. T⇘ap i P(X_minus a)[^]Pi)"
proof-
  have 0: "(T⇘ap) of X_minus a = p" 
    using P_def taylor_id assms(1) assms(3) 
    by fastforce
  have 1: "degree (T⇘ap)  n"
    using assms 
    by (simp add: taylor_deg)
  have 2: "T⇘ap of X_minus a = (Pi{..n}. T⇘ap i PX_minus a [^]Pi)"
    using 1 X_minus_closed[of a]  poly_comp_expansion[of "T⇘ap" "X_minus a" n] 
          assms  taylor_closed 
    by blast
  then show ?thesis
    using 0 
    by simp
qed

text‹The $i^{th}$ term in the taylor expansion›
definition taylor_term where
"taylor_term c p i = (taylor_expansion R c p i) UP R(UP_cring.X_minus R c) [^]UP Ri" 

lemma (in UP_cring) taylor_term_closed:
assumes "p  carrier P"
assumes "a  carrier R"
shows "taylor_term a p i  carrier (UP R)"
 unfolding taylor_term_def 
  using P.nat_pow_closed P_def taylor_closed taylor_def X_minus_closed assms(1) assms(2) smult_closed 
  by (simp add: cfs_closed)

lemma(in UP_cring) taylor_term_sum:
  assumes "p  carrier P"
  assumes "degree p  n"
  assumes "a  carrier R"
  shows "p = (Pi  {..n}. taylor_term a p i)"
  unfolding taylor_term_def taylor_def
  using assms taylor_sum[of p n a] P_def 
  using taylor_def by auto
  
lemma (in UP_cring) taylor_expansion_add:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "c  carrier R"
  shows "taylor_expansion R c (p UP Rq) = (taylor_expansion R c p) UP R(taylor_expansion R c q)"
  unfolding taylor_expansion_def 
  using assms X_plus_closed[of c] P_def sub_add 
  by blast

lemma (in UP_cring) taylor_term_add:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "a  carrier R"
  shows "taylor_term a (p UP Rq) i = taylor_term a p i UP Rtaylor_term a q i"
  using assms   taylor_expansion_add[of p q a]  
  unfolding taylor_term_def 
  using P.nat_pow_closed P_def taylor_closed X_minus_closed cfs_add smult_l_distr 
  by (simp add: taylor_def cfs_closed)
  
lemma (in UP_cring) to_fun_taylor_term:
assumes "p  carrier P"
assumes "a  carrier R"
assumes "c  carrier R"
shows "to_fun (taylor_term c p i) a = (T⇘cp i)  (a  c)[^]i"
  using assms to_fun_smult[of "X_minus c [^]UP Ri" a "taylor_expansion R c p i"] 
              to_fun_X_minus[of c a] to_fun_nat_pow[of "X_minus c" a i]  
  unfolding taylor_term_def 
  using P.nat_pow_closed P_def taylor_closed taylor_def X_minus_closed 
  by (simp add: cfs_closed)    


end

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Defining the (Scalar-Valued) Derivative of a Polynomial Using the Taylor Expansion›
(**************************************************************************************************)
(**************************************************************************************************)
definition derivative where
"derivative R f a = (taylor_expansion R a f) 1"

context UP_cring
begin 

abbreviation(in UP_cring) deriv where
"deriv  derivative R"

lemma(in UP_cring) deriv_closed:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "(derivative R f a)  carrier R"
  unfolding derivative_def  
  using taylor_closed taylor_def assms(1) assms(2) cfs_closed by auto

lemma(in UP_cring) deriv_add:
  assumes "f  carrier P"
  assumes "g  carrier P"
  assumes "a  carrier R"
  shows "deriv (f Pg) a = deriv f a  deriv g a"
  unfolding derivative_def taylor_expansion_def  using assms 
  by (simp add: X_plus_closed sub_add sub_closed)

end 
(**************************************************************************************************)
(**************************************************************************************************)
section‹The Polynomial-Valued Derivative Operator›
(**************************************************************************************************)
(**************************************************************************************************)

context UP_cring
begin

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Operator Which Shifts Coefficients›
  (**********************************************************************)
  (**********************************************************************)
lemma cfs_times_X:
  assumes "g  carrier P"
  shows "(X Pg) (Suc n) = g n"
  apply(rule poly_induct3[of g])
  apply (simp add: assms)
  apply (metis (no_types, lifting) P.m_closed P.r_distr X_closed cfs_add)
  by (metis (no_types, lifting) P_def R.l_one R.one_closed R.r_null Suc_eq_plus1 X_poly_def
      cfs_monom coeff_monom_mult coeff_simp monom_closed monom_mult)

lemma times_X_pow_coeff:
  assumes "g  carrier P"
  shows "(monom P 𝟭 k Pg) (n + k) = g n"
  using coeff_monom_mult P.m_closed P_def assms coeff_simp monom_closed 
  by (simp add: cfs_closed)

lemma zcf_eq_zero_unique:
  assumes "f  carrier P"
  assumes "g  carrier P  (f = X Pg)"
  shows " h. h   carrier P  (f = X Ph)  h = g"
proof-
  fix h
  assume A: "h   carrier P  (f = X Ph)"
  then have 0: " X Pg =  X Ph"
    using assms(2) by auto
 show "h = g"
   using 0 A assms 
   by (metis P_def coeff_simp  cfs_times_X up_eqI)
qed

lemma f_minus_ctrm:
  assumes "f  carrier P"
  shows "zcf(f Pctrm f) = 𝟬"
  using assms 
  by (smt ctrm_is_poly P.add.inv_closed P.minus_closed P_def R.r_neg R.zero_closed zcf_to_fun 
        to_fun_minus to_fun_plus UP_cring_axioms zcf_ctrm zcf_def a_minus_def cfs_closed)

definition poly_shift where
"poly_shift f n = f (Suc n)"
   
lemma poly_shift_closed:
  assumes "f  carrier P"
  shows "poly_shift f  carrier P"
  apply(rule UP_car_memI[of "deg R f"])
  unfolding poly_shift_def 
proof -
  fix n :: nat
  assume "deg R f < n"
  then have "deg R f < Suc n"
    using Suc_lessD by blast
  then have "f (Suc n) = 𝟬P(Suc n)"
    by (metis P.l_zero UP_zero_closed assms coeff_of_sum_diff_degree0)
  then show "f (Suc n) = 𝟬"
    by simp
next
  show " n. f (Suc n)  carrier R"
    by(rule cfs_closed, rule assms)
qed

lemma poly_shift_eq_0:
  assumes "f  carrier P"
  shows "f  n = (ctrm f PX Ppoly_shift f) n"
  apply(cases "n = 0")
  apply (smt ctrm_degree ctrm_is_poly ltrm_of_X One_nat_def P.r_null P.r_zero P_def UP_cring.lcf_monom(1) UP_cring_axioms UP_mult_closed UP_r_one UP_zero_closed X_closed zcf_ltrm_mult zcf_def zcf_zero assms cfs_add cfs_closed deg_zero degree_X lessI monom_one poly_shift_closed to_poly_inverse)
proof- assume A: "n  0"
  then obtain k where k_def: " n = Suc k"
    by (meson lessI less_Suc_eq_0_disj)
  show ?thesis 
    using  cfs_times_X[of "poly_shift f" k] poly_shift_def[of f k] poly_shift_closed assms
        cfs_add[of "ctrm f" "X Ppoly_shift f" n] unfolding k_def 
    by (simp add: X_closed cfs_closed cfs_monom)
qed   
        
lemma poly_shift_eq:
  assumes "f  carrier P"
  shows "f = (ctrm f PX Ppoly_shift f)"
by(rule ext, rule poly_shift_eq_0, rule assms)

lemma poly_shift_id:
  assumes "f  carrier P"
  shows "f Pctrm f = X Ppoly_shift f"
  using assms poly_shift_eq[of f] poly_shift_closed unfolding a_minus_def 
  by (metis ctrm_is_poly P.add.inv_solve_left P.m_closed UP_a_comm UP_a_inv_closed X_closed)

lemma poly_shift_degree_zero:
  assumes "p  carrier P"
  assumes "degree p = 0"
  shows "poly_shift p = 𝟬P⇙"
  by (metis ltrm_deg_0 P.r_neg P.r_null UP_ring UP_zero_closed X_closed zcf_eq_zero_unique 
      abelian_group.minus_eq assms(1) assms(2) poly_shift_closed poly_shift_id ring_def)

lemma poly_shift_degree:
  assumes "p  carrier P"
  assumes "degree p > 0"
  shows "degree (poly_shift p) = degree p - 1 "
  using poly_shift_id[of p] 
  by (metis ctrm_degree ctrm_is_poly P.r_null X_closed add_diff_cancel_right' assms(1) assms(2) 
      deg_zero degree_of_difference_diff_degree degree_times_X nat_less_le poly_shift_closed)
  
lemma poly_shift_monom:
  assumes "a  carrier R"
  shows "poly_shift (monom P a (Suc k)) = (monom P a k)"
proof-
  have "(monom P a (Suc k)) = ctrm (monom P a (Suc k)) PX Ppoly_shift (monom P a (Suc k))"
    using poly_shift_eq[of "monom P a (Suc k)"] assms monom_closed 
    by blast
  then have "(monom P a (Suc k)) = 𝟬PPX Ppoly_shift (monom P a (Suc k))"
    using assms by simp
  then have "(monom P a (Suc k)) = X Ppoly_shift (monom P a (Suc k))"
    using X_closed assms poly_shift_closed by auto  
  then have "X P(monom P a k) = X Ppoly_shift (monom P a (Suc k))"
    by (metis P_def R.l_one R.one_closed X_poly_def assms monom_mult plus_1_eq_Suc)
  then show ?thesis 
    using X_closed X_not_zero assms 
    by (meson UP_mult_closed zcf_eq_zero_unique monom_closed poly_shift_closed)        
qed

lemma(in UP_cring) poly_shift_add:
  assumes "f  carrier P"
  assumes "g  carrier P"
  shows "poly_shift (f Pg) = (poly_shift f) P(poly_shift g)"
  apply(rule ext)
  using cfs_add[of "poly_shift f" "poly_shift g"] poly_shift_closed poly_shift_def 
  by (simp add: poly_shift_def assms(1) assms(2))

lemma(in UP_cring) poly_shift_s_mult:
  assumes "f  carrier P"
  assumes "s  carrier R"
  shows "poly_shift (s Pf) = s P(poly_shift f)"
proof-
  have "(s Pf) = (ctrm (s Pf)) P(X Ppoly_shift (s Pf))"
    using poly_shift_eq[of "(s Pf)"] assms(1) assms(2) 
    by blast
  then have 0: "(s Pf) = (s P(ctrm f)) P(X Ppoly_shift (s Pf))"
    using ctrm_smult assms(1) assms(2) by auto   
  have 1: "(s Pf) = s P((ctrm f) P(X P(poly_shift f)))"
    using assms(1) poly_shift_eq by auto
  have 2:  "(s Pf) = (s P(ctrm f)) P(s P(X P(poly_shift f)))"
    by (simp add: "1" X_closed assms(1) assms(2) ctrm_is_poly poly_shift_closed smult_r_distr)    
  have 3:  "(s Pf) = (s P(ctrm f)) P(X P(s P(poly_shift f)))"
    using "2" UP_m_comm X_closed assms(1) assms(2) smult_assoc2
    by (simp add: poly_shift_closed)    
  have 4: "(X Ppoly_shift (s Pf)) =  (X P(s P(poly_shift f)))"
    using 3 0  X_closed assms(1) assms(2) ctrm_is_poly poly_shift_closed 
    by auto     
  then show ?thesis 
    using X_closed X_not_zero assms(1) assms(2) 
    by (metis UP_mult_closed UP_smult_closed zcf_eq_zero_unique poly_shift_closed)  
qed

lemma zcf_poly_shift:
  assumes "f   carrier P"
  shows "zcf (poly_shift f) = f 1"  
  apply(rule poly_induct3)
  apply (simp add: assms)
  using poly_shift_add zcf_add cfs_add poly_shift_closed apply metis 
  unfolding zcf_def using poly_shift_monom poly_shift_degree_zero
  by (simp add: poly_shift_def)

fun poly_shift_iter ("shift") where
Base:"poly_shift_iter 0 f = f"|
Step:"poly_shift_iter (Suc n) f = poly_shift (poly_shift_iter n f)"

lemma shift_closed:
  assumes "f  carrier P"
  shows "shift n f  carrier P"
  apply(induction n)
  using assms poly_shift_closed by auto 

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Operator Which Multiplies Coefficients by Their Degree›
  (**********************************************************************)
  (**********************************************************************)

definition n_mult where 
"n_mult f = (λn. [n]R(f n))"

lemma(in UP_cring) n_mult_closed:
  assumes "f  carrier P"
  shows "n_mult f  carrier P"
  apply(rule UP_car_memI[of "deg R f"])
  unfolding n_mult_def 
  apply (metis P.l_zero R.add.nat_pow_one UP_zero_closed assms cfs_zero coeff_of_sum_diff_degree0)
  using assms cfs_closed by auto 

text‹Facts about the shift function›

lemma shift_one:
"shift (Suc 0) = poly_shift"
  by auto 

lemma shift_factor0:
  assumes "f  carrier P"
  shows "degree f  (Suc k)  degree (f P((shift (Suc k) f) P(X[^]P(Suc k)))) < (Suc k)"
proof(induction k)
  case 0
  have 0: " f P(ctrm f) =  (shift (Suc 0) f)PX"
    by (metis UP_m_comm X_closed assms poly_shift_id shift_closed shift_one)         
  then have  " f P(shift (Suc 0) f)PX =    (ctrm f) "
  proof-
    have " f P(ctrm f) P(shift (Suc 0) f)PX=  (shift (Suc 0) f)PX P(shift (Suc 0) f)PX"
      using 0   by simp
    then have " f P(ctrm f) P(shift (Suc 0) f)PX = 𝟬P⇙"
      using UP_cring.UP_cring[of R] assms 
      by (metis "0" P.ring_simprules(4) P_def UP_ring.UP_ring UP_ring_axioms 
          a_minus_def abelian_group.r_neg ctrm_is_poly ring_def)
    then have " f P((ctrm f) P(shift (Suc 0) f)PX) = 𝟬P⇙"
      using assms P.ring_simprules  
      by (metis "0" poly_shift_id poly_shift_eq)
    then have " f P((shift (Suc 0) f)PX P(ctrm f) ) = 𝟬P⇙"
      using P.m_closed UP_a_comm X_closed assms ctrm_is_poly shift_closed 
      by presburger      
    then have "f P((shift (Suc 0) f)PX) P(ctrm f)= 𝟬P⇙"
      using P.add.m_assoc P.ring_simprules(14) P.ring_simprules(19) assms "0" 
          P.add.inv_closed P.r_neg P.r_zero ctrm_is_poly
      by smt      
    then show ?thesis 
      by (metis "0" P.add.m_comm P.m_closed P.ring_simprules(14) P.ring_simprules(18) 
          P.ring_simprules(3) X_closed assms ctrm_is_poly poly_shift_id poly_shift_eq 
          shift_closed)
  qed
  then have  " f P(shift (Suc 0) f)P(X[^]P(Suc 0)) =    (ctrm f) "
  proof-
    have "X = X[^]P(Suc 0)"
      by (simp add: X_closed)
    then show ?thesis 
      using 0 f Pshift (Suc 0) f PX = ctrm f 
      by auto   
  qed
  then have " degree (f P(shift (Suc 0) f)P(X[^]P(Suc 0))) < 1"
    using ctrm_degree[of f] assms   by simp
  then show ?case 
    by blast
next
  case (Suc n)
  fix k
  assume IH: "degree f  (Suc k)  degree (f P((shift (Suc k) f) P(X[^]P(Suc k)))) < (Suc k)"
  show  "degree f  (Suc (Suc k))  degree (f P((shift (Suc (Suc k)) f) P(X[^]P(Suc (Suc k))))) < (Suc (Suc k))"
  proof-
    obtain n where n_def: "n = Suc k"
      by simp
    have IH': "degree f  n  degree (f P((shift n f) P(X[^]Pn))) < n"
      using n_def IH by auto 
    have P: "degree f  (Suc n)  degree (f P((shift (Suc n) f) P(X[^]P(Suc n)))) < (Suc n)"
    proof-
      obtain g where g_def: "g = (f P((shift n f) P(X[^]Pn)))"
        by simp
      obtain s where s_def: "s = shift n f"
        by simp
      obtain s' where s'_def: "s' = shift (Suc n) f"
        by simp
      have P: "g  carrier P" "s  carrier P" "s'  carrier P" "(X[^]Pn)  carrier P"
        using s_def s'_def g_def assms shift_closed[of f n]  
        apply (simp add: X_closed)
        apply (simp add: f  carrier P  shift n f  carrier P assms s_def)
        using P_def UP_cring.shift_closed UP_cring_axioms assms s'_def apply blast
        using X_closed by blast
      have g_def': "g = (f P(s P(X[^]Pn)))"
        using g_def s_def by auto 
      assume "degree f  (Suc n)"
      then have " degree (f P(s P(X[^]Pn))) < n" 
        using IH' Suc_leD s_def by blast
      then have d_g: "degree g < n" using g_def' by auto 
      have P0: "f P(s' P(X[^]P(Suc n))) = ((ctrm s)P(X[^]Pn)) Pg"
      proof-
        have "s = (ctrm s) P(X Ps')"
          using s_def s'_def  P_def poly_shift_eq UP_cring_axioms assms shift_closed 
          by (simp add: UP_cring.poly_shift_eq)
        then have 0: "g = f P((ctrm s) P(X Ps')) P(X[^]Pn)"
          using g_def'  by auto
        then have "g = f P((ctrm s)P(X[^]Pn)) P((X Ps') P(X[^]Pn))"
          using P cring_axioms X_closed  P.l_distr P.ring_simprules(19) UP_a_assoc a_minus_def assms
          by (simp add: a_minus_def ctrm_is_poly)
        then have "g P((X Ps') P(X[^]Pn)) =  f P((ctrm s)P(X[^]Pn))"
          using P cring_axioms X_closed  P.l_distr P.ring_simprules UP_a_assoc a_minus_def assms
          by (simp add: P.r_neg2 ctrm_is_poly)                                                   
        then have " ((ctrm s)P(X[^]Pn)) =  f P(g P((X Ps') P(X[^]Pn)))"
          using P cring_axioms X_closed P.ring_simprules UP_a_assoc a_minus_def assms
          by (simp add: P.ring_simprules(17) ctrm_is_poly)           
        then have " ((ctrm s)P(X[^]Pn)) =  f P(((X Ps') P(X[^]Pn)) Pg)"
          by (simp add: P(1) P(3) UP_a_comm X_closed)          
        then have "((ctrm s)P(X[^]Pn)) =  f P((X Ps') P(X[^]Pn)) Pg"
          using P(1) P(3) P.ring_simprules(19) UP_a_assoc a_minus_def assms
          by (simp add: a_minus_def X_closed) 
        then have "((ctrm s)P(X[^]Pn)) Pg=  f P((X Ps') P(X[^]Pn))"
          by (metis P(1) P(3) P(4) P.add.inv_solve_right P.m_closed P.ring_simprules(14)
              P.ring_simprules(4) P_def UP_cring.X_closed UP_cring_axioms assms)
        then have "((ctrm s)P(X[^]Pn)) Pg=  f P((s' PX) P(X[^]Pn))"
          by (simp add: P(3) UP_m_comm X_closed)         
        then have "((ctrm s)P(X[^]Pn)) Pg=  f P(s' P(X[^]P(Suc n)))"
          using P(3) P.nat_pow_Suc2 UP_m_assoc X_closed by auto
        then show ?thesis
          by auto 
      qed
      have P1: "degree (((ctrm s)P(X[^]Pn)) Pg)  n"
      proof-
        have Q0: "degree ((ctrm s)P(X[^]Pn))   n"
        proof(cases "ctrm s = 𝟬P⇙")
          case True
          then show ?thesis 
            by (simp add: P(4))
        next
          case False
          then have F0: "degree ((ctrm s)P(X[^]Pn))  degree (ctrm s) + degree (X[^]Pn) "
            by (meson ctrm_is_poly P(2) P(4) deg_mult_ring)             
          have F1: "𝟭𝟬 degree (X[^]Pn) = n"
            unfolding X_poly_def 
            using P_def cring_monom_degree by auto
          show ?thesis 
            by (metis (no_types, opaque_lifting) F0 F1 ltrm_deg_0 P(2) P.r_null P_def R.l_null R.l_one 
                R.nat_pow_closed R.zero_closed X_poly_def assms cfs_closed 
                add_0 deg_const deg_zero deg_ltrm 
                monom_pow monom_zero zero_le)                                    
        qed
        then show ?thesis 
          using d_g
          by (simp add: P(1) P(2) P(4) bound_deg_sum ctrm_is_poly) 
      qed
      then show ?thesis 
        using s'_def P0 by auto
    qed
    assume "degree f  (Suc (Suc k)) "
    then show "degree (f P((shift (Suc (Suc k)) f) P(X[^]P(Suc (Suc k))))) < (Suc (Suc k))"
      using P by(simp add: n_def)
  qed
qed

lemma(in UP_cring) shift_degree0:
  assumes "f  carrier P"
  shows "degree f >n  Suc (degree (shift (Suc n) f)) = degree (shift n f)"
proof(induction n) 
  case 0
  assume B: "0< degree f"
  have 0: "degree (shift 0 f) = degree f"
    by simp
  have 1: "degree f = degree (f P(ctrm f))"
    using assms(1) B ctrm_degree degree_of_difference_diff_degree
    by (simp add: ctrm_is_poly)   
  have "(f P(ctrm f)) = X P(shift 1 f)"
    using P_def poly_shift_id UP_cring_axioms assms(1) by auto
  then have "degree (f P(ctrm f)) = 1 + (degree (shift 1 f))"
    by (metis "1" B P.r_null X_closed add.commute assms deg_nzero_nzero degree_times_X not_gr_zero shift_closed)        
  then have  "degree (shift 0 f) = 1 + (degree (shift 1 f))"
    using 0 1 by auto 
  then show ?case 
    by simp
next
  case (Suc n)
  fix n 
  assume IH: "(n < degree f  Suc (degree (shift (Suc n) f)) = degree (shift n f))"
  show "Suc n < degree f  Suc (degree (shift (Suc (Suc n)) f)) = degree (shift (Suc n) f)"
  proof-
    assume A: " Suc n < degree f"
    then have 0: "(shift (Suc n) f) = ctrm ((shift (Suc n) f)) P(shift (Suc (Suc n)) f)PX"
      by (metis UP_m_comm X_closed assms local.Step poly_shift_eq shift_closed)                
    have N: "(shift (Suc (Suc n)) f)  𝟬P⇙"
    proof
      assume C: "shift (Suc (Suc n)) f = 𝟬P⇙"
      obtain g where g_def: "g = f P(shift (Suc (Suc n)) f)P(X[^]P(Suc (Suc n)))"
        by simp
      have C0: "degree g < degree f"
        using g_def assms A 
        by (meson Suc_leI Suc_less_SucD Suc_mono less_trans_Suc shift_factor0)
      have C1: "g = f" 
        using C
        by (simp add: P.minus_eq X_closed assms g_def)          
      then show False 
        using C0 by auto 
    qed
    have 1: "degree (shift (Suc n) f) = degree ((shift (Suc n) f) Pctrm ((shift (Suc n) f)))"
    proof(cases "degree (shift (Suc n) f) = 0")
      case True
      then show ?thesis 
        using N assms poly_shift_degree_zero poly_shift_closed shift_closed by auto        
    next
      case False
      then have "degree (shift (Suc n) f) > degree  (ctrm ((shift (Suc n) f)))"
      proof -
        have "shift (Suc n) f  carrier P"
          using assms shift_closed by blast
        then show ?thesis
          using False ctrm_degree by auto
      qed        
      then show ?thesis
      proof -
        show ?thesis
          using degree (ctrm (shift (Suc n) f)) < degree (shift (Suc n) f) 
            assms ctrm_is_poly degree_of_difference_diff_degree shift_closed by presburger
      qed         
    qed
    have 2: "(shift (Suc n) f) Pctrm ((shift (Suc n) f)) =  (shift (Suc (Suc n)) f)PX"
      using 0
      by (metis Cring_Poly.INTEG.Step P.m_comm X_closed assms poly_shift_id shift_closed)                                   
    have 3: "degree ((shift (Suc n) f) Pctrm ((shift (Suc n) f))) =  degree (shift (Suc (Suc n)) f) + 1"
      using 2 N X_closed X_not_zero assms  degree_X shift_closed 
      by (metis UP_m_comm degree_times_X)
    then show ?thesis using 1
      by linarith
  qed
qed

lemma(in UP_cring) shift_degree:
  assumes "f  carrier P"
  shows "degree f  n   degree (shift n f) + n = degree f"
proof(induction n)
  case 0
  then show ?case 
    by auto
next
  case (Suc n)
  fix n
  assume IH: "(n  degree f  degree (shift n f) + n = degree f)"
  show "Suc n  degree f  degree (shift (Suc n) f) + Suc n = degree f"
  proof-
    assume A: "Suc n  degree f "
    have 0: "degree (shift n f) + n = degree f" 
      using IH A by auto  
    have 1: "degree (shift n f) = Suc (degree (shift (Suc n) f))"
      using A assms shift_degree0 by auto 
    show "degree (shift (Suc n) f) + Suc n = degree f" 
      using 0 1  by simp
  qed
qed

lemma(in UP_cring) shift_degree':
  assumes "f  carrier P"
  shows "degree (shift (degree f) f) = 0"
  using shift_degree assms 
  by fastforce

lemma(in UP_cring) shift_above_degree:
  assumes "f  carrier P"
  assumes "k > degree f"
  shows "(shift k f) = 𝟬P⇙"
proof-
  have "n. shift ((degree f)+ (Suc n)) f = 𝟬P⇙"
  proof-
    fix n
    show "shift ((degree f)+ (Suc n)) f = 𝟬P⇙"
    proof(induction n)
      case 0
      have B0:"shift (degree f) f  = ctrm(shift (degree f) f) P(shift (degree f + Suc 0) f)PX"
      proof -
        have f1: "f n. f  carrier P  shift n f  carrier P"
          by (meson shift_closed)
        then have "shift (degree f + Suc 0) f  carrier P"
          using assms(1) by blast
        then show ?thesis
          using f1 by (simp add: P.m_comm X_closed assms(1) poly_shift_eq)
      qed        
      have B1:"shift (degree f) f = ctrm(shift (degree f) f)"
      proof -
        have "shift (degree f) f  carrier P"
          using assms(1) shift_closed by blast
        then show ?thesis
          using ltrm_deg_0 assms(1) shift_degree' by auto          
      qed        
      have B2: "(shift (degree f + Suc 0) f)PX = 𝟬P⇙"
        using B0 B1 X_closed assms(1)
      proof -
        have "f n. f  carrier P  shift n f  carrier P"
          using shift_closed by blast
        then show ?thesis
          by (metis (no_types) B0 B1 P.add.l_cancel_one UP_mult_closed X_closed assms(1))
      qed         
      then show ?case 
        by (metis P.r_null UP_m_comm UP_zero_closed X_closed assms(1) zcf_eq_zero_unique shift_closed)        
    next
      case (Suc n)
      fix n
      assume "shift (degree f + Suc n) f = 𝟬P⇙"
      then show "shift (degree f + Suc (Suc n)) f = 𝟬P⇙"
        by (simp add: poly_shift_degree_zero)
    qed
  qed
  then show ?thesis 
    using assms(2) less_iff_Suc_add by auto
qed

lemma(in UP_domain) shift_cfs0:
  assumes "f  carrier P"
  shows "zcf(shift 1 f) = f 1"
  using assms 
  by (simp add: zcf_poly_shift) 

lemma(in UP_cring) X_mult_cf:
  assumes "p  carrier P"
  shows "(p PX) (k+1) = p k"
  unfolding X_poly_def
  using assms 
  by (metis UP_m_comm X_closed X_poly_def add.commute plus_1_eq_Suc cfs_times_X)
 
lemma(in UP_cring) X_pow_cf:
  assumes "p  carrier P"
  shows "(p PX[^]P(n::nat)) (n + k) = p k"
proof-
  have P: "f. f  carrier P  (f PX[^]P(n::nat)) (n + k) = f k"
  proof(induction n)
    show "f. f  carrier P  (f PX [^]P(0::nat)) (0 + k) = f k"
    proof-
      fix f
      assume B0: "f  carrier P"
      show "(f PX [^]P(0::nat)) (0 + k) = f k"
        by (simp add: B0)
    qed
    fix n
    fix f
    assume IH: "(f. f  carrier P  (f PX [^]Pn) (n + k) = f k)"
    assume A0: " f  carrier P"
    show "(f PX [^]PSuc n) (Suc n + k) = f k"
    proof-
      have 0: "(f PX [^]Pn)(n + k) = f k"
        using A0 IH by simp
      have 1: "((f PX [^]Pn)PX) (Suc n + k) = (f PX [^]Pn)(n + k)"
        using X_mult_cf A0 P.m_closed P.nat_pow_closed 
              Suc_eq_plus1 X_closed add_Suc by presburger
      have 2: "(f P(X [^]Pn PX)) (Suc n + k) = (f PX [^]Pn)(n + k)"
        using 1
        by (simp add: A0 UP_m_assoc X_closed)         
      then show ?thesis 
        by (simp add: "0")
    qed
  qed
  show ?thesis using assms P[of p] by auto 
qed

lemma poly_shift_cfs:
  assumes "f  carrier P"
  shows "poly_shift f n = f (Suc n)"
proof-
  have "(f Pctrm f) (Suc n) = (X P(poly_shift f)) (Suc n)"
    using assms poly_shift_id by auto
  then show ?thesis unfolding X_poly_def using poly_shift_closed assms 
    by (metis (no_types, lifting) ctrm_degree ctrm_is_poly 
        P.add.m_comm P.minus_closed coeff_of_sum_diff_degree0 poly_shift_id poly_shift_eq cfs_times_X zero_less_Suc)
qed

lemma(in UP_cring) shift_cfs:
  assumes "p  carrier P"
  shows "(shift k p) n = p (k + n)"
  apply(induction k arbitrary: n)
  by (auto simp: assms poly_shift_cfs shift_closed)

  (**********************************************************************)
  (**********************************************************************)
  subsection‹The Derivative Operator›
  (**********************************************************************)
  (**********************************************************************)
definition pderiv where
"pderiv  p = poly_shift (n_mult p)"

lemma pderiv_closed:
  assumes "p  carrier P"
  shows "pderiv p  carrier P"
  unfolding  pderiv_def 
  using assms n_mult_closed[of p]  poly_shift_closed[of "n_mult p"] 
  by blast

text‹Function which obtains the first n+1 terms of f, in ascending order of degree:›

definition trms_of_deg_leq  where
"trms_of_deg_leq n f  f (UP R)((shift (Suc n) f) UP Rmonom P 𝟭 (Suc n))"

lemma trms_of_deg_leq_closed:
  assumes "f  carrier P"
  shows "trms_of_deg_leq n f  carrier P"
  unfolding trms_of_deg_leq_def using assms 
  by (metis P.m_closed P.minus_closed P_def R.one_closed monom_closed shift_closed)

lemma trms_of_deg_leq_id:
  assumes "f  carrier P"
  shows "f P(trms_of_deg_leq k f) = shift (Suc k) f Pmonom P 𝟭 (Suc k)"
  unfolding trms_of_deg_leq_def 
  using assms 
  by (smt P.add.inv_closed P.l_zero P.m_closed P.minus_add P.minus_minus P.r_neg
      P_def R.one_closed UP_a_assoc a_minus_def monom_closed shift_closed)

lemma trms_of_deg_leq_id':
  assumes "f  carrier P"
  shows "f = (trms_of_deg_leq k f)  Pshift (Suc k) f Pmonom P 𝟭 (Suc k)"
  using trms_of_deg_leq_id assms  trms_of_deg_leq_closed[of f]
  by (smt P.add.inv_closed P.l_zero P.m_closed P.minus_add P.minus_minus P.r_neg R.one_closed UP_a_assoc a_minus_def monom_closed shift_closed)

lemma deg_leqI:
  assumes "p  carrier P"
  assumes "n. n > k  p n = 𝟬"
  shows "degree p  k"
  by (metis assms(1) assms(2) deg_zero deg_ltrm le0 le_less_linear monom_zero)

lemma deg_leE:
  assumes "p  carrier P"
  assumes "degree p < k"
  shows "p k = 𝟬"
  using assms  coeff_of_sum_diff_degree0 P_def coeff_simp deg_aboveD 
  by auto

lemma trms_of_deg_leq_deg:
  assumes "f  carrier P"
  shows "degree (trms_of_deg_leq k f)  k"
proof-
  have "n. (trms_of_deg_leq k f) (Suc k + n) = 𝟬"
  proof-
    fix n
    have 0: "(shift (Suc k) f UP Rmonom P 𝟭 (Suc k)) (Suc k + n) = shift (Suc k) f n"
      using assms shift_closed cfs_monom_mult_l 
      by (metis P.m_comm P_def R.one_closed add.commute monom_closed times_X_pow_coeff)
    then show "trms_of_deg_leq k f (Suc k + n) = 𝟬" 
      unfolding trms_of_deg_leq_def
      using shift_cfs[of f "Suc k" n] 
            cfs_minus[of f "shift (Suc k) f UP Rmonom P 𝟭 (Suc k)" "Suc k + n"]
      by (metis P.m_closed P.r_neg P_def R.one_closed a_minus_def assms 
          cfs_minus cfs_zero monom_closed shift_closed)
  qed
  then show ?thesis using deg_leqI 
    by (metis (no_types, lifting) assms le_iff_add less_Suc_eq_0_disj less_Suc_eq_le trms_of_deg_leq_closed)
qed
  
lemma trms_of_deg_leq_zero_is_ctrm:
  assumes "f  carrier P"
  assumes "degree f > 0"
  shows "trms_of_deg_leq 0 f = ctrm f"
proof-
  have "f = ctrm f P(X P(shift (Suc 0) f))"
    using assms poly_shift_eq 
    by simp
  then have "f = ctrm f P(X [^]UP RSuc 0 P(shift (Suc 0) f))"
    using P.nat_pow_eone P_def X_closed  by auto
  then show ?thesis 
  unfolding trms_of_deg_leq_def 
  by (metis (no_types, lifting) ctrm_is_poly One_nat_def P.add.right_cancel P.m_closed 
      P.minus_closed P.nat_pow_eone P_def UP_m_comm X_closed X_poly_def assms(1) shift_closed
      trms_of_deg_leq_def trms_of_deg_leq_id')
qed

lemma cfs_monom_mult:
  assumes "p  carrier P"
  assumes "a  carrier R"
  assumes "k < n"
  shows "(p P(monom P a n)) k = 𝟬"
  apply(rule poly_induct3[of p])
    apply (simp add: assms(1))
   apply (metis (no_types, lifting) P.l_distr P.m_closed R.r_zero R.zero_closed assms(2) cfs_add monom_closed)
  using assms monom_mult[of _ a _ n] 
  by (metis R.m_closed R.m_comm add.commute cfs_monom not_add_less1)

lemma(in UP_cring) cfs_monom_mult_2:
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "m < n"
  shows "((monom P a n) Pf) m = 𝟬"
  using cfs_monom_mult
  by (simp add: P.m_comm assms(1) assms(2) assms(3))

lemma trms_of_deg_leq_cfs:
  assumes "f  carrier P"
  shows "trms_of_deg_leq n f k = (if k  n then (f k) else 𝟬)"
  unfolding trms_of_deg_leq_def 
  apply(cases "k  n")
  using cfs_minus[of f "shift (Suc n) f UP Rmonom P 𝟭 (Suc n)"] 
        cfs_monom_mult[of _ 𝟭 k "Suc n"] 
   apply (metis (no_types, lifting) P.m_closed P.minus_closed P_def R.one_closed R.r_zero assms
      cfs_add cfs_closed le_refl monom_closed nat_less_le nat_neq_iff not_less_eq_eq shift_closed 
      trms_of_deg_leq_def trms_of_deg_leq_id')
  using trms_of_deg_leq_deg[of f n] deg_leE   
  unfolding trms_of_deg_leq_def 
  using assms trms_of_deg_leq_closed trms_of_deg_leq_def by auto
 
lemma trms_of_deg_leq_iter:
  assumes "f  carrier P"
  shows "trms_of_deg_leq (Suc k) f = (trms_of_deg_leq k f) Pmonom P (f (Suc k)) (Suc k)"
proof fix x
  show "trms_of_deg_leq (Suc k) f x = (trms_of_deg_leq k f Pmonom P (f (Suc k)) (Suc k)) x"
    apply(cases "x  k")
    using trms_of_deg_leq_cfs trms_of_deg_leq_closed cfs_closed[of f "Suc k"]
          cfs_add[of "trms_of_deg_leq k f" "monom P (f (Suc k)) (Suc k)" x]   
     apply (simp add: assms)
    using deg_leE assms cfs_closed cfs_monom apply auto[1]
    by (simp add: assms cfs_closed cfs_monom trms_of_deg_leq_cfs trms_of_deg_leq_closed)   
qed 

lemma trms_of_deg_leq_0:
  assumes "f  carrier P"
  shows "trms_of_deg_leq 0 f = ctrm f"
  by (metis One_nat_def P.r_null P_def UP_m_comm UP_zero_closed X_closed X_poly_def assms not_gr_zero
      poly_shift_degree_zero shift_one trms_of_deg_leq_def trms_of_deg_leq_zero_is_ctrm trunc_simps(2) trunc_zero)
  
lemma trms_of_deg_leq_degree_f:
  assumes "f   carrier P"
  shows "trms_of_deg_leq (degree f) f = f"
proof fix x 
  show "trms_of_deg_leq (deg R f) f x = f x"
    using assms trms_of_deg_leq_cfs deg_leE[of f x] 
  by simp
qed

definition(in UP_cring) lin_part where
"lin_part f = trms_of_deg_leq 1 f"

lemma(in UP_cring) lin_part_id:
  assumes "f  carrier P"
  shows "lin_part f = (ctrm f) Pmonom P (f 1) 1"
  unfolding lin_part_def
  by (simp add: assms trms_of_deg_leq_0 trms_of_deg_leq_iter)

lemma(in UP_cring) lin_part_eq:
  assumes "f  carrier P"
  shows "f = lin_part f P(shift 2 f) Pmonom P 𝟭 2"
  unfolding lin_part_def 
  by (metis Suc_1 assms trms_of_deg_leq_id')

text‹Constant term of a substitution:›

lemma zcf_eval:
  assumes "f  carrier P"
  shows "zcf f = to_fun f 𝟬"
  using assms zcf_to_fun by blast

lemma ctrm_of_sub:
  assumes "f  carrier P"
  assumes "g  carrier P"
  shows "zcf(f of g) = to_fun f (zcf g)"
  apply(rule poly_induct3[of f]) 
  apply (simp add: assms(1))
  using P_def UP_cring.to_fun_closed UP_cring_axioms zcf_add zcf_to_fun assms(2) to_fun_plus sub_add sub_closed apply fastforce
  using R.zero_closed zcf_to_fun assms(2) to_fun_sub monom_closed sub_closed by presburger
 
text‹Evaluation of linear part:›

lemma to_fun_lin_part:
  assumes "f  carrier P"
  assumes "b  carrier R"
  shows "to_fun (lin_part f) b = (f 0)  (f 1)  b"
  using assms lin_part_id[of f] to_fun_ctrm to_fun_monom monom_closed 
   by (simp add: cfs_closed to_fun_plus) 

text‹Constant term of taylor expansion:›

lemma taylor_zcf:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "zcf(T⇘af) = to_fun f a"
  unfolding taylor_expansion_def
  using ctrm_of_sub assms P_def zcf_eval X_plus_closed taylor_closed taylor_eval by auto

lemma(in UP_cring) taylor_eq_1:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "(T⇘af) P(trms_of_deg_leq 1 (T⇘af)) = (shift (2::nat) (T⇘af))P(X[^]P(2::nat))"
  by (metis P.nat_pow_eone P.nat_pow_mult P_def Suc_1 taylor_closed X_closed X_poly_def assms(1) 
      assms(2) monom_one_Suc2 one_add_one trms_of_deg_leq_id)

lemma(in UP_cring) taylor_deg_1:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "f of (X_plus a) = (lin_part (T⇘af)) P(shift (2::nat) (T⇘af))P(X[^]P(2::nat))"
  using taylor_eq_1[of f a]
  unfolding taylor_expansion_def lin_part_def
  using One_nat_def X_plus_closed assms(1) 
        assms(2) trms_of_deg_leq_id' numeral_2_eq_2 sub_closed 
  by (metis P.nat_pow_Suc2 P.nat_pow_eone P_def taylor_def X_closed X_poly_def monom_one_Suc taylor_expansion_def)

lemma(in UP_cring) taylor_deg_1_eval:
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "c = to_fun (shift (2::nat) (T⇘af)) b"
  assumes "fa = to_fun f a"
  assumes "f'a = deriv f a"
  shows "to_fun f (b  a) = fa  (f'a  b)  (c  b[^](2::nat))"
  using assms taylor_deg_1 unfolding derivative_def   
proof-
  have 0: "to_fun f (b  a) = to_fun (f of (X_plus a)) b"
    using to_fun_sub assms X_plus_closed by auto    
  have 1: "to_fun (lin_part (T⇘af)) b = fa  (f'a  b) "
    using assms to_fun_lin_part[of "(T⇘af)" b]  
    by (metis P_def taylor_def UP_cring.taylor_zcf UP_cring.taylor_closed UP_cring_axioms zcf_def derivative_def)    
  have 2: "(T⇘af) = (lin_part (T⇘af)) P((shift 2 (T⇘af))PX[^]P(2::nat))"
    using lin_part_eq[of "(T⇘af)"] assms(1) assms(2) taylor_closed 
    by (metis taylor_def taylor_deg_1 taylor_expansion_def)        
  then have "to_fun (T⇘af) b = fa  (f'a  b)  to_fun ((shift 2 (T⇘af))PX[^]P(2::nat)) b"
    using 1 2 
    by (metis P.nat_pow_closed taylor_closed UP_mult_closed X_closed assms(1) assms(2) assms(3) 
        to_fun_plus lin_part_def shift_closed trms_of_deg_leq_closed)    
  then have  "to_fun (T⇘af) b = fa  (f'a  b)  c  to_fun (X[^]P(2::nat)) b"
    by (simp add: taylor_closed X_closed assms(1) assms(2) assms(3) assms(4) to_fun_mult shift_closed)    
  then have 3: "to_fun f (b  a)= fa  (f'a  b)  c  to_fun (X[^]P(2::nat)) b"
    using taylor_eval assms(1) assms(2) assms(3) by auto
  have "to_fun (X[^]P(2::nat)) b = b[^](2::nat)"
    by (metis P.nat_pow_Suc2 P.nat_pow_eone R.nat_pow_Suc2 
        R.nat_pow_eone Suc_1 to_fun_X 
        X_closed assms(3) to_fun_mult)
  then show ?thesis 
    using 3 by auto 
qed

lemma(in UP_cring) taylor_deg_1_eval':
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "c = to_fun (shift (2::nat) (T⇘af)) b"
  assumes "fa = to_fun f a"
  assumes "f'a = deriv f a"
  shows "to_fun f (a  b) = fa  (f'a  b)  (c  b[^](2::nat))"
  using R.add.m_comm taylor_deg_1_eval assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) 
  by auto

lemma(in UP_cring) taylor_deg_1_eval'':
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "c = to_fun (shift (2::nat) (T⇘af)) (b)"
  shows "to_fun f (a  b) = (to_fun f a)  (deriv f a  b)  (c  b[^](2::nat))"
proof-
  have "b  carrier R"
    using assms 
    by blast
  then have 0: "to_fun f (a  b) = (to_fun f a) (deriv f a  (b))  (c  (b)[^](2::nat))"
  unfolding a_minus_def
  using taylor_deg_1_eval'[of f a "b" c "(to_fun f a)" "deriv f a"] assms
  by auto 
  have 1: " (deriv f a  b) = (deriv f a  (b))"
    using assms 
    by (simp add: R.r_minus deriv_closed)
  have 2: "(c  b[^](2::nat)) = (c  (b)[^](2::nat))"
    using assms 
    by (metis R.add.inv_closed R.add.inv_solve_right R.l_zero R.nat_pow_Suc2 
        R.nat_pow_eone R.zero_closed Suc_1 UP_ring_axioms UP_ring_def
        ring.ring_simprules(26) ring.ring_simprules(27))
  show ?thesis 
    using 0 1 2 
    unfolding a_minus_def
    by simp 
qed

lemma(in UP_cring) taylor_deg_1_expansion:
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "c = to_fun (shift (2::nat) (T⇘af)) (b  a)"
  assumes "fa = to_fun f a"
  assumes "f'a = deriv f a"
  shows "to_fun f (b) = fa  f'a  (b  a)  (c  (b  a)[^](2::nat))"
proof-
  obtain b' where b'_def: "b'= b  a "
    by simp
  then have b'_def': "b = b'  a"
    using assms 
    by (metis R.add.inv_solve_right R.minus_closed R.minus_eq)
  have "to_fun f (b'  a) = fa  (f'a  b')  (c  b'[^](2::nat))" 
    using assms taylor_deg_1_eval[of f a b' c fa f'a]  b'_def 
    by blast
  then have "to_fun f (b) = fa  (f'a  b')  (c  b'[^](2::nat))" 
    using b'_def' 
    by auto 
  then show "to_fun f (b) = fa  f'a  (b  a)  c  (b  a) [^] (2::nat)"
    using b'_def
    by auto   
qed

lemma(in UP_cring) Taylor_deg_1_expansion':
  assumes "f  carrier (UP R)"
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "c  carrier R. to_fun f (b) = (to_fun f a)  (deriv f a)  (b  a)  (c  (b  a)[^](2::nat))"
  using taylor_deg_1_expansion[of f a b] assms unfolding P_def 
  by (metis P_def R.minus_closed taylor_closed shift_closed to_fun_closed)


text‹Basic Properties of deriv and pderiv:›

lemma n_mult_degree_bound:
  assumes "f  carrier P"
  shows "degree (n_mult f)  degree f"
  apply(rule deg_leqI)
  apply (simp add: assms n_mult_closed)
  by (simp add: assms deg_leE n_mult_def)

lemma pderiv_deg_0[simp]:
  assumes "f  carrier P"
  assumes "degree f = 0"
  shows "pderiv f = 𝟬P⇙"
proof-
  have "degree (n_mult f) = 0"
    using P_def n_mult_degree_bound  assms(1) assms(2) by fastforce
  then show ?thesis 
    unfolding pderiv_def
    by (simp add: assms(1) n_mult_closed poly_shift_degree_zero)    
qed

lemma deriv_deg_0:
  assumes "f  carrier P"
  assumes "degree f = 0"
  assumes "a  carrier R"
  shows "deriv f a = 𝟬"
  unfolding derivative_def taylor_expansion_def
  using X_plus_closed assms(1) assms(2) assms(3) deg_leE sub_const by force

lemma poly_shift_monom':
  assumes "a  carrier R"
  shows "poly_shift (a P(X[^]P(Suc n))) = aP(X[^]Pn)"
  using assms monom_rep_X_pow poly_shift_monom by auto

lemma monom_coeff:
  assumes "a  carrier R"
  shows "(a PX [^]P(n::nat)) k = (if (k = n) then a else 𝟬)"
  using assms cfs_monom monom_rep_X_pow by auto

lemma cfs_n_mult:
  assumes "p  carrier P"
  shows "n_mult p n = [n](p n)"
  by (simp add: n_mult_def)

lemma cfs_add_nat_pow:
  assumes "p  carrier P"
  shows "([(n::nat)]Pp) k = [n](p k)"
  apply(induction n) by (auto simp: assms)

lemma cfs_add_int_pow:
  assumes "p  carrier P"
  shows "([(n::int)]Pp) k = [n](p k)"
  apply(induction n)
  by(auto simp: add_pow_int_ge assms cfs_add_nat_pow add_pow_int_lt)

lemma add_nat_pow_monom:
  assumes "a  carrier R"
  shows "[(n::nat)]Pmonom P a k = monom P ([n]a) k"
  apply(rule ext) 
  by (simp add: assms cfs_add_nat_pow cfs_monom)
  
lemma add_int_pow_monom:
  assumes "a  carrier R"
  shows "[(n::int)]Pmonom P a k = monom P ([n]a) k"
  apply(rule ext) 
  by (simp add: assms cfs_add_int_pow cfs_monom)

lemma n_mult_monom:
  assumes "a  carrier R"
  shows "n_mult (monom P a (Suc n)) = monom P ([Suc n]a) (Suc n)"
  apply(rule ext)
  unfolding n_mult_def 
  using assms cfs_monom by auto

lemma pderiv_monom:
  assumes "a  carrier R"
  shows "pderiv (monom P a n) = monom P ([n]a) (n-1)"
  apply(cases "n = 0")
   apply (simp add: assms)
  unfolding pderiv_def 
  using assms Suc_diff_1[of n] n_mult_monom[of a "n-1"] poly_shift_monom[of "[Suc (n-1)]a" "Suc (n-1)"]
  by (metis R.add.nat_pow_closed neq0_conv poly_shift_monom)

lemma pderiv_monom':
  assumes "a  carrier R"
  shows "pderiv (a PX[^]P(n::nat)) = ([n]a)PX[^]P(n-1)"  
  using assms pderiv_monom[of a n ] 
  by (simp add: P_def UP_cring.monom_rep_X_pow UP_cring_axioms)

lemma n_mult_add:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "n_mult (p Pq) = n_mult p Pn_mult q"
proof(rule ext) fix x show "n_mult (p Pq) x = (n_mult p Pn_mult q) x"
  using assms R.add.nat_pow_distrib[of "p x" "q x" x]  cfs_add[of p q x] 
        cfs_add[of "n_mult p" "n_mult q" x] n_mult_closed
  unfolding n_mult_def 
  by (simp add: cfs_closed) 
qed

lemma pderiv_add:
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "pderiv (p Pq) = pderiv p Ppderiv q"
  unfolding pderiv_def 
  using assms poly_shift_add n_mult_add 
  by (simp add: n_mult_closed)

lemma zcf_monom_sub:
  assumes "p  carrier P"
  shows "zcf ((monom P 𝟭 (Suc n)) of  p) = zcf p [^] (Suc n)"
  apply(induction n)
  using One_nat_def P.nat_pow_eone R.nat_pow_eone R.one_closed R.zero_closed zcf_to_fun assms to_fun_closed monom_sub smult_one apply presburger
  using P_def UP_cring.ctrm_of_sub UP_cring_axioms zcf_to_fun assms to_fun_closed to_fun_monom monom_closed 
  by fastforce

lemma zcf_monom_sub':
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "zcf ((monom P a (Suc n)) of  p) = a  zcf p [^] (Suc n)"
  using zcf_monom_sub assms P_def R.zero_closed UP_cring.ctrm_of_sub UP_cring.to_fun_monom UP_cring_axioms
    zcf_to_fun to_fun_closed monom_closed by fastforce
  
lemma deriv_monom:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "deriv (monom P a n) b =  ([n]a)(b[^](n-1))"  
proof(induction n)
  case 0
  have 0: "b [^] ((0::nat) - 1)  carrier R"
    using assms 
    by simp
  then show ?case unfolding derivative_def using assms  
    by (smt One_nat_def P_def R.add.nat_pow_0 R.nat_pow_Suc2 R.nat_pow_eone R.zero_closed 
        taylor_def taylor_deg UP_cring.taylor_closed UP_cring.zcf_monom UP_cring.shift_one 
        UP_cring_axioms zcf_degree_zero zcf_zero_degree_zero degree_monom monom_closed 
        monom_rep_X_pow plus_1_eq_Suc poly_shift_degree_zero shift_cfs to_fun_monom to_fun_zero zero_diff)
next
  case (Suc n)
  show ?case
  proof(cases "n = 0")
    case True
    have T0: "[Suc n]  a  b [^] (Suc n - 1) = a"
      by (simp add: True assms(1))
    have T1: "(X_poly R UP Rto_polynomial R b) [^]UP RSuc n = X_poly R UP Rto_polynomial R b "
      using P.nat_pow_eone P_def True UP_a_closed X_closed assms(2) to_poly_closed by auto
    then show ?thesis 
      unfolding derivative_def taylor_expansion_def
      using T0 T1 True sub_monom(2)[of "X_plus b" a "Suc n"] cfs_add assms
      unfolding P_def X_poly_plus_def to_polynomial_def X_poly_def  
      by (smt Group.nat_pow_0 lcf_eq lcf_monom(2) ltrm_of_X_plus One_nat_def P_def R.one_closed 
          R.r_one R.r_zero UP_cring.zcf_monom UP_cring.degree_of_X_plus 
          UP_cring.poly_shift_degree_zero UP_cring_axioms X_closed X_plus_closed X_poly_def 
          X_poly_plus_def zcf_zero_degree_zero cfs_monom_mult_l degree_to_poly to_fun_X_pow 
          plus_1_eq_Suc poly_shift_cfs poly_shift_monom to_poly_closed to_poly_mult_simp(2)
          to_poly_nat_pow to_polynomial_def)
  next
    case False
  have "deriv (monom P a (Suc n)) b = ((monom P a (Suc n)) of (X_plus b)) 1"
    unfolding derivative_def taylor_expansion_def 
    by auto
  then have "deriv (monom P a (Suc n)) b = (((monom P a n) of (X_plus b)) P(X_plus b)) 1"
    using  monom_mult[of a 𝟭 n 1] sub_mult[of "X_plus b" "monom P a n" "monom P 𝟭 1" ]  X_plus_closed[of b] assms   
    by (metis lcf_monom(1) P.l_one P.nat_pow_eone P_def R.one_closed R.r_one Suc_eq_plus1 
        deg_one monom_closed monom_one sub_monom(1) to_poly_inverse)
  then have "deriv (monom P a (Suc n)) b = (((monom P a n) of (X_plus b)) P(monom P 𝟭 1) P(((monom P a n) of (X_plus b)) Pto_poly b)) 1"
    unfolding X_poly_plus_def 
    by (metis P.r_distr P_def X_closed X_plus_closed X_poly_def X_poly_plus_def assms(1) assms(2) monom_closed sub_closed to_poly_closed)
  then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0  b  ((monom P a n) of (X_plus b)) 1"
    unfolding X_poly_plus_def 
    by (smt One_nat_def P.m_closed P_def UP_m_comm X_closed X_plus_closed X_poly_def X_poly_plus_def
        assms(1) assms(2) cfs_add cfs_monom_mult_l monom_closed plus_1_eq_Suc sub_closed cfs_times_X to_polynomial_def)
  then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0  b  (deriv (monom P a n) b)"
    by (simp add: derivative_def taylor_expansion_def)
  then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0  b  ( ([n]a)(b[^](n-1)))"
    by (simp add: Suc)
  then have 0: "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0  ([n]a)(b[^]n)"
    using assms R.m_comm[of b] R.nat_pow_mult[of b "n-1" 1] False 
    by (metis (no_types, lifting) R.add.nat_pow_closed R.m_lcomm R.nat_pow_closed R.nat_pow_eone add.commute add_eq_if plus_1_eq_Suc)
  have 1: "((monom P a n) of (X_plus b)) 0 = a  b[^]n"
     unfolding X_poly_plus_def using zcf_monom_sub' 
     by (smt ctrm_of_sub One_nat_def P_def R.l_zero R.one_closed UP_cring.zcf_to_poly 
         UP_cring.f_minus_ctrm UP_cring_axioms X_plus_closed X_poly_def X_poly_plus_def zcf_add
         zcf_def assms(1) assms(2) to_fun_monom monom_closed monom_one_Suc2 poly_shift_id poly_shift_monom to_poly_closed)
   show ?thesis 
     using 0 1  R.add.nat_pow_Suc2 R.add.nat_pow_closed R.l_distr R.nat_pow_closed assms(1) assms(2) diff_Suc_1 by presburger
 qed
qed

lemma deriv_smult:
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "g  carrier P"
  shows "deriv (a Pg) b = a  (deriv g b)"
  unfolding derivative_def taylor_expansion_def 
  using assms sub_smult X_plus_closed cfs_smult
  by (simp add: sub_closed)

lemma deriv_const:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "deriv (monom P a 0) b = 𝟬"
  unfolding derivative_def 
  using assms taylor_closed taylor_def taylor_deg deg_leE by auto 

lemma deriv_monom_deg_one:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "deriv (monom P a 1) b = a"
  unfolding derivative_def taylor_expansion_def 
  using assms cfs_X_plus[of b 1] sub_monom_deg_one X_plus_closed[of b]
  by simp  

lemma monom_Suc:
  assumes "a  carrier R"
  shows "monom P a (Suc n) = monom P 𝟭 1 Pmonom P a n"
        "monom P a (Suc n) = monom P a n Pmonom P 𝟭 1"
  apply (metis R.l_one R.one_closed Suc_eq_plus1_left assms monom_mult)
  by (metis R.one_closed R.r_one Suc_eq_plus1 assms monom_mult)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹The Product Rule›
(**************************************************************************************************)
(**************************************************************************************************)

lemma(in UP_cring) times_x_product_rule: 
  assumes "f  carrier P"
  shows "pderiv (f Pup_ring.monom P 𝟭 1) =  f Ppderiv f Pup_ring.monom P 𝟭 1"
proof(rule poly_induct3[of f])
  show "f  carrier P"
    using assms by blast 
  show "p q. q  carrier P 
           p  carrier P 
           pderiv (p Pup_ring.monom P 𝟭 1) = p Ppderiv p Pup_ring.monom P 𝟭 1 
           pderiv (q Pup_ring.monom P 𝟭 1) = q Ppderiv q Pup_ring.monom P 𝟭 1 
           pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = p Pq Ppderiv (p Pq) Pup_ring.monom P 𝟭 1"
  proof- fix p q assume A: "q  carrier P"
           "p  carrier P"
           "pderiv (p Pup_ring.monom P 𝟭 1) = p Ppderiv p Pup_ring.monom P 𝟭 1"
           "pderiv (q Pup_ring.monom P 𝟭 1) = q Ppderiv q Pup_ring.monom P 𝟭 1"
    have 0: "(p Pq) Pup_ring.monom P 𝟭 1 =  (p Pup_ring.monom P 𝟭 1) P(q Pup_ring.monom P 𝟭 1)"
      using A assms  by (meson R.one_closed UP_l_distr is_UP_monomE(1) is_UP_monomI)
    have  1: "pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = pderiv (p Pup_ring.monom P 𝟭 1) Ppderiv (q Pup_ring.monom P 𝟭 1)"
      unfolding 0 apply(rule pderiv_add) 
      using A is_UP_monomE(1) monom_is_UP_monom(1) apply blast
      using A is_UP_monomE(1) monom_is_UP_monom(1) by blast
    have 2: "pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = p Ppderiv p Pup_ring.monom P 𝟭 1 P(q Ppderiv q Pup_ring.monom P 𝟭 1)"
      unfolding 1 A  by blast 
    have 3: "pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = p Pq P(pderiv p Pup_ring.monom P 𝟭 1 Ppderiv q Pup_ring.monom P 𝟭 1)"
      unfolding 2 
      using A P.add.m_lcomm R.one_closed UP_a_assoc UP_a_closed UP_mult_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed by presburger
    have 4: "pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = p Pq P((pderiv p Ppderiv q) Pup_ring.monom P 𝟭 1)"
      unfolding 3 using A P.l_distr R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed by presburger
    show 5: "pderiv ((p Pq) Pup_ring.monom P 𝟭 1) = p Pq Ppderiv (p Pq) Pup_ring.monom P 𝟭 1"
      unfolding 4 using pderiv_add A by presburger
  qed
  show "a n. a  carrier R 
           pderiv (up_ring.monom P a n Pup_ring.monom P 𝟭 1) = up_ring.monom P a n Ppderiv (up_ring.monom P a n) Pup_ring.monom P 𝟭 1"
  proof- fix a n assume A: "a  carrier R"
    have 0: "up_ring.monom P a n Pup_ring.monom P 𝟭 1 = up_ring.monom P a (Suc n)"
      using A monom_Suc(2) by presburger
    have 1: "pderiv (up_ring.monom P a n Pup_ring.monom P 𝟭 1) =  [(Suc n)] P(up_ring.monom P a n)"
      unfolding 0 using A add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom 
      by (simp add: P_def)
    have 2: "pderiv (up_ring.monom P a n Pup_ring.monom P 𝟭 1) = (up_ring.monom P a n) P[n] P(up_ring.monom P a n)"
      unfolding 1 using A P.add.nat_pow_Suc2 is_UP_monomE(1) monom_is_UP_monom(1) by blast
    have 3: "pderiv (up_ring.monom P a n) Pup_ring.monom P 𝟭 1 =  [n] P(up_ring.monom P a n)"
      apply(cases "n = 0")
      using A  add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom pderiv_deg_0 apply auto[1]
      using monom_Suc(2)[of a "n-1"] A  add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom
      by (metis R.add.nat_pow_closed Suc_eq_plus1 add_eq_if monom_Suc(2) pderiv_monom)     
    show "pderiv (up_ring.monom P a n Pup_ring.monom P 𝟭 1) = up_ring.monom P a n Ppderiv (up_ring.monom P a n) Pup_ring.monom P 𝟭 1"
      unfolding 2 3 by blast 
  qed
qed

lemma(in UP_cring) deg_one_eval:
  assumes "g  carrier (UP R)"
  assumes "deg R g = 1"
  shows "t. t  carrier R  to_fun g t = g 0  (g 1)t"
proof-
  obtain h where h_def: "h = ltrm g"
    by blast 
  have 0: "deg R (g UP Rh) = 0"
    using assms unfolding h_def 
    by (metis ltrm_closed ltrm_eq_imp_deg_drop ltrm_monom P_def UP_car_memE(1) less_one)
  have 1: "g UP Rh = to_poly (g 0)"
  proof(rule ext) fix x show "(g UP Rh) x = to_polynomial R (g 0) x"
    proof(cases "x = 0")
      case True
      have T0: "h 0 = 𝟬"
        unfolding h_def using assms UP_car_memE(1) cfs_monom by presburger
      have T1: "(g UP Rh) 0 = g 0  h 0"
        using ltrm_closed P_def assms(1) cfs_minus h_def by blast
      then show ?thesis using T0 assms  
        by (smt "0" ltrm_closed ltrm_deg_0 P.minus_closed P_def UP_car_memE(1) UP_zero_closed zcf_def zcf_zero deg_zero degree_to_poly h_def to_poly_closed to_poly_inverse to_poly_minus trunc_simps(2) trunc_zero)        
    next
      case False
      then have "x > 0"
        by presburger 
      then show ?thesis 
        by (metis "0" ltrm_closed P.minus_closed P_def UP_car_memE(1) UP_cring.degree_to_poly UP_cring_axioms assms(1) deg_leE h_def to_poly_closed)     
    qed
  qed
  have 2: "g = (g UP Rh) UP Rh"
    unfolding h_def using assms 
    by (metis "1" P_def h_def lin_part_def lin_part_id to_polynomial_def trms_of_deg_leq_degree_f)    
  fix t assume A: "t  carrier R"
  have 3: " to_fun g t = to_fun (g UP Rh) t  to_fun h t"
    using 2 
    by (metis "1" A P_def UP_car_memE(1) assms(1) h_def monom_closed to_fun_plus to_polynomial_def)
  then show "to_fun g t = g 0  g 1  t "
    unfolding 1 h_def 
    using A P_def UP_cring.lin_part_def UP_cring_axioms assms(1) assms(2) to_fun_lin_part trms_of_deg_leq_degree_f by fastforce
qed

lemma nmult_smult:
  assumes "a  carrier R"
  assumes "f  carrier P"
  shows "n_mult (a Pf) = a P(n_mult f)"
  apply(rule poly_induct4[of f])
  apply (simp add: assms(2))
  using assms(1) n_mult_add n_mult_closed smult_closed smult_r_distr apply presburger
  using assms  apply(intro ext, metis (no_types, lifting) ctrm_smult ltrm_deg_0 P_def R.add.nat_pow_0 UP_cring.ctrm_degree UP_cring.n_mult_closed UP_cring.n_mult_def UP_cring_axioms UP_smult_closed UP_zero_closed zcf_degree_zero zcf_zero deg_const deg_zero le_0_eq monom_closed n_mult_degree_bound smult_r_null)
  using monom_mult_smult n_mult_monom assms 
  by (smt lcf_monom(1) P_def R.add.nat_pow_closed R.add_pow_rdistr R.zero_closed UP_cring.to_poly_mult_simp(1) UP_cring_axioms UP_smult_closed cfs_closed cring_lcf_mult monom_closed to_polynomial_def)

lemma pderiv_smult:
  assumes "a  carrier R"
  assumes "f  carrier P"
  shows "pderiv (a Pf) = a P(pderiv f)"
  unfolding pderiv_def
  using assms
  by (simp add: n_mult_closed nmult_smult poly_shift_s_mult) 

lemma(in UP_cring) pderiv_minus:
  assumes "a  carrier P"
  assumes "b  carrier P"
  shows "pderiv (a Pb) = pderiv a Ppderiv b"
proof-
  have "Pb = (𝟭)Pb"
    using R.one_closed UP_smult_one assms(2) smult_l_minus by presburger
  thus ?thesis unfolding a_minus_def using pderiv_add assms pderiv_smult
    by (metis P.add.inv_closed R.add.inv_closed R.one_closed UP_smult_one pderiv_closed smult_l_minus)
qed

lemma(in UP_cring) pderiv_const:
  assumes "b  carrier R"
  shows "pderiv (up_ring.monom P b 0) = 𝟬P⇙"
  using assms pderiv_monom[of b 0] deg_const is_UP_monomE(1) monom_is_UP_monom(1) pderiv_deg_0 
  by blast
   
lemma(in UP_cring) pderiv_minus_const:
  assumes "a  carrier P"
  assumes "b  carrier R"
  shows "pderiv (a Pup_ring.monom P b 0) = pderiv a"
  using pderiv_minus[of a "up_ring.monom P b 0" ] assms pderiv_const[of b] 
  by (smt P.l_zero P.minus_closed P_def UP_cring.pderiv_const UP_cring.pderiv_minus UP_cring.poly_shift_eq UP_cring_axioms cfs_closed monom_closed pderiv_add pderiv_closed poly_shift_id)

lemma(in UP_cring) monom_product_rule: 
  assumes "f  carrier P"
  assumes "a  carrier R" 
  shows "pderiv (f Pup_ring.monom P a n) =  f Ppderiv (up_ring.monom P a n) Ppderiv f Pup_ring.monom P a n"
proof-
  have "f. f  carrier P  pderiv (f Pup_ring.monom P a n) =  f Ppderiv (up_ring.monom P a n) Ppderiv f Pup_ring.monom P a n"
  proof(induction n)
    case 0
    show ?case 
    proof fix f show "f  carrier P  pderiv (f Pup_ring.monom P a 0) = f Ppderiv (up_ring.monom P a 0) Ppderiv f Pup_ring.monom P a 0 "
      proof assume A: "f  carrier P"
        have 0: "f Pup_ring.monom P a 0 = a Pf"
          using assms A  UP_m_comm is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult by presburger
        have 1: "f Ppderiv (up_ring.monom P a 0) = 𝟬P⇙"
          using A assms P.r_null pderiv_const by presburger
        have 2: "pderiv f Pup_ring.monom P a 0 = a Ppderiv f"
          using assms A  UP_m_comm is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult  pderiv_closed by presburger
        show "pderiv (f Pup_ring.monom P a 0) = f Ppderiv (up_ring.monom P a 0) Ppderiv f Pup_ring.monom P a 0"
          unfolding 0 1 2 using A UP_l_zero UP_smult_closed assms(2) pderiv_closed pderiv_smult by presburger
      qed
    qed
  next
    case (Suc n)
    show "f. f  carrier P 
             pderiv (f Pup_ring.monom P a (Suc n)) = f Ppderiv (up_ring.monom P a (Suc n)) Ppderiv f Pup_ring.monom P a (Suc n)"
    proof fix f
      show "f  carrier P 
         pderiv (f Pup_ring.monom P a (Suc n)) = f Ppderiv (up_ring.monom P a (Suc n)) Ppderiv f Pup_ring.monom P a (Suc n)"
      proof
       assume A: "f  carrier P"
      show " pderiv (f Pup_ring.monom P a (Suc n)) = f Ppderiv (up_ring.monom P a (Suc n)) Ppderiv f Pup_ring.monom P a (Suc n)"
      proof(cases "n = 0")
        case True
        have 0: "(f Pup_ring.monom P a (Suc n)) = a Pf Pup_ring.monom P 𝟭 1"
        proof -
          have "n. up_ring.monom P a n  carrier P"
            using assms(2) is_UP_monomE(1) monom_is_UP_monom(1) by presburger
          then show ?thesis
            by (metis A P.m_assoc P.m_comm R.one_closed True assms(2) is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_mult_is_smult)
        qed
        have 1: "f Ppderiv (up_ring.monom P a (Suc n)) = a Pf"
          using assms True 
          by (metis A One_nat_def P.m_comm R.add.nat_pow_eone diff_Suc_1 is_UP_monomE(1) is_UP_monomI monom_mult_is_smult pderiv_monom)
        have 2: "pderiv f Pup_ring.monom P a (Suc n) = a P(pderiv f Pup_ring.monom P 𝟭 1)"
          using A assms unfolding True 
          by (metis P.m_lcomm R.one_closed UP_mult_closed is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_mult_is_smult pderiv_closed)
        have 3: "a Pf Pa P(pderiv f Pup_ring.monom P 𝟭 1) = a P(f P(pderiv f Pup_ring.monom P 𝟭 1))"
          using assms A  P.m_closed R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed smult_r_distr by presburger
        show ?thesis
          unfolding 0 1 2 3 
          using A times_x_product_rule P.m_closed R.one_closed UP_smult_assoc2 assms(2) is_UP_monomE(1) monom_is_UP_monom(1) pderiv_smult by presburger
      next
        case False
      have IH: "pderiv ((f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n) = (f Pup_ring.monom P 𝟭 1) Ppderiv (up_ring.monom P a n) Ppderiv (f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n"
        using Suc A P.m_closed R.one_closed is_UP_monomE(1) is_UP_monomI by presburger
      have 0: "f Pup_ring.monom P a (Suc n) = (f Pup_ring.monom P 𝟭 1)  Pup_ring.monom P a n"
        using A R.one_closed UP_m_assoc assms(2) is_UP_monomE(1) monom_Suc(1) monom_is_UP_monom(1) by presburger
      have 1: "(f Pup_ring.monom P 𝟭 1) Ppderiv (up_ring.monom P a n) Ppderiv (f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n = 
               (f Pup_ring.monom P 𝟭 1) Ppderiv (up_ring.monom P a n) P(f Ppderiv f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n "
        using A times_x_product_rule by presburger
      have 2: "(f Pup_ring.monom P 𝟭 1) Ppderiv (up_ring.monom P a n) =(f Pup_ring.monom P ([n]a) n)"
      proof-
        have 20: "up_ring.monom P ([n]  a) (n) = up_ring.monom P 𝟭 1 Pup_ring.monom P ([n]  a) (n - 1)"
          using A assms False monom_mult[of 𝟭 "[n]a" 1 "n-1"] 
          by (metis R.add.nat_pow_closed R.l_one R.one_closed Suc_eq_plus1 add.commute add_eq_if )          
        show ?thesis unfolding 20 using assms A False   pderiv_monom[of a n] 
          using P.m_assoc R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) by simp
      qed
      have 3: "(f Pup_ring.monom P ([n]a) n) = [n]P(f Pup_ring.monom P a n)"
        using A assms by (metis P.add_pow_rdistr add_nat_pow_monom is_UP_monomE(1) monom_is_UP_monom(1))
      have 4: "pderiv (f Pup_ring.monom P 𝟭 1) =  (f Ppderiv f Pup_ring.monom P 𝟭 1)"
        using times_x_product_rule  A by blast
      have 5: " (f Ppderiv f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n  = 
 (f Pup_ring.monom P a n ) P(pderiv f Pup_ring.monom P 𝟭 1 Pup_ring.monom P a n )"
        using A assms  by (meson P.l_distr P.m_closed R.one_closed is_UP_monomE(1) is_UP_monomI pderiv_closed)
      have 6: " (f Ppderiv f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n  = 
 (f Pup_ring.monom P a n ) P(pderiv f Pup_ring.monom P 𝟭 1 Pup_ring.monom P a n )"
        using A assms False 5  by blast
      have 7: "(f Pup_ring.monom P 𝟭 1) Ppderiv (up_ring.monom P a n) Ppderiv (f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n =
               [(Suc n)] P(f Pup_ring.monom P a n) Ppderiv f Pup_ring.monom P 𝟭 1 Pup_ring.monom P a n"
        unfolding 2 3  5 6 using assms A P.a_assoc 
        by (smt "1" "2" "3" "6" P.add.nat_pow_Suc P.m_closed R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed)
      have 8: "pderiv (f Pup_ring.monom P a (Suc n)) = pderiv ((f Pup_ring.monom P 𝟭 1) Pup_ring.monom P a n)"
        using A assms 0  by presburger
      show " pderiv (f Pup_ring.monom P a (Suc n)) = f Ppderiv (up_ring.monom P a (Suc n)) Ppderiv f Pup_ring.monom P a (Suc n)"
        unfolding 8 IH 0 1 2 3 4 5 6 
        by (smt "2" "4" "6" "7" A P.add_pow_rdistr R.one_closed UP_m_assoc add_nat_pow_monom assms(2) diff_Suc_1 is_UP_monomE(1) is_UP_monomI monom_Suc(1) pderiv_closed pderiv_monom)        
       qed
     qed
   qed
 qed
  thus ?thesis using assms by blast 
qed

lemma(in UP_cring) product_rule: 
  assumes "f  carrier (UP R)"
  assumes "g  carrier (UP R)"
  shows "pderiv (f UP Rg) = (pderiv f UP Rg) UP R(f UP Rpderiv g)"
proof(rule poly_induct3[of f])
  show "f  carrier P"
    using assms unfolding P_def by blast 
  show "p q. q  carrier P 
           p  carrier P 
           pderiv (p UP Rg) = pderiv p UP Rg UP Rp UP Rpderiv g 
           pderiv (q UP Rg) = pderiv q UP Rg UP Rq UP Rpderiv g 
           pderiv ((p Pq) UP Rg) = pderiv (p Pq) UP Rg UP R(p Pq) UP Rpderiv g"
  proof- fix p q 
    assume A: "q  carrier P" "p  carrier P"
           "pderiv (p UP Rg) = pderiv p UP Rg UP Rp UP Rpderiv g"
           "pderiv (q UP Rg) = pderiv q UP Rg UP Rq UP Rpderiv g"
    have 0: "(p Pq) UP Rg = p UP Rg UP Rq UP Rg"
      using A assms unfolding P_def using P_def UP_l_distr by blast
    have 1: "pderiv ((p Pq) UP Rg) = pderiv (p UP Rg) UP Rpderiv (q UP Rg)"
      unfolding 0  using pderiv_add[of "p Pg" "q Pg"] unfolding P_def 
      using A(1) A(2) P_def UP_mult_closed assms(2) by blast
    have 2: "pderiv ((p Pq) UP Rg) = pderiv p UP Rg UP Rp UP Rpderiv g UP R(pderiv q UP Rg UP Rq UP Rpderiv g)"
      unfolding 1 A by blast 
    have 3: "pderiv ((p Pq) UP Rg) = pderiv p UP Rg UP Rpderiv q UP Rg  UP Rp UP Rpderiv g  UP Rq UP Rpderiv g"
      using A assms 
      by (smt "2" P.add.m_lcomm P.m_closed P_def UP_a_assoc pderiv_closed)
    have 4: "pderiv ((p Pq) UP Rg) = (pderiv p UP Rg UP Rpderiv q UP Rg)  UP R(p UP Rpderiv g  UP Rq UP Rpderiv g)"
      unfolding 3 using A assms P_def UP_a_assoc UP_a_closed UP_mult_closed pderiv_closed by auto
    have 5: "pderiv ((p Pq) UP Rg) = ((pderiv p  UP Rpderiv q) UP Rg)  UP R((p UP Rq)  UP Rpderiv g)"
      unfolding 4 using A assms  by (metis P.l_distr P_def pderiv_closed)
    have 6: "pderiv ((p Pq) UP Rg) = ((pderiv (p Pq)) UP Rg)  UP R((p UP Rq)  UP Rpderiv g)"
      unfolding 5 using A assms  
      by (metis P_def pderiv_add)
    show "pderiv ((p Pq) UP Rg) = pderiv (p Pq) UP Rg UP R(p Pq) UP Rpderiv g"
      unfolding 6 using A assms P_def by blast     
  qed
  show "a n. a  carrier R 
           pderiv (up_ring.monom P a n UP Rg) = pderiv (up_ring.monom P a n) UP Rg UP Rup_ring.monom P a n UP Rpderiv g"
    using P_def UP_m_comm assms(2) is_UP_monomE(1) monom_is_UP_monom(1) monom_product_rule pderiv_closed by presburger
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹The Chain Rule›
(**************************************************************************************************)
(**************************************************************************************************)

lemma(in UP_cring) chain_rule: 
  assumes "f  carrier P"
  assumes "g  carrier P"
  shows "pderiv (compose R f g) = compose R (pderiv f) g UP Rpderiv g"
proof(rule poly_induct3[of f])
  show "f  carrier P"
    using assms by blast
  show "p q. q  carrier P 
           p  carrier P 
           pderiv (Cring_Poly.compose R p g) = Cring_Poly.compose R (pderiv p) g UP Rpderiv g 
           pderiv (Cring_Poly.compose R q g) = Cring_Poly.compose R (pderiv q) g UP Rpderiv g 
           pderiv (Cring_Poly.compose R (p Pq) g) = Cring_Poly.compose R (pderiv (p Pq)) g UP Rpderiv g"
  using pderiv_add sub_add 
  by (smt P_def UP_a_closed UP_m_comm UP_r_distr assms(2) pderiv_closed sub_closed)
  show "a n. a  carrier R 
           pderiv (compose R (up_ring.monom P a n) g) = compose R (pderiv (up_ring.monom P a n)) g UP Rpderiv g"
  proof-
    fix a n assume A: "a  carrier R"
    show "pderiv (compose R (up_ring.monom P a n) g) = compose R (pderiv (up_ring.monom P a n)) g UP Rpderiv g"
    proof(induction n)
      case 0
      have 00: "(compose R (up_ring.monom P a 0) g) = (up_ring.monom P a 0)"
        using A P_def assms(2) deg_const is_UP_monom_def monom_is_UP_monom(1) sub_const by presburger
      have 01: "pderiv (up_ring.monom P a 0) = 𝟬P⇙"
        using A pderiv_const by blast
      show ?case unfolding 00 01 
        by (metis P.l_null P_def UP_zero_closed assms(2) deg_zero pderiv_closed sub_const)
    next
      case (Suc n)
      show "pderiv (Cring_Poly.compose R (up_ring.monom P a (Suc n)) g) = Cring_Poly.compose R (pderiv (up_ring.monom P a (Suc n))) g UP Rpderiv g"
      proof(cases "n = 0")
        case True
        have 0: "compose R (up_ring.monom P a (Suc n)) g = a Pg"
          using A assms sub_monom_deg_one[of g a] unfolding True using One_nat_def 
          by presburger
        have 1: "(pderiv (up_ring.monom P a (Suc n))) = up_ring.monom P a 0"
          unfolding True 
        proof -
          have "pderiv (up_ring.monom P a 0) = 𝟬P⇙"
            using A pderiv_const by blast
          then show "pderiv (up_ring.monom P a (Suc 0)) = up_ring.monom P a 0"
            using A lcf_monom(1) P_def  X_closed deg_const deg_nzero_nzero  is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_rep_X_pow pderiv_monom poly_shift_degree_zero poly_shift_eq sub_monom(2) sub_monom_deg_one to_poly_inverse to_poly_mult_simp(2)
            by (metis (no_types, lifting) P.l_null P.r_zero X_poly_def times_x_product_rule)
        qed
        then show ?thesis unfolding 0 1 
          using A P_def assms(2) deg_const is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult pderiv_closed pderiv_smult sub_const 
          by presburger
      next
        case False
      have 0: "compose R (up_ring.monom P a (Suc n)) g = (compose R (up_ring.monom P a n) g) P(compose R (up_ring.monom P 𝟭 1) g)"
        using assms A by (metis R.one_closed monom_Suc(2) monom_closed sub_mult)
      have 1: "compose R (up_ring.monom P a (Suc n)) g = (compose R (up_ring.monom P a n) g) Pg"
        unfolding 0 using A assms 
        by (metis P_def R.one_closed UP_cring.lcf_monom(1) UP_cring.to_poly_inverse UP_cring_axioms UP_l_one UP_one_closed deg_one monom_one sub_monom_deg_one to_poly_mult_simp(1))
      have 2: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) =
                    ((pderiv (compose R (up_ring.monom P a n) g)) Pg) P((compose R (up_ring.monom P a n) g) Ppderiv g)"
        unfolding 1 unfolding P_def apply(rule product_rule)
        using A assms unfolding P_def using P_def is_UP_monomE(1) is_UP_monomI rev_sub_closed sub_rev_sub apply presburger
        using assms unfolding P_def  by blast 
      have 3: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = 
                    (compose R (pderiv (up_ring.monom P a n)) g UP Rpderiv g Pg) P((compose R (up_ring.monom P a n) g) Ppderiv g)"
        unfolding 2 Suc by blast 
      have 4: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = 
                    ((compose R (pderiv (up_ring.monom P a n)) g  Pg) UP Rpderiv g) P((compose R (up_ring.monom P a n) g) Ppderiv g)"
        unfolding 3 using A assms m_assoc m_comm 
        by (smt P_def monom_closed monom_rep_X_pow pderiv_closed sub_closed)
      have 5: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = 
                    ((compose R (pderiv (up_ring.monom P a n)) g  Pg) P(compose R (up_ring.monom P a n) g)) Ppderiv g"
        unfolding 4 using A assms 
        by (metis P.l_distr P.m_closed P_def UP_cring.pderiv_closed UP_cring_axioms monom_closed sub_closed)
      have 6: "compose R (pderiv (up_ring.monom P a n)) g  Pg = [n]Pcompose R ((up_ring.monom P a n)) g"
      proof-
        have 60: "(pderiv (up_ring.monom P a n)) = (up_ring.monom P ([n]a) (n-1))" 
          using A assms pderiv_monom by blast
        have 61: "compose R (pderiv (up_ring.monom P a n)) g Pg  = compose R ((up_ring.monom P ([n]a) (n-1))) g P(compose R (up_ring.monom P 𝟭 1) g)"
          unfolding 60 using A assms sub_monom_deg_one[of g 𝟭 ] R.one_closed smult_one by presburger
        have 62: "compose R (pderiv (up_ring.monom P a n)) g  Pg = compose R (up_ring.monom P ([n]a) n) g"
          unfolding 61 using False A assms sub_mult[of g "up_ring.monom P ([n]  a) (n - 1)" "up_ring.monom P 𝟭 1" ] monom_mult[of "[n]a" 𝟭 "n-1" 1] 
          by (metis Nat.add_0_right R.add.nat_pow_closed R.one_closed R.r_one Suc_eq_plus1 add_eq_if monom_closed)
        have 63: "k::nat. Cring_Poly.compose R (up_ring.monom P ([k]  a) n) g = [k] PCring_Poly.compose R (up_ring.monom P a n) g"
        proof- fix k::nat show "Cring_Poly.compose R (up_ring.monom P ([k]  a) n) g = [k] PCring_Poly.compose R (up_ring.monom P a n) g"
            apply(induction k)
            using  UP_zero_closed assms(2) deg_zero monom_zero sub_const 
             apply (metis A P.add.nat_pow_0 add_nat_pow_monom)
          proof- 
            fix k::nat
            assume a: "Cring_Poly.compose R (monom P ([k]  a) n) g =
         [k] PCring_Poly.compose R (monom P a n) g"
            have 0: "(monom P ([Suc k]  a) n) = [Suc k]  a P(monom P 𝟭 n)"
              by (simp add: A monic_monom_smult)
            have 1: "(monom P ([Suc k]  a) n) = [k]  a P(monom P 𝟭 n) Pa P(monom P 𝟭 n) "
              unfolding 0 
              by (simp add: A UP_smult_l_distr)
            show "Cring_Poly.compose R (monom P ([Suc k]  a) n) g =
         [Suc k] P(Cring_Poly.compose R (monom P a n) g) "
              unfolding 1 
              by (simp add: A a assms(2) monic_monom_smult sub_add)
          qed
        qed
        have 64: "Cring_Poly.compose R (up_ring.monom P ([n]  a) n) g = [n] PCring_Poly.compose R (up_ring.monom P a n) g"
          using 63 by blast 
        show ?thesis unfolding 62 64 by blast 
      qed
      have 63: "k::nat. Cring_Poly.compose R (up_ring.monom P ([k]  a) n) g = [k] PCring_Poly.compose R (up_ring.monom P a n) g"
      proof- fix k::nat show "Cring_Poly.compose R (up_ring.monom P ([k]  a) n) g = [k] PCring_Poly.compose R (up_ring.monom P a n) g"
            apply(induction k)
            using UP_zero_closed assms(2) deg_zero monom_zero sub_const 
            apply (metis A P.add.nat_pow_0 add_nat_pow_monom)           
            using A P.add.nat_pow_Suc add_nat_pow_monom assms(2) is_UP_monomE(1) monom_is_UP_monom(1) rev_sub_add sub_rev_sub
            by (metis P.add.nat_pow_closed)            
      qed
      have 7: "([n] PCring_Poly.compose R (up_ring.monom P a n) g PCring_Poly.compose R (up_ring.monom P a n) g) = 
                                [Suc n] P(Cring_Poly.compose R (up_ring.monom P a n) g)"
        using A assms P.add.nat_pow_Suc by presburger
      have 8: "[Suc n] PCring_Poly.compose R (up_ring.monom P a n) g Ppderiv g = Cring_Poly.compose R (up_ring.monom P ([Suc n]  a) n) g Ppderiv g"
        unfolding  63[of "Suc n"] by blast 
      show ?thesis unfolding 5 6 7 8 using A assms pderiv_monom[of  "a" "Suc n"] 
        using P_def diff_Suc_1 by metis  
    qed
  qed
  qed
qed

lemma deriv_prod_rule_times_monom:
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "q  carrier P"
  shows "deriv ((monom P a n) Pq) b = (deriv (monom P a n) b)  (to_fun q b)  (to_fun (monom P a n) b)  deriv q b"
proof(rule poly_induct3[of q])
  show "q  carrier P"
    using assms by simp 
  show " p q. q  carrier P 
           p  carrier P 
           deriv (monom P a n Pp) b = deriv (monom P a n) b  to_fun p b  to_fun (monom P a n) b  deriv p b 
           deriv (monom P a n Pq) b = deriv (monom P a n) b  to_fun q b  to_fun (monom P a n) b  deriv q b 
           deriv (monom P a n P(p Pq)) b = deriv (monom P a n) b  to_fun (p Pq) b  to_fun (monom P a n) b  deriv (p Pq) b"
  proof- fix p q assume A: "q  carrier P" " p  carrier P"
          "deriv (monom P a n Pp) b = deriv (monom P a n) b  to_fun p b  to_fun (monom P a n) b  deriv p b"
          "deriv (monom P a n Pq) b = deriv (monom P a n) b  to_fun q b  to_fun (monom P a n) b  deriv q b"
    have "deriv (monom P a n P(p Pq)) b =  deriv (monom P a n) b  to_fun p b  to_fun (monom P a n) b  deriv p b
                                                  deriv (monom P a n) b  to_fun q b  to_fun (monom P a n) b  deriv q b"
      using A assms 
      by (simp add: P.r_distr R.add.m_assoc deriv_add deriv_closed to_fun_closed)    
    hence "deriv (monom P a n P(p Pq)) b =  deriv (monom P a n) b  to_fun p b deriv (monom P a n) b  to_fun q b 
               to_fun (monom P a n) b  deriv p b   to_fun (monom P a n) b  deriv q b"
      using A(1) A(2) R.add.m_assoc R.add.m_comm assms(1) assms(2) deriv_closed to_fun_closed by auto
    hence "deriv (monom P a n P(p Pq)) b =  deriv (monom P a n) b  (to_fun p b   to_fun q b)
               to_fun (monom P a n) b  (deriv p b   deriv q b)"
      by (simp add: A(1) A(2) R.add.m_assoc R.r_distr assms(1) assms(2) deriv_closed to_fun_closed)
    thus "deriv (monom P a n P(p Pq)) b = deriv (monom P a n) b  to_fun (p Pq) b  to_fun (monom P a n) b  deriv (p Pq) b"
      by (simp add: A(1) A(2) assms(2) deriv_add to_fun_plus)
  qed
  show "c m. c  carrier R   deriv (monom P a n Pmonom P c m) b = 
                                    deriv (monom P a n) b  to_fun (monom P c m) b 
                                   to_fun (monom P a n) b  deriv (monom P c m) b"
  proof- fix c m assume A: "c  carrier R"
    show "deriv (monom P a n Pmonom P c m) b = deriv (monom P a n) b  to_fun (monom P c m) b  to_fun (monom P a n) b  deriv (monom P c m) b"
    proof(cases "n = 0")
      case True
      have LHS: "deriv (monom P a n Pmonom P c m) b = deriv (monom P (a  c) m) b"
        by (metis A True add.left_neutral assms(1) monom_mult)
      have RHS: "deriv (monom P a n) b  to_fun (monom P c m) b  to_fun (monom P a n) b  deriv (monom P c m) b 
                = a  deriv (monom P c m) b "
        using deriv_const to_fun_monom A True assms(1) assms(2) deriv_closed by auto
      show ?thesis using A assms  LHS RHS deriv_monom 
        by (smt R.add.nat_pow_closed R.add_pow_rdistr R.m_assoc R.m_closed R.nat_pow_closed)
    next
      case False
      show ?thesis 
      proof(cases "m = 0")
        case True
        have LHS: "deriv (monom P a n Pmonom P c m) b = deriv (monom P (a  c) n) b"
          by (metis A True add.comm_neutral assms(1) monom_mult)
        have RHS: "deriv (monom P a n) b  to_fun (monom P c m) b  to_fun (monom P a n) b  deriv (monom P c m) b 
                = c  deriv (monom P a n) b "
          by (metis (no_types, lifting) A lcf_monom(1) P_def R.m_closed R.m_comm R.r_null
              R.r_zero True UP_cring.to_fun_ctrm UP_cring_axioms assms(1) assms(2) deg_const 
              deriv_closed deriv_const to_fun_closed monom_closed)
        show ?thesis using LHS RHS deriv_monom A assms 
          by (smt R.add.nat_pow_closed R.add_pow_ldistr R.m_assoc R.m_closed R.m_comm R.nat_pow_closed)
      next
        case F: False
        have pos: "n > 0" "m >0"
          using F False by auto 
        have RHS: "deriv (monom P a n Pmonom P c m) b = [(n + m)]  (a  c)  b [^] (n + m - 1)"
          using deriv_monom[of "a  c" b "n + m"] monom_mult[of a c n m]  
          by (simp add: A assms(1) assms(2))
        have LHS: "deriv (monom P a n) b  to_fun (monom P c m) b   to_fun (monom P a n) b  deriv (monom P c m) b
                    = [n]a (b[^](n-1))  c  b[^]m   a  b[^]n    [m]c (b[^](m-1))"
          using deriv_monom[of a b n] to_fun_monom[of a b n] 
            deriv_monom[of c b m] to_fun_monom[of c b m] A assms 
          by (simp add: R.m_assoc)
        have 0: "[n]a  (b[^](n-1))  c  b[^]m = [n]a  c  b[^](n + m -1) "
        proof-
          have "[n]a  (b[^](n-1))  c  b[^]m = [n]a   c  (b[^](n-1))  b[^]m"
            by (simp add: A R.m_lcomm R.semiring_axioms assms(1) assms(2) semiring.semiring_simprules(8))
          hence "[n]a  (b[^](n-1))  c  b[^]m = [n]a   c  ((b[^](n-1))  b[^]m)"
            by (simp add: A R.m_assoc assms(1) assms(2))
          thus ?thesis 
            by (simp add: False R.nat_pow_mult add_eq_if assms(2))
        qed
        have 1: "a  b[^]n    [m]c (b[^](m-1)) = a  [m]c  b[^](n + m -1)" 
        proof-
          have "a  b[^]n    [m]c (b[^](m-1)) = a   [m]c  b[^]n  (b[^](m-1))"
            using A R.m_comm R.m_lcomm assms(1) assms(2) by auto
          hence "a  b[^]n    [m]c (b[^](m-1)) = a   [m]c  (b[^]n  (b[^](m-1)))"
            by (simp add: A R.m_assoc assms(1) assms(2))
          thus ?thesis 
            by (simp add: F R.nat_pow_mult add.commute add_eq_if assms(2))
        qed
        have LHS: "deriv (monom P a n) b  to_fun (monom P c m) b   to_fun (monom P a n) b  deriv (monom P c m) b
                    = [n]a  c  b[^](n + m -1)    a  [m]c  b[^](n + m -1)"      
          using LHS 0 1 
          by simp
        hence LHS: "deriv (monom P a n) b  to_fun (monom P c m) b   to_fun (monom P a n) b  deriv (monom P c m) b
                    = [n] (a  c  b[^](n + m -1))    [m] (a  c  b[^](n + m -1))"                
          by (simp add: A R.add_pow_ldistr R.add_pow_rdistr assms(1) assms(2))
        show ?thesis using LHS RHS 
          by (simp add: A R.add.nat_pow_mult R.add_pow_ldistr assms(1) assms(2))
      qed
    qed
  qed
qed
               
lemma deriv_prod_rule:
  assumes "p  carrier P"
  assumes "q  carrier P"
  assumes "a  carrier R"
  shows "deriv (p Pq) a = deriv p a  (to_fun q a)  (to_fun p a)  deriv q a"
proof(rule poly_induct3[of p])
  show "p  carrier P"
    using assms(1) by simp 
  show " p qa.
       qa  carrier P 
       p  carrier P 
       deriv (p Pq) a = deriv p a  to_fun q a  to_fun p a  deriv q a 
       deriv (qa Pq) a = deriv qa a  to_fun q a  to_fun qa a  deriv q a 
       deriv ((p Pqa) Pq) a = deriv (p Pqa) a  to_fun q a  to_fun (p Pqa) a  deriv q a"
  proof- fix f g assume A: "f  carrier P" "g  carrier P"
                  "deriv (f Pq) a = deriv f a  to_fun q a  to_fun f a  deriv q a"
                  "deriv (g Pq) a = deriv g a  to_fun q a  to_fun g a  deriv q a"
    have "deriv ((f Pg) Pq) a = deriv f a  to_fun q a  to_fun f a  deriv q a 
                                         deriv g a  to_fun q a  to_fun g a  deriv q a"
      using A deriv_add 
      by (simp add: P.l_distr R.add.m_assoc assms(2) assms(3) deriv_closed to_fun_closed)
    hence "deriv ((f Pg) Pq) a = deriv f a  to_fun q a  deriv g a  to_fun q a  
                                       to_fun f a  deriv q a   to_fun g a  deriv q a"
      using R.a_comm R.a_assoc deriv_closed to_fun_closed assms 
      by (simp add: A(1) A(2))      
    hence "deriv ((f Pg) Pq) a = (deriv f a  to_fun q a  deriv g a  to_fun q a)  
                                         (to_fun f a  deriv q a   to_fun g a  deriv q a)"
      by (simp add: A(1) A(2) R.add.m_assoc assms(2) assms(3) deriv_closed to_fun_closed)
    thus "deriv ((f Pg) Pq) a = deriv (f Pg) a  to_fun q a  to_fun (f Pg) a  deriv q a"
      by (simp add: A(1) A(2) R.l_distr assms(2) assms(3) deriv_add deriv_closed to_fun_closed to_fun_plus)
  qed
  show "aa n. aa  carrier R  deriv (monom P aa n Pq) a = deriv (monom P aa n) a  to_fun q a  to_fun (monom P aa n) a  deriv q a"
    using deriv_prod_rule_times_monom 
    by (simp add: assms(2) assms(3))
qed

lemma pderiv_eval_deriv_monom:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "to_fun (pderiv (monom P a n)) b = deriv (monom P a n) b"
  using deriv_monom assms pderiv_monom  
  by (simp add: P_def UP_cring.to_fun_monom UP_cring_axioms)

lemma pderiv_eval_deriv:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "deriv f a = to_fun (pderiv f) a"
  apply(rule poly_induct3[of f])
  apply (simp add: assms(1))
    using assms(2) deriv_add to_fun_plus pderiv_add pderiv_closed apply presburger
      using assms(2) pderiv_eval_deriv_monom 
  by presburger

text‹Taking taylor expansions commutes with taking derivatives:›

lemma(in UP_cring) taylor_expansion_pderiv_comm:
  assumes "f  carrier (UP R)"
  assumes "c  carrier R"
  shows "pderiv (taylor_expansion R c f) = taylor_expansion R c (pderiv f)"
  apply(rule poly_induct3[of f])
  using assms unfolding P_def  apply blast
proof-
  fix p q assume A: " q  carrier (UP R)" "p  carrier (UP R)"
           "pderiv (taylor_expansion R c p) = taylor_expansion R c (pderiv p)"
           "pderiv (taylor_expansion R c q) = taylor_expansion R c (pderiv q)"
  have 0: " pderiv (taylor_expansion R c (p UP Rq)) =  pderiv (taylor_expansion R c p UP Rtaylor_expansion R c q)"
    using A P_def taylor_expansion_add assms(2) by presburger
  show "pderiv (taylor_expansion R c (p UP Rq)) = taylor_expansion R c (pderiv (p UP Rq))"
    unfolding 0
    using A(1) A(2) A(3) A(4) taylor_def UP_cring.taylor_closed UP_cring.taylor_expansion_add UP_cring.pderiv_add UP_cring.pderiv_closed UP_cring_axioms assms(2) by fastforce
next
  fix a n assume A: "a  carrier R"
  show "pderiv (taylor_expansion R c (up_ring.monom (UP R) a n)) = taylor_expansion R c (pderiv (up_ring.monom (UP R) a n))"
  proof(cases "n = 0")
    case True
    have 0: "deg R (taylor_expansion R c (up_ring.monom (UP R) a n)) = 0"
      unfolding True 
      using P_def A assms taylor_def taylor_deg deg_const is_UP_monomE(1) monom_is_UP_monom(2) by presburger
    have 1: "(pderiv (up_ring.monom (UP R) a n)) = 𝟬P⇙"
      unfolding True using P_def A assms pderiv_const by blast
    show ?thesis unfolding 1 using 0 A assms P_def 
      by (metis P.add.right_cancel taylor_closed taylor_def taylor_expansion_add UP_l_zero UP_zero_closed monom_closed pderiv_deg_0)
  next
    case False
    have 0: "pderiv (up_ring.monom (UP R) a n) =  (up_ring.monom (UP R) ([n]a) (n-1))"
      using A  
      by (simp add: UP_cring.pderiv_monom UP_cring_axioms)      
    have 1: "pderiv (taylor_expansion R c (up_ring.monom (UP R) a n)) = (Cring_Poly.compose R (up_ring.monom (UP R) ([n]a) (n-1)) (X_plus c)) Ppderiv (X_plus c)"
      using chain_rule[of "up_ring.monom (UP R) a n" "X_plus c"] unfolding 0 taylor_expansion_def
      using A P_def X_plus_closed assms(2) is_UP_monom_def monom_is_UP_monom(1) by presburger
    have 2: "pderiv (X_plus c) = 𝟭P⇙"
      using pderiv_add[of "X_poly R" "to_poly c"]  P.l_null P.l_one P.r_zero P_def R.one_closed  X_closed 
        X_poly_def X_poly_plus_def assms(2) monom_one pderiv_const to_poly_closed to_polynomial_def
      by (metis times_x_product_rule)
      show ?thesis unfolding 1 0 2 taylor_expansion_def
        by (metis "1" "2" A P.l_one P_def R.add.nat_pow_closed UP_m_comm UP_one_closed X_plus_closed assms(2) monom_closed sub_closed taylor_expansion_def)      
  qed
qed

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Linear Substitutions›
  (**********************************************************************)
  (**********************************************************************)
lemma(in UP_ring) lcoeff_Lcf:
  assumes "f  carrier P"
  shows "lcoeff f = lcf f"
  unfolding P_def
  using assms  coeff_simp[of f] by metis

lemma(in UP_cring) linear_sub_cfs:
  assumes "f  carrier (UP R)"
  assumes "d  carrier R"
  assumes "g = compose R f (up_ring.monom (UP R) d 1)"
  shows "g i = d[^]i  f i"
proof-
  have 0: "(up_ring.monom (UP R) d 1)  carrier (UP R)"
    using assms by (meson R.ring_axioms UP_ring.intro UP_ring.monom_closed)
  have 1: "(i. compose R f (up_ring.monom (UP R) d 1) i = d[^]i  f i)"
    apply(rule poly_induct3[of f])
    using assms unfolding P_def apply blast
  proof-
    show "p q. q  carrier (UP R) 
           p  carrier (UP R) 
           i. Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i = d [^] i  p i 
           i. Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i = d [^] i  q i 
           i. Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i = d [^] i  (p UP Rq) i"
    proof 
    fix p q i
    assume A: "q  carrier (UP R)"
           "p  carrier (UP R)"
           "i. Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i = d [^] i  p i"
           "i. Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i = d [^] i  q i"
    show "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i = d [^] i  (p UP Rq) i"
    proof- 
      have 1: "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) = 
              Cring_Poly.compose R p (up_ring.monom (UP R) d 1) UP RCring_Poly.compose R q (up_ring.monom (UP R) d 1)"
        using A(1) A(2) sub_add[of "up_ring.monom (UP R) d 1" q p] unfolding P_def 
        using "0" P_def sub_add by blast
      have 2: "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i = 
              Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i  Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i"
        using 1 by (metis (no_types, lifting) "0" A(1) A(2) P_def cfs_add sub_closed)
      have 3: "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i =   d [^] i  p i  d [^] i  q i"
        unfolding 2 using A by presburger
      have 4: "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i =   d [^] i  (p i   q i)"
        using "3" A(1) A(2) R.nat_pow_closed R.r_distr UP_car_memE(1) assms(2) by presburger
      thus "Cring_Poly.compose R (p UP Rq) (up_ring.monom (UP R) d 1) i = d [^] i  (p UP Rq) i"
        unfolding 4  using A(1) A(2) P_def cfs_add by presburger
    qed
  qed
  show "a n. a  carrier R 
           i. Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) i = d [^] i  up_ring.monom (UP R) a n i"
  proof fix a n i assume A: "a  carrier R"
    have 0: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1)  = 
                         a UP R(up_ring.monom (UP R) d 1)[^]UP Rn"
      using assms A 0 P_def monom_sub by blast
    have 1: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1)  = 
                         a UP R(d[^]n UP R(up_ring.monom (UP R) 𝟭 n))"
      unfolding  0 using A assms 
      by (metis P_def R.nat_pow_closed monic_monom_smult monom_pow mult.left_neutral)
    have 2: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1)  = 
                         (a d[^]n)UP R(up_ring.monom (UP R) 𝟭 n)"
      unfolding 1 using A assms  
      by (metis R.nat_pow_closed R.one_closed R.ring_axioms UP_ring.UP_smult_assoc1 UP_ring.intro UP_ring.monom_closed)
    show "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) i = d [^] i  up_ring.monom (UP R) a n i"
      apply(cases "i = n")
      unfolding 2 using A P_def R.m_closed R.m_comm R.nat_pow_closed assms(2) cfs_monom monic_monom_smult apply presburger
      using A P_def R.m_closed R.nat_pow_closed R.r_null assms(2) cfs_monom monic_monom_smult by presburger
  qed
  qed
  show ?thesis using 1 unfolding assms   
    by blast 
qed

lemma(in UP_cring) linear_sub_deriv: 
  assumes "f  carrier (UP R)"
  assumes "d  carrier R"
  assumes "g = compose R f (up_ring.monom (UP R) d 1)"
  assumes "c  carrier R"
  shows "pderiv g = d UP Rcompose R (pderiv f) (up_ring.monom (UP R) d 1)"
  unfolding assms 
proof(rule  poly_induct3[of f])
  show "f  carrier P"
    using assms unfolding P_def  by blast
  show " p q. q  carrier P 
           p  carrier P 
           pderiv (Cring_Poly.compose R p (up_ring.monom (UP R) d 1)) = d UP RCring_Poly.compose R (pderiv p) (up_ring.monom (UP R) d 1) 
           pderiv (Cring_Poly.compose R q (up_ring.monom (UP R) d 1)) = d UP RCring_Poly.compose R (pderiv q) (up_ring.monom (UP R) d 1) 
           pderiv (Cring_Poly.compose R (p Pq) (up_ring.monom (UP R) d 1)) =
           d UP RCring_Poly.compose R (pderiv (p Pq)) (up_ring.monom (UP R) d 1)"
  proof- fix p q assume A: "q  carrier P" "p  carrier P"
           "pderiv (Cring_Poly.compose R p (up_ring.monom (UP R) d 1)) = d UP RCring_Poly.compose R (pderiv p) (up_ring.monom (UP R) d 1)"
           "pderiv (Cring_Poly.compose R q (up_ring.monom (UP R) d 1)) = d UP RCring_Poly.compose R (pderiv q) (up_ring.monom (UP R) d 1)"
    show " pderiv (Cring_Poly.compose R (p Pq) (up_ring.monom (UP R) d 1)) =
           d UP RCring_Poly.compose R (pderiv (p Pq)) (up_ring.monom (UP R) d 1)"
      using A assms 
      by (smt P_def UP_a_closed UP_r_distr monom_closed monom_mult_is_smult pderiv_add pderiv_closed rev_sub_add sub_closed sub_rev_sub)
  qed
  show "a n. a  carrier R 
           pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) =
           d UP RCring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)"
  proof- fix a n assume A: "a  carrier R"
    have "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = a UP R(up_ring.monom P d 1)[^]UP Rn"
      using A assms sub_monom(2)  P_def is_UP_monomE(1) monom_is_UP_monom(1) by blast
    hence 0: "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = a UP R(up_ring.monom P (d[^]n) n)"
      using A assms P_def monom_pow nat_mult_1 by metis 
    show "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) =
           d UP RCring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)"
    proof(cases  "n = 0")
      case True
      have T0: "pderiv (up_ring.monom P a n) = 𝟬UP R⇙" unfolding True 
        using A P_def pderiv_const by blast
      have T1: "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = up_ring.monom P a n"
        unfolding True 
        using A assms  P_def deg_const is_UP_monomE(1) monom_is_UP_monom(1) sub_const by presburger
      thus ?thesis unfolding T0 T1  
        by (metis P.nat_pow_eone P_def UP_smult_closed UP_zero_closed X_closed assms(2) deg_zero monom_rep_X_pow smult_r_null sub_const)
    next
      case False
      have F0: "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = (a UP R(up_ring.monom P ([n]R(d[^]n))(n-1)))"
        using A assms pderiv_monom unfolding 0 
        using P_def R.nat_pow_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_smult by metis 
      have F1: "(pderiv (up_ring.monom P a n)) = up_ring.monom P ([n]  a) (n - 1)"
        using A assms pderiv_monom[of a n] by blast 
      hence F2: "(pderiv (up_ring.monom P a n)) = ([n]  a) UP Rup_ring.monom P 𝟭 (n - 1)"
        using A P_def monic_monom_smult by auto        
      have F3: "Cring_Poly.compose R ((([n]  a) UP R(up_ring.monom P 𝟭 (n - 1)))) (up_ring.monom (UP R) d 1) = 
                ([n]  a) UP R((up_ring.monom (UP R) d 1)[^]UP R(n-1))"
        using A F1 F2 P_def assms(2) monom_closed sub_monom(2) by fastforce
      have F4: "Cring_Poly.compose R ((([n]  a) UP R(up_ring.monom P 𝟭 (n - 1)))) (up_ring.monom (UP R) d 1) =
                ([n]  a) UP R((up_ring.monom (UP R) (d[^](n-1)) (n-1)))"
        by (metis F3 P_def assms(2) monom_pow nat_mult_1)
      have F5: "d UP R(Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)) = 
                (d ([n]  a)) UP Rup_ring.monom (UP R) (d [^] (n - 1)) (n - 1)"
        unfolding F4 F2 
        using A P_def assms(2) monom_closed smult_assoc1 by auto       
      have F6: "d UP R(Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)) = 
                (d  d[^](n-1) [n]  a) UP R((up_ring.monom (UP R) 𝟭 (n-1)))"
        unfolding F5 using False A  assms P_def  R.m_assoc R.m_closed R.m_comm R.nat_pow_closed monic_monom_smult monom_mult_smult
        by (smt R.add.nat_pow_closed)        
      have F7: "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = (a  ([n]R(d[^]n)) UP R(up_ring.monom P 𝟭 (n-1)))"
        unfolding F0 using A assms  P_def  R.m_closed R.nat_pow_closed monic_monom_smult monom_mult_smult
        by simp
      have F8: "a  [n]  (d [^] n)  = d  d [^] (n - 1)  [n]  a"
      proof-
        have F80: "d  d [^] (n - 1)  [n]  a = d [^] n  [n]  a"
          using A assms  False by (metis R.nat_pow_Suc2 add.right_neutral add_eq_if)
        show ?thesis unfolding F80 
          using A R.add_pow_rdistr R.m_comm R.nat_pow_closed assms(2) by presburger
      qed
      show ?thesis unfolding F6 F7 unfolding  F8 P_def by blast  
    qed
  qed
qed

lemma(in UP_cring) linear_sub_deriv': 
  assumes "f  carrier (UP R)"
  assumes "d  carrier R"
  assumes "g = compose R f (up_ring.monom (UP R) d 1)"
  assumes "c  carrier R"
  shows "pderiv g = compose R (d UP Rpderiv f) (up_ring.monom (UP R) d 1)"
  using assms linear_sub_deriv[of f d g c] P_def is_UP_monomE(1) is_UP_monomI pderiv_closed sub_smult by metis 

lemma(in UP_cring) linear_sub_inv:
  assumes "f  carrier (UP R)"
  assumes "d  Units R"
  assumes "g = compose R f (up_ring.monom (UP R) d 1)"
  shows "compose R g (up_ring.monom (UP R) (inv d) 1) = f"
  unfolding assms 
proof fix x 
  have 0: "Cring_Poly.compose R (Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) (up_ring.monom (UP R) (inv d) 1) x = 
        (inv d)[^]x  ((Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) x)"
    apply(rule linear_sub_cfs) 
    using P_def R.Units_closed assms(1) assms(2) monom_closed sub_closed apply auto[1]
     apply (simp add: assms(2))
    by blast
  show "Cring_Poly.compose R (Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) (up_ring.monom (UP R) (inv d) 1) x = f x "
    unfolding 0 using linear_sub_cfs[of f d "Cring_Poly.compose R f (up_ring.monom (UP R) d 1)" x]
      assms 
  by (smt R.Units_closed R.Units_inv_closed R.Units_l_inv R.m_assoc R.m_comm R.nat_pow_closed R.nat_pow_distrib R.nat_pow_one R.r_one UP_car_memE(1))
qed

lemma(in UP_cring) linear_sub_deg: 
  assumes "f  carrier (UP R)"
  assumes "d  Units R"
  assumes "g = compose R f (up_ring.monom (UP R) d 1)"
  shows "deg R g = deg R f"
proof(cases "deg R f = 0")
  case True
  show ?thesis using assms 
    unfolding True assms using P_def True monom_closed 
    by (simp add: R.Units_closed sub_const)    
next
  case False
  have 0: "monom (UP R) d 1 (deg R (monom (UP R) d 1)) = d"
    using assms  lcf_monom(2) by blast
  have 1: "d[^](deg R f)  Units R" 
    using assms(2)
    by (metis Group.comm_monoid.axioms(1) R.units_comm_group R.units_of_pow comm_group_def monoid.nat_pow_closed units_of_carrier)
  have 2: "f (deg R f)  𝟬"
    using assms False P_def UP_cring.ltrm_rep_X_pow UP_cring_axioms deg_ltrm degree_monom by fastforce
  have "deg R g = deg R f * deg R (up_ring.monom (UP R) d 1)" 
    unfolding assms 
    apply(rule cring_sub_deg[of "up_ring.monom (UP R) d 1" f] )
    using assms P_def monom_closed apply blast
    unfolding P_def apply(rule assms)
    unfolding 0 using 2 1 
    by (metis R.Units_closed R.Units_l_cancel R.m_comm R.r_null R.zero_closed UP_car_memE(1) assms(1))  
  thus ?thesis using assms  unfolding assms
    by (metis (no_types, lifting) P_def R.Units_closed deg_monom deg_zero is_UP_monomE(1) linear_sub_inv monom_is_UP_monom(2) monom_zero mult.right_neutral mult_0_right sub_closed sub_const)
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
section‹Lemmas About Polynomial Division›
(**************************************************************************************************)
(**************************************************************************************************)
context UP_cring
begin

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Division by Linear Terms›
  (**********************************************************************)
  (**********************************************************************)
definition UP_root_div where
"UP_root_div f a =  (poly_shift (T⇘af)) of (X_minus a)"

definition UP_root_rem where
"UP_root_rem f a = ctrm (T⇘af)"

lemma UP_root_div_closed:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "UP_root_div f a  carrier P"
  using assms 
  unfolding UP_root_div_def
  by (simp add: taylor_closed X_minus_closed poly_shift_closed sub_closed)

lemma rem_closed:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "UP_root_rem f a  carrier P"
  using assms 
  unfolding UP_root_rem_def
  by (simp add: taylor_closed ctrm_is_poly) 

lemma rem_deg:
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "degree (UP_root_rem f a) = 0"
  by (simp add: taylor_closed assms(1) assms(2) ctrm_degree UP_root_rem_def)

lemma remainder_theorem:
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "g = UP_root_div f a"
  assumes "r = UP_root_rem f a"
  shows "f = r P(X_minus a) Pg"
proof-
  have  "T⇘af = (ctrm (T⇘af)) PX Ppoly_shift (T⇘af)"
    using poly_shift_eq[of "T⇘af"] assms taylor_closed 
    by blast     
  hence 1: "T⇘af of (X_minus a) = (ctrm (T⇘af)) P(X_minus a) P(poly_shift (T⇘af) of (X_minus a))"
    using assms taylor_closed[of f a] X_minus_closed[of a] X_closed 
          sub_add[of "X_minus a" "ctrm (T⇘af)" "X Ppoly_shift (T⇘af)"] 
          sub_const[of "X_minus a"]  sub_mult[of "X_minus a" X  "poly_shift (T⇘af)"] 
          ctrm_degree ctrm_is_poly P.m_closed poly_shift_closed sub_X 
    by presburger
  have 2: "f = (ctrm (T⇘af)) P(X_minus a) P(poly_shift (T⇘af) of (X_minus a))"
    using 1 taylor_id[of a f] assms 
    by simp 
  then show ?thesis
    using assms
    unfolding UP_root_div_def UP_root_rem_def
    by auto 
qed

lemma remainder_theorem':
  assumes "f  carrier P"
  assumes "a  carrier R"
  shows "f = UP_root_rem f a P(X_minus a) PUP_root_div f a"
  using assms remainder_theorem by auto 

lemma factor_theorem:
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "g = UP_root_div f a"
  assumes "to_fun f a = 𝟬"
  shows "f = (X_minus a) Pg"
  using remainder_theorem[of f a g _] assms
  unfolding UP_root_rem_def 
  by (simp add: ctrm_zcf taylor_zcf taylor_closed UP_root_div_closed X_minus_closed)

lemma factor_theorem':
  assumes "f  carrier P"
  assumes "a  carrier R"
  assumes "to_fun f a = 𝟬"
  shows "f = (X_minus a) PUP_root_div f a"
  by (simp add: assms(1) assms(2) assms(3) factor_theorem)

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Geometric Sums›
  (**********************************************************************)
  (**********************************************************************)
lemma geom_quot:
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "p = monom P 𝟭 (Suc n) Pmonom P (b[^](Suc n)) 0 "
  assumes "g = UP_root_div p b"
  shows "a[^](Suc n)  b[^] (Suc n) = (a  b)  (to_fun g a)"
proof-
  have root: "to_fun p b = 𝟬"
    using assms to_fun_const[of "b[^](Suc n)" b] to_fun_monic_monom[of b "Suc n"] R.nat_pow_closed[of b "Suc n"]
          to_fun_diff[of "monom P 𝟭 (Suc n)" "monom P (b[^](Suc n)) 0" b] monom_closed 
    by (metis P.minus_closed P_def R.one_closed R.zero_closed UP_cring.f_minus_ctrm 
        UP_cring.to_fun_diff UP_cring_axioms zcf_to_fun cfs_monom to_fun_const)
  have LHS: "to_fun p a = a[^](Suc n)  b[^] (Suc n)"
    using assms to_fun_const to_fun_monic_monom to_fun_diff 
    by auto
  have RHS: "to_fun ((X_minus b) Pg) a = (a  b)  (to_fun g a)"
    using to_fun_mult[of g "X_minus b"] assms X_minus_closed 
    by (metis P.minus_closed P_def R.nat_pow_closed R.one_closed UP_cring.UP_root_div_closed UP_cring_axioms to_fun_X_minus monom_closed)
  show ?thesis 
    using RHS LHS root factor_theorem' assms(2) assms(3) assms(4) 
    by auto
qed

end

context UP_cring
begin 

definition geometric_series where
"geometric_series n a b = to_fun (UP_root_div (monom P 𝟭 (Suc n) UP R(monom P (b[^](Suc n)) 0)) b) a"

lemma geometric_series_id:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "a[^](Suc n) b[^] (Suc n) = (a  b)  (geometric_series n a b)"
  using assms geom_quot 
  by (simp add: P_def geometric_series_def)

lemma geometric_series_closed:
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "(geometric_series n a b)  carrier R"
  unfolding geometric_series_def 
  using assms  P.minus_closed P_def UP_root_div_closed to_fun_closed monom_closed 
  by auto

text‹Shows that $a^n - b^n$ has $a - b$ as a factor:›
lemma to_fun_monic_monom_diff:
    assumes "a  carrier R"
    assumes "b  carrier R"
    shows "c. c  carrier R  to_fun (monom P 𝟭 n) a  to_fun (monom P 𝟭 n) b = (a  b)  c"
proof(cases "n = 0")
  case True
  have "to_fun (monom P 𝟭 0) a  to_fun (monom P 𝟭 0) b = (a  b)  𝟬"
    unfolding a_minus_def using to_fun_const[of 𝟭] assms 
    by (simp add: R.r_neg)
  then show ?thesis 
    using True by blast 
next
  case False
  then show ?thesis 
    using Suc_diff_1[of n] geometric_series_id[of a b "n-1"] geometric_series_closed[of a b "n-1"]
        assms(1) assms(2) to_fun_monic_monom 
    by auto
qed

lemma to_fun_diff_factor:
  assumes "a  carrier R"
  assumes "b  carrier R"
  assumes "f  carrier P"
  shows "c. c  carrier R (to_fun f a)  (to_fun f b) = (a  b)c"
proof(rule poly_induct5[of f])
  show "f  carrier P" using assms by simp 
  show "p q. q  carrier P 
           p  carrier P 
           c. c  carrier R  to_fun p a  to_fun p b = (a  b)  c 
           c. c  carrier R  to_fun q a  to_fun q b = (a  b)  c 
     c. c  carrier R  to_fun (p Pq) a  to_fun (p Pq) b = (a  b)  c"
  proof- fix p q assume A: "q  carrier P" "p  carrier P" 
                          "c. c  carrier R  to_fun p a  to_fun p b = (a  b)  c"
                          "c. c  carrier R  to_fun q a  to_fun q b = (a  b)  c"
    obtain c  where c_def: "c  carrier R  to_fun p a  to_fun p b = (a  b)  c"
      using A by blast 
    obtain c'  where c'_def: "c'  carrier R  to_fun q a  to_fun q b = (a  b)  c'"
      using A by blast 
    have 0: "(a  b)  c  (a  b)  c' = (a  b)(c  c')"
      using assms c_def c'_def  unfolding a_minus_def       
      by (simp add: R.r_distr R.r_minus)   
    have 1: "to_fun (p Pq) a  to_fun (p Pq) b = to_fun p a  to_fun q a  to_fun p b  to_fun q b"
      using A to_fun_plus[of p q a] to_fun_plus[of p q b] assms to_fun_closed 
            R.ring_simprules(19)[of "to_fun p b" "to_fun q b"] 
      by (simp add: R.add.m_assoc R.minus_eq to_fun_plus)
    hence "to_fun (p Pq) a  to_fun (p Pq) b = to_fun p a  to_fun p b  to_fun q a  to_fun q b"
      using 0 A assms R.ring_simprules  to_fun_closed a_assoc a_comm 
      unfolding a_minus_def 
      by smt
    hence "to_fun (p Pq) a  to_fun (p Pq) b = to_fun p a  to_fun p b  (to_fun q a  to_fun q b)"
      using 0 A assms R.ring_simprules  to_fun_closed 
      unfolding a_minus_def 
      by metis 
    hence "to_fun (p Pq) a  to_fun (p Pq) b = (a  b)(c  c')"
      using 0 A c_def c'_def  
      by simp
    thus "c. c  carrier R  to_fun (p Pq) a  to_fun (p Pq) b = (a  b)  c"
      using R.add.m_closed c'_def c_def by blast
  qed
  show "n. c. c  carrier R  to_fun (monom P 𝟭 n) a  to_fun (monom P 𝟭 n) b = (a  b)  c"
    by (simp add: assms(1) assms(2) to_fun_monic_monom_diff)
  show "p aa.
       aa  carrier R 
       p  carrier P  c. c  carrier R  to_fun p a  to_fun p b = (a  b)  c  c. c  carrier R  to_fun (aa Pp) a  to_fun (aa Pp) b = (a  b)  c"
  proof- fix p c assume A: "c  carrier R" " p  carrier P" 
                          "e. e  carrier R  to_fun p a  to_fun p b = (a  b)  e"
    then obtain d where d_def: "d  carrier R  to_fun p a  to_fun p b = (a  b)  d"
      by blast 
    have "to_fun (c Pp) a  to_fun (c Pp) b = c  (to_fun p a  to_fun p b)"
      using A d_def assms to_fun_smult[of p a c] to_fun_smult[of p b c] 
            to_fun_closed[of p a] to_fun_closed[of p b]  R.ring_simprules 
      by smt
    hence "cd  carrier R  to_fun (c Pp) a  to_fun (c Pp) b = (a  b)  (c d)"
      by (simp add: A(1) R.m_lcomm assms(1) assms(2) d_def)
    thus "e. e  carrier R  to_fun (c Pp) a  to_fun (c Pp) b = (a  b)  e" 
      by blast 
  qed
qed

text‹Any finite set over a domain is the zero set of a polynomial:›
lemma(in UP_domain) fin_set_poly_roots:
  assumes "F  carrier R"
  assumes "finite F"
  shows " P  carrier (UP R).  x  carrier R. to_fun P x = 𝟬  x  F"
  apply(rule finite.induct)
    apply (simp add: assms(2))
proof-
  show "Pcarrier (UP R). xcarrier R. (to_fun P x = 𝟬) = (x  {})"
  proof-
    have "xcarrier R. (to_fun (𝟭UP R) x = 𝟬) = (x  {})"
    proof
      fix x
      assume A: "x  carrier R"
      then have "(to_fun (𝟭UP R)) x = 𝟭"
        by (metis P_def R.one_closed UP_cring.to_fun_to_poly UP_cring_axioms ring_hom_one to_poly_is_ring_hom)
      then show "(to_fun 𝟭UP Rx = 𝟬) = (x  {})"
        by simp      
    qed
    then show ?thesis 
      using P_def UP_one_closed 
      by blast
  qed
  show "A a. finite A 
           Pcarrier (UP R). xcarrier R. (to_fun P x = 𝟬) = (x  A)  Pcarrier (UP R). xcarrier R. (to_fun P x = 𝟬) = (x  insert a A)"
  proof-
    fix A :: "'a set" fix a
    assume fin_A: "finite A"
    assume IH: "Pcarrier (UP R). xcarrier R. (to_fun P x = 𝟬) = (x  A)"
    then obtain p  where p_def: "p carrier (UP R)  (xcarrier R. (to_fun p x = 𝟬) = (x  A))"
      by blast
    show "Pcarrier (UP R). xcarrier R. (to_fun P x = 𝟬) = (x  insert a A)"
    proof(cases "a  carrier R")
      case True
        obtain Q where Q_def: "Q = p UP R(X UP Rto_poly a)"
      by blast
      have "xcarrier R. (to_fun Q x = 𝟬) = (x  insert a A)" 
      proof fix x
        assume P: "x  carrier R"
        have P0: "to_fun (X UP Rto_poly a) x = x  a" 
          using to_fun_plus[of X "UP Rto_poly a" x] True P 
          unfolding a_minus_def 
          by (metis X_poly_minus_def a_minus_def to_fun_X_minus)
        then have "to_fun Q x = (to_fun p x)  (x  a)"
        proof-
          have 0: " p  carrier P"
          by (simp add: P_def p_def)
          have 1: "    X UP Rto_poly a  carrier P"
          using P.minus_closed P_def True X_closed to_poly_closed by auto
          have 2: "x  carrier R"
          by (simp add: P)
          then show ?thesis 
          using to_fun_mult[of p "(X UP Rto_poly a)" x] P0 0 1 2 Q_def True P_def to_fun_mult 
          by auto
        qed
       then show "(to_fun Q x = 𝟬) = (x  insert a A)" 
        using p_def 
        by (metis P R.add.inv_closed R.integral_iff R.l_neg R.minus_closed R.minus_unique True UP_cring.to_fun_closed UP_cring_axioms a_minus_def insert_iff)        
    qed
    then have "Q  carrier (UP R)  (xcarrier R. (to_fun Q x = 𝟬) = (x  insert a A))" 
      using P.minus_closed P_def Q_def True UP_mult_closed X_closed p_def to_poly_closed by auto
    then show ?thesis 
      by blast
  next
    case False 
    then show ?thesis 
      using IH subsetD by auto
  qed
  qed
qed

  (**********************************************************************)
  (**********************************************************************)
  subsection‹Polynomial Evaluation at Multiplicative Inverses›
  (**********************************************************************)
  (**********************************************************************)
text‹For every polynomial $p(x)$ of degree $n$, there is a unique polynomial $q(x)$ which satisfies the equation $q(x) = x^n p(1/x)$. This section defines this polynomial and proves this identity.›
definition(in UP_cring) one_over_poly where
"one_over_poly p = (λ n. if n  degree p then p ((degree  p) - n) else 𝟬)"

lemma(in UP_cring) one_over_poly_closed: 
  assumes "p  carrier P"
  shows "one_over_poly p  carrier P"
  apply(rule UP_car_memI[of "degree p" ])
  unfolding one_over_poly_def using assms  apply simp
  by (simp add: assms cfs_closed)
  
lemma(in UP_cring) one_over_poly_monom:
  assumes "a  carrier R"
  shows "one_over_poly (monom P a n) = monom P a 0"
  apply(rule ext)
  unfolding one_over_poly_def using assms 
  by (metis cfs_monom deg_monom diff_diff_cancel diff_is_0_eq diff_self_eq_0 zero_diff)

lemma(in UP_cring) one_over_poly_monom_add:
  assumes "a  carrier R"
  assumes "a  𝟬"
  assumes "p  carrier P"
  assumes "degree p < n"
  shows "one_over_poly (p Pmonom P a n) = monom P a 0 Pmonom P 𝟭 (n - degree p) Pone_over_poly p"
proof-
  have 0: "degree (p Pmonom P a n) = n"
    by (simp add: assms(1) assms(2) assms(3) assms(4) equal_deg_sum)
  show ?thesis 
  proof(rule ext) fix x show "one_over_poly (p Pmonom P a n) x =
                                 (monom P a 0 Pmonom P 𝟭 (n - deg R p) Pone_over_poly p) x"
    proof(cases "x = 0")
      case T: True
      have T0: "one_over_poly (p Pmonom P a n) 0 = a"
        unfolding one_over_poly_def 
        by (metis lcf_eq lcf_monom(1) ltrm_of_sum_diff_deg P.add.m_closed assms(1) assms(2) assms(3) assms(4) diff_zero le0 monom_closed)        
      have T1: "(monom P a 0 Pmonom P 𝟭 (n - degree p) Pone_over_poly p) 0 = a"
        using one_over_poly_closed 
        by (metis (no_types, lifting) lcf_monom(1) R.one_closed R.r_zero UP_m_comm UP_mult_closed assms(1) assms(3) assms(4) cfs_add cfs_monom_mult deg_const monom_closed zero_less_diff)                    
      show ?thesis  using T0 T1 T by auto 
    next
      case F: False
      show ?thesis
      proof(cases "x < n - degree p")
        case True
        then have T0: "degree p < n - x  n - x < n"
          using F by auto
        then have T1: "one_over_poly (p Pmonom P a n) x = 𝟬"
          using True F 0 unfolding one_over_poly_def 
          using assms(1) assms(3) coeff_of_sum_diff_degree0 
          by (metis ltrm_cfs ltrm_of_sum_diff_deg P.add.m_closed P.add.m_comm assms(2) assms(4) monom_closed nat_neq_iff)          
        have "(monom P a 0 Pmonom P 𝟭 (n - degree p) Pone_over_poly p) x = 𝟬"  
          using True F 0  one_over_poly_def one_over_poly_closed 
          by (metis (no_types, lifting) P.add.m_comm P.m_closed R.one_closed UP_m_comm assms(1)
              assms(3) cfs_monom_mult coeff_of_sum_diff_degree0 deg_const monom_closed neq0_conv)
        then show ?thesis using T1 by auto 
      next
        case False
        then have "n - degree p  x"
          by auto 
        then obtain k where k_def: "k + (n - degree p) = x"
          using le_Suc_ex  diff_add 
          by blast
        have F0: "(monom P a 0 Pmonom P 𝟭 (n - deg R p) Pone_over_poly p) x
              = one_over_poly p k"
          using k_def one_over_poly_closed assms 
                times_X_pow_coeff[of "one_over_poly p" "n - deg R p" k]
                P.m_closed 
          by (metis (no_types, lifting) P.add.m_comm R.one_closed add_gr_0 coeff_of_sum_diff_degree0 deg_const monom_closed zero_less_diff)                
        show ?thesis 
        proof(cases "x  n")
          case True
          have T0: "n - x = degree p - k"
            using assms(4) k_def by linarith
          have T1: "n - x < n"
            using True F 
            by linarith
          then have F1: "(p Pmonom P a n) (n - x) = p (degree p - k)"
            using True False F0 0 k_def cfs_add 
            by (simp add: F0 T0 assms(1) assms(3) cfs_closed cfs_monom)            
          then show ?thesis 
            using "0" F0 assms(1) assms(2) assms(3) degree_of_sum_diff_degree k_def one_over_poly_def
            by auto          
        next
          case False
          then show ?thesis 
            using "0" F0 assms(1) assms(2) assms(3) degree_of_sum_diff_degree k_def one_over_poly_def 
            by auto          
        qed
      qed
    qed 
  qed
qed

lemma( in UP_cring) one_over_poly_eval: 
  assumes "p  carrier P"
  assumes "x  carrier R"
  assumes "x  Units R"
  shows "to_fun (one_over_poly p) x = (x[^](degree p))  (to_fun p (invRx))" 
proof(rule poly_induct6[of p])
  show " p  carrier P"
    using assms by simp 
  show "a n. a  carrier R  
        to_fun (one_over_poly (monom P a 0)) x = x [^] deg R (monom P a 0)  to_fun (monom P a 0) (inv x)"
    using assms to_fun_const one_over_poly_monom by auto
  show "a n p.
       a  carrier R 
       a  𝟬 
       p  carrier P 
       deg R p < n 
       to_fun (one_over_poly p) x = x [^] deg R p  to_fun p (inv x) 
      to_fun (one_over_poly (p Pmonom P a n)) x = x [^] deg R (p Pmonom P a n)  to_fun (p Pmonom P a n) (inv x)"
  proof- fix a n p assume A: "a  carrier R" "a  𝟬" "p  carrier P" "deg R p < n"
                             "to_fun (one_over_poly p) x = x [^] deg R p  to_fun p (inv x)"
    have "one_over_poly (p Pmonom P a n) = monom P a 0 Pmonom P 𝟭 (n - degree p) Pone_over_poly p"
      using A by (simp add: one_over_poly_monom_add)
    hence "to_fun ( one_over_poly (p Pmonom P a n)) x = 
                a  to_fun ( monom P 𝟭 (n - degree p) Pone_over_poly p) x"
      using A to_fun_plus one_over_poly_closed cfs_add 
      by (simp add: assms(2) to_fun_const)
    hence "to_fun (one_over_poly (p Pmonom P a n)) x = a  x[^](n - degree p)  x [^] degree p  to_fun p (inv x)"
      by (simp add: A(3) A(5) R.m_assoc assms(2) assms(3) to_fun_closed to_fun_monic_monom to_fun_mult one_over_poly_closed)
    hence 0:"to_fun (one_over_poly (p Pmonom P a n)) x = a  x[^]n  to_fun p (inv x)"
      using A R.nat_pow_mult assms(2)
      by auto
    have 1: "to_fun (one_over_poly (p Pmonom P a n)) x = x[^]n  (a  inv x [^]n  to_fun p (inv x))"
    proof-
      have "x[^]n  a  inv x [^]n = a"
        by (metis (no_types, opaque_lifting) A(1) R.Units_inv_closed R.Units_r_inv R.m_assoc 
            R.m_comm R.nat_pow_closed R.nat_pow_distrib R.nat_pow_one R.r_one assms(2) assms(3))
      thus ?thesis 
        using A R.ring_simprules(23)[of _ _  "x[^]n"] 0 R.m_assoc assms(2) assms(3) to_fun_closed 
        by auto
    qed
    have 2: "degree (p Pmonom P a n) = n"
      by (simp add: A(1) A(2) A(3) A(4) equal_deg_sum)
    show " to_fun (one_over_poly (p Pmonom P a n)) x = x [^] deg R (p Pmonom P a n)  to_fun (p Pmonom P a n) (inv x)"
      using 1 2 
      by (metis (no_types, lifting) A(1) A(3) P_def R.Units_inv_closed R.add.m_comm 
          UP_cring.to_fun_monom UP_cring_axioms assms(3) to_fun_closed to_fun_plus monom_closed)
  qed
qed

end

(**************************************************************************************************)
(**************************************************************************************************)
section‹Lifting Homomorphisms of Rings to Polynomial Rings by Application to Coefficients›
(**************************************************************************************************)
(**************************************************************************************************)

definition poly_lift_hom where
"poly_lift_hom R S φ = eval R (UP S) (to_polynomial S  φ) (X_poly S)"

context UP_ring
begin

lemma(in UP_cring) pre_poly_lift_hom_is_hom:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "ring_hom_ring R (UP S) (to_polynomial S  φ)"
  apply(rule ring_hom_ringI) 
       apply (simp add: R.ring_axioms)
  apply (simp add: UP_ring.UP_ring UP_ring.intro assms(1) cring.axioms(1))
  using UP_cring.intro UP_cring.to_poly_closed assms(1) assms(2) ring_hom_closed apply fastforce
  using assms UP_cring.to_poly_closed[of S] ring_hom_closed[of φ R S] comp_apply[of "to_polynomial S" φ]
  unfolding UP_cring_def 
  apply (metis UP_cring.to_poly_mult UP_cring_def ring_hom_mult)
  using assms UP_cring.to_poly_closed[of S] ring_hom_closed[of φ R S] comp_apply[of "to_polynomial S" φ]
  unfolding UP_cring_def 
    apply (metis UP_cring.to_poly_add UP_cring_def ring_hom_add)
  using assms UP_cring.to_poly_closed[of S] ring_hom_one[of φ R S] comp_apply[of "to_polynomial S" φ]
  unfolding UP_cring_def 
  by (simp add: φ  ring_hom R S  φ 𝟭 = 𝟭S⇙› UP_cring.intro UP_cring.to_poly_is_ring_hom ring_hom_one)

lemma(in UP_cring) poly_lift_hom_is_hom:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "poly_lift_hom R S φ  ring_hom (UP R) (UP S)"
  unfolding poly_lift_hom_def
  apply( rule UP_pre_univ_prop.eval_ring_hom[of R "UP S" ])
  unfolding UP_pre_univ_prop_def
  apply (simp add: R_cring RingHom.ring_hom_cringI UP_cring.UP_cring UP_cring_def assms(1) assms(2) pre_poly_lift_hom_is_hom)
  by (simp add: UP_cring.X_closed UP_cring.intro assms(1))

lemma(in UP_cring) poly_lift_hom_closed: 
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  shows "poly_lift_hom R S φ p  carrier (UP S)"
  by (metis assms(1) assms(2) assms(3) poly_lift_hom_is_hom ring_hom_closed)

lemma(in UP_cring) poly_lift_hom_add: 
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  assumes "q  carrier (UP R)"
  shows "poly_lift_hom R S φ (p UP Rq) = poly_lift_hom R S φ p UP Spoly_lift_hom R S φ q"
  using assms poly_lift_hom_is_hom[of S φ] ring_hom_add[of "poly_lift_hom R S φ" "UP R" "UP S" p q] 
  by blast

lemma(in UP_cring) poly_lift_hom_mult: 
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  assumes "q  carrier (UP R)"
  shows "poly_lift_hom R S φ (p UP Rq) = poly_lift_hom R S φ p UP Spoly_lift_hom R S φ q"
  using assms poly_lift_hom_is_hom[of  S φ] ring_hom_mult[of "poly_lift_hom R S φ" "UP R" "UP S" p q] 
  by blast 

lemma(in UP_cring) poly_lift_hom_extends_hom:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "r  carrier R"
  shows "poly_lift_hom R S φ (to_polynomial R r) = to_polynomial S (φ r)"
  using UP_pre_univ_prop.eval_const[of R "UP S" "to_polynomial S  φ" "X_poly S" r ] assms 
        comp_apply[of "λa. monom (UP S) a 0"  φ r] pre_poly_lift_hom_is_hom[of  S φ]
  unfolding poly_lift_hom_def to_polynomial_def UP_pre_univ_prop_def 
  by (simp add: R_cring RingHom.ring_hom_cringI UP_cring.UP_cring UP_cring.X_closed UP_cring.intro)

lemma(in UP_cring) poly_lift_hom_extends_hom':
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "r  carrier R"
  shows "poly_lift_hom R S φ (monom P r 0) = monom (UP S) (φ r) 0"
  using poly_lift_hom_extends_hom[of S φ r] assms 
  unfolding to_polynomial_def P_def 
  by blast

lemma(in UP_cring) poly_lift_hom_smult: 
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ  (a UP Rp) = φ a UP S(poly_lift_hom R S φ p)"
  using assms poly_lift_hom_is_hom[of S φ] poly_lift_hom_extends_hom'[of S φ a] 
        poly_lift_hom_mult[of  S φ "monom P a 0" p] ring_hom_closed[of φ R S a]
        UP_ring.monom_mult_is_smult[of S "φ a" "poly_lift_hom R S φ p"]
        monom_mult_is_smult[of a p] monom_closed[of a 0] poly_lift_hom_closed[of S φ p]
  unfolding to_polynomial_def UP_ring_def P_def cring_def 
  by simp        

lemma(in UP_cring) poly_lift_hom_monom:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "r  carrier R"
  shows "poly_lift_hom R S φ (monom (UP R) r n) = (monom (UP S) (φ r) n)"
proof-
  have "eval R (UP S) (to_polynomial S  φ) (X_poly S) (monom (UP R) r n) = (to_polynomial S  φ) r UP SX_poly S [^]UP Sn"
    using assms UP_pre_univ_prop.eval_monom[of R "UP S" "to_polynomial S  φ" r "X_poly S" n]
    unfolding UP_pre_univ_prop_def UP_cring_def ring_hom_cring_def
    by (meson UP_cring.UP_cring UP_cring.X_closed UP_cring.pre_poly_lift_hom_is_hom UP_cring_axioms 
        UP_cring_def ring_hom_cring_axioms.intro ring_hom_ring.homh)     
   then have "eval R (UP S) (to_polynomial S  φ) (X_poly S) (monom (UP R) r n) = (to_polynomial S (φ r)) UP SX_poly S [^]UP Sn"
    by simp
  then show ?thesis
  unfolding poly_lift_hom_def
  using assms UP_cring.monom_rep_X_pow[of S "φ r" n]  ring_hom_closed[of φ R S r]
    by (metis UP_cring.X_closed UP_cring.intro UP_cring.monom_sub UP_cring.sub_monom(1))
qed

lemma(in UP_cring) poly_lift_hom_X_var:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "poly_lift_hom R S φ (monom (UP R) 𝟭R1) = (monom (UP S) 𝟭S1)"
  using assms(1) assms(2) poly_lift_hom_monom ring_hom_one by fastforce

lemma(in UP_cring) poly_lift_hom_X_var':
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "poly_lift_hom R S φ (X_poly R) = (X_poly S)"
  unfolding X_poly_def 
  using assms(1) assms(2) poly_lift_hom_X_var by blast
  
lemma(in UP_cring) poly_lift_hom_X_var'':
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "poly_lift_hom R S φ (monom (UP R) 𝟭Rn) = (monom (UP S) 𝟭Sn)"
  using assms(1) assms(2) poly_lift_hom_monom ring_hom_one by fastforce

lemma(in UP_cring) poly_lift_hom_X_var''':
  assumes "cring S"
  assumes "φ  ring_hom R S"
  shows "poly_lift_hom R S φ (X_poly R [^]UP R(n::nat)) = (X_poly S) [^]UP S(n::nat)"
  using assms 
  by (smt ltrm_of_X P.nat_pow_closed P_def R.ring_axioms UP_cring.to_fun_closed UP_cring.intro
      UP_cring.monom_pow UP_cring.poly_lift_hom_monom UP_cring_axioms X_closed cfs_closed 
      cring.axioms(1) to_fun_X_pow poly_lift_hom_X_var' ring_hom_closed ring_hom_nat_pow)  

lemma(in UP_cring) poly_lift_hom_X_plus:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (X_poly_plus R a) = X_poly_plus S (φ a)"
  using ring_hom_add
  unfolding X_poly_plus_def 
  using P_def X_closed assms(1) assms(2) assms(3) poly_lift_hom_X_var' poly_lift_hom_add poly_lift_hom_extends_hom to_poly_closed by fastforce

lemma(in UP_cring) poly_lift_hom_X_plus_nat_pow:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (X_poly_plus R a [^]UP R(n::nat)) = X_poly_plus S (φ a) [^]UP S(n::nat)"
  using assms poly_lift_hom_X_plus[of S φ a] 
        ring_hom_nat_pow[of "UP R" "UP S" "poly_lift_hom R S φ" "X_poly_plus R a" n] 
        poly_lift_hom_is_hom[of S φ] X_plus_closed[of a] UP_ring.UP_ring[of S]
  unfolding P_def cring_def UP_cring_def
  using P_def UP_ring UP_ring.intro
  by (simp add: UP_ring.intro)
 
lemma(in UP_cring) X_poly_plus_nat_pow_closed:
  assumes "a  carrier R"
  shows " X_poly_plus R a [^]UP R(n::nat)  carrier (UP R)"
  using assms P.nat_pow_closed P_def X_plus_closed by auto

lemma(in UP_cring) poly_lift_hom_X_plus_nat_pow_smult:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "poly_lift_hom R S φ (b UP RX_poly_plus R a [^]UP R(n::nat)) = φ b UP SX_poly_plus S (φ a) [^]UP S(n::nat)"
  by (simp add: X_poly_plus_nat_pow_closed assms(1) assms(2) assms(3) assms(4) poly_lift_hom_X_plus_nat_pow poly_lift_hom_smult)

lemma(in UP_cring) poly_lift_hom_X_minus:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (X_poly_minus R a) = X_poly_minus S (φ a)"
  using poly_lift_hom_X_plus[of S φ " a"] X_minus_plus[of a] UP_cring.X_minus_plus[of S "φ a"]
        R.ring_hom_a_inv[of S φ a]
  unfolding UP_cring_def P_def
  by (metis R.add.inv_closed assms(1) assms(2) assms(3) cring.axioms(1) ring_hom_closed)

lemma(in UP_cring) poly_lift_hom_X_minus_nat_pow:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (X_poly_minus R a [^]UP R(n::nat)) = X_poly_minus S (φ a) [^]UP S(n::nat)"
  using assms poly_lift_hom_X_minus ring_hom_nat_pow X_minus_plus UP_cring.X_minus_plus 
      poly_lift_hom_X_plus poly_lift_hom_X_plus_nat_pow by fastforce

lemma(in UP_cring) X_poly_minus_nat_pow_closed:
  assumes "a  carrier R"
  shows "X_poly_minus R a [^]UP R(n::nat)  carrier (UP R)"
  using assms monoid.nat_pow_closed[of "UP R" "X_poly_minus R a" n] 
        P.nat_pow_closed P_def X_minus_closed by auto

lemma(in UP_cring) poly_lift_hom_X_minus_nat_pow_smult:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "a  carrier R"
  assumes "b  carrier R"
  shows "poly_lift_hom R S φ (b UP RX_poly_minus R a [^]UP R(n::nat)) = φ b UP SX_poly_minus S (φ a) [^]UP S(n::nat)"
  by (simp add: X_poly_minus_nat_pow_closed assms(1) assms(2) assms(3) assms(4) poly_lift_hom_X_minus_nat_pow poly_lift_hom_smult)

lemma(in UP_cring) poly_lift_hom_cf:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier P"
  shows "poly_lift_hom R S φ p k = φ (p k)"
  apply(rule poly_induct3[of p])
  apply (simp add: assms(3))
proof-
  show "p q. q  carrier P 
           p  carrier P 
           poly_lift_hom R S φ p k = φ (p k)  poly_lift_hom R S φ q k = φ (q k)  poly_lift_hom R S φ (p Pq) k = φ ((p Pq) k)"
  proof- fix p q assume A: "p  carrier P"  "q  carrier P"
          "poly_lift_hom R S φ p k = φ (p k)" "poly_lift_hom R S φ q k = φ (q k)"
    show "poly_lift_hom R S φ q k = φ (q k)  poly_lift_hom R S φ (p Pq) k = φ ((p Pq) k)"
      using A assms poly_lift_hom_add[of S φ p q] 
            poly_lift_hom_closed[of S φ p] poly_lift_hom_closed[of S φ q]   
           UP_ring.cfs_closed[of S "poly_lift_hom R S φ q " k] UP_ring.cfs_closed[of S "poly_lift_hom R S φ p" k] 
           UP_ring.cfs_add[of S "poly_lift_hom R S φ p" "poly_lift_hom R S φ q" k]       
      unfolding P_def UP_ring_def  
      by (metis (full_types) P_def cfs_add cfs_closed cring.axioms(1) ring_hom_add)
  qed
  show "a n. a  carrier R  poly_lift_hom R S φ (monom P a n) k = φ (monom P a n k)"
  proof- fix a m assume A: "a  carrier R" 
    show "poly_lift_hom R S φ (monom P a m) k = φ (monom P a m k)"     
      apply(cases "m = k")
    using cfs_monom[of a m k] assms poly_lift_hom_monom[of S φ a m] UP_ring.cfs_monom[of S "φ a" m k] 
       unfolding P_def UP_ring_def  
       apply (simp add: A cring.axioms(1) ring_hom_closed)
           using cfs_monom[of a m k] assms poly_lift_hom_monom[of S φ a m] UP_ring.cfs_monom[of S "φ a" m k] 
       unfolding P_def UP_ring_def  
       by (metis A P_def R.ring_axioms cring.axioms(1) ring_hom_closed ring_hom_zero)     
  qed
qed

lemma(in ring) ring_hom_monom_term:
  assumes "a  carrier R"
  assumes "c  carrier R"
  assumes "ring S"
  assumes "h  ring_hom R S"
  shows "h (a  c[^](n::nat)) = h a S(h c)[^]Sn"
  apply(induction n)
  using assms ringE(2) ring_hom_closed apply fastforce
  by (metis assms(1) assms(2) assms(3) assms(4) local.ring_axioms nat_pow_closed ring_hom_mult ring_hom_nat_pow)

lemma(in UP_cring) poly_lift_hom_eval:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "UP_cring.to_fun  S (poly_lift_hom R S φ p) (φ a) = φ (to_fun p a) "
  apply(rule poly_induct3[of p])
    apply (simp add: assms(3))
proof-
  show "p q. q  carrier P 
           p  carrier P 
           UP_cring.to_fun S (poly_lift_hom R S φ p) (φ a) = φ (to_fun p a) 
           UP_cring.to_fun S (poly_lift_hom R S φ q) (φ a) = φ (to_fun q a) 
           UP_cring.to_fun S (poly_lift_hom R S φ (p Pq)) (φ a) = φ (to_fun (p Pq) a)"
  proof- fix p q  assume A: "q  carrier P"  "p  carrier P"
           "UP_cring.to_fun S (poly_lift_hom R S φ p) (φ a) = φ (to_fun p a)"
           "UP_cring.to_fun S (poly_lift_hom R S φ q) (φ a) = φ (to_fun q a)"
    have "(poly_lift_hom R S φ (p Pq)) = poly_lift_hom R S φ p UP Spoly_lift_hom R S φ q"
      using A(1) A(2) P_def assms(1) assms(2) poly_lift_hom_add by auto
    hence "UP_cring.to_fun S (poly_lift_hom R S φ (p Pq)) (φ a) = 
          UP_cring.to_fun S (poly_lift_hom R S φ p) (φ a) SUP_cring.to_fun S (poly_lift_hom R S φ q) (φ a)"
      using UP_cring.to_fun_plus[of S] assms 
      unfolding UP_cring_def 
      by (metis (no_types, lifting) A(1) A(2) P_def poly_lift_hom_closed ring_hom_closed)
    thus "UP_cring.to_fun S (poly_lift_hom R S φ (p Pq)) (φ a) = φ (to_fun (p Pq) a)"
      using A to_fun_plus assms ring_hom_add[of φ R S]
        poly_lift_hom_closed[of S φ]  UP_cring.to_fun_def[of S] to_fun_def
      unfolding P_def UP_cring_def 
      using UP_cring.to_fun_closed UP_cring_axioms 
      by metis 
  qed
  show "c n. c  carrier R  UP_cring.to_fun  S (poly_lift_hom R S φ (monom P c n)) (φ a) = φ (to_fun (monom P c n) a)"
    unfolding P_def
  proof - fix c n assume A: "c  carrier R"
    have 0: "φ (a [^]R(n::nat)) =  φ a [^]Sn" 
      using assms ring_hom_nat_pow[of R S φ a n] 
      unfolding cring_def 
      using R.ring_axioms by blast
    have 1: "φ (c Ra [^]Rn) = φ c Sφ a [^]Sn"
      using ring_hom_mult[of φ R S c "a [^]Rn" ] 0 assms A monoid.nat_pow_closed [of R a n]
      by (simp add: cring.axioms(1) ringE(2))
    show "UP_cring.to_fun  S (poly_lift_hom R S φ (monom (UP R) c n)) (φ a) = φ (to_fun(monom (UP R) c n) a)"
      using assms A poly_lift_hom_monom[of S φ c n] UP_cring.to_fun_monom[of S "φ c" "φ a" n]
           to_fun_monom[of c a n]  0 1 ring_hom_closed[of φ R S] unfolding UP_cring_def
      by (simp add: P_def to_fun_def)      
  qed
qed

lemma(in UP_cring) poly_lift_hom_sub:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier P"
  assumes "q  carrier P"
  shows "poly_lift_hom R S φ (compose R p q) = compose S (poly_lift_hom R S φ p) (poly_lift_hom R S φ q)"
  apply(rule poly_induct3[of p])
  apply (simp add: assms(3))
proof-
  show " p qa.
       qa  carrier P 
       p  carrier P 
       poly_lift_hom R S φ (Cring_Poly.compose R p q) = Cring_Poly.compose S (poly_lift_hom R S φ p) (poly_lift_hom R S φ q) 
       poly_lift_hom R S φ (Cring_Poly.compose R qa q) = Cring_Poly.compose S (poly_lift_hom R S φ qa) (poly_lift_hom R S φ q) 
       poly_lift_hom R S φ (Cring_Poly.compose R (p Pqa) q) = Cring_Poly.compose S (poly_lift_hom R S φ (p Pqa)) (poly_lift_hom R S φ q)"
  proof- fix a b assume A: "a  carrier P"
       "b  carrier P"
       "poly_lift_hom R S φ (Cring_Poly.compose R a q) = Cring_Poly.compose S (poly_lift_hom R S φ a) (poly_lift_hom R S φ q)"
       "poly_lift_hom R S φ (Cring_Poly.compose R b q) = Cring_Poly.compose S (poly_lift_hom R S φ b) (poly_lift_hom R S φ q)"
    show "poly_lift_hom R S φ (Cring_Poly.compose R (a Pb) q) = Cring_Poly.compose S (poly_lift_hom R S φ (a Pb)) (poly_lift_hom R S φ q)"
      using assms UP_cring.sub_add[of R q a b ] UP_cring.sub_add[of S  ] 
      unfolding UP_cring_def 
      by (metis A(1) A(2) A(3) A(4) P_def R_cring UP_cring.sub_closed UP_cring_axioms poly_lift_hom_add poly_lift_hom_closed)
  qed
  show "a n. a  carrier R 
           poly_lift_hom R S φ (Cring_Poly.compose R (monom P a n) q) = 
            Cring_Poly.compose S (poly_lift_hom R S φ (monom P a n)) (poly_lift_hom R S φ q)"
  proof-
    fix a n assume A: "a  carrier R"
    have 0: "(poly_lift_hom R S φ (monom (UP R) a n)) = monom (UP S) (φ a) n"
      by (simp add: A assms(1) assms(2) assms(3) assms(4) poly_lift_hom_monom)
    have 1: " q [^]UP Rn  carrier (UP R)"
      using monoid.nat_pow_closed[of "UP R" q n] UP_ring.UP_ring UP_ring.intro assms(1) assms
            P.monoid_axioms P_def by blast
    have 2: "poly_lift_hom R S φ (to_polynomial R a UP Rq [^]UP Rn) = 
            to_polynomial S (φ a) UP S(poly_lift_hom R S φ q) [^]UP Sn"
      using poly_lift_hom_mult[of S φ "to_polynomial R a" "q [^]UP Rn"] poly_lift_hom_is_hom[of S φ]
            ring_hom_nat_pow[of P "UP S" "poly_lift_hom R S φ" q n]   UP_cring.UP_cring[of S]
            UP_cring poly_lift_hom_monom[of S φ a 0] ring_hom_closed[of φ R S a] 
            monom_closed[of a 0] nat_pow_closed[of q n] assms A
      unfolding to_polynomial_def P_def UP_cring_def cring_def
      by auto 
    have 3: "poly_lift_hom R S φ (Cring_Poly.compose R (monom (UP R) a n) q) = to_polynomial S (φ a) UP S(poly_lift_hom R S φ q) [^]UP Sn"
      using "2" A P_def assms(4) sub_monom(1) by auto      
    have 4: "Cring_Poly.compose S (poly_lift_hom R S φ (monom (UP R) a n)) (poly_lift_hom R S φ q) 
                              = Cring_Poly.compose S (monom (UP S) (φ a) n) (poly_lift_hom R S φ q)"
      by (simp add: "0")    
    have "poly_lift_hom R S φ q  carrier (UP S)"
      using P_def UP_cring.poly_lift_hom_closed UP_cring_axioms assms(1) assms(2) assms(4) by blast
    then have 5: "Cring_Poly.compose S (poly_lift_hom R S φ (monom (UP R) a n)) (poly_lift_hom R S φ q)
                              = to_polynomial S (φ a) UP S(poly_lift_hom R S φ q) [^]UP Sn"
      using 4 UP_cring.sub_monom[of S "poly_lift_hom R S φ q" "φ a" n] assms 
      unfolding UP_cring_def  
      by (simp add: A ring_hom_closed)
    thus "poly_lift_hom R S φ (Cring_Poly.compose R (monom P a n) q) = 
            Cring_Poly.compose S (poly_lift_hom R S φ (monom P a n)) (poly_lift_hom R S φ q)"
      using 0 1 2 3 4 assms A 
      by (simp add: P_def)
  qed
qed

lemma(in UP_cring) poly_lift_hom_comm_taylor_expansion:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (taylor_expansion R a p) = taylor_expansion S (φ a) (poly_lift_hom R S φ p)"
  unfolding taylor_expansion_def 
  using  poly_lift_hom_sub[of S φ p "(X_poly_plus R a)"] poly_lift_hom_X_plus[of S φ a] assms 
  by (simp add: P_def UP_cring.X_plus_closed UP_cring_axioms)
 
lemma(in UP_cring) poly_lift_hom_comm_taylor_expansion_cf:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  assumes "a  carrier R"
  shows "φ (taylor_expansion R a p i) = taylor_expansion S (φ a) (poly_lift_hom R S φ p) i"
  using poly_lift_hom_cf assms poly_lift_hom_comm_taylor_expansion P_def 
        taylor_def UP_cring.taylor_closed UP_cring_axioms by fastforce
  
lemma(in UP_cring) taylor_expansion_cf_closed:
  assumes "p  carrier P"
  assumes "a  carrier R"
  shows "taylor_expansion R a p i  carrier R"
  using assms taylor_closed
  by (simp add: taylor_def cfs_closed)

lemma(in UP_cring) poly_lift_hom_comm_taylor_term:
  assumes "cring S"
  assumes "φ  ring_hom R S"
  assumes "p  carrier (UP R)"
  assumes "a  carrier R"
  shows "poly_lift_hom R S φ (taylor_term a p i) = UP_cring.taylor_term S (φ a) (poly_lift_hom R S φ p) i"
  using poly_lift_hom_X_minus_nat_pow_smult[of S φ a "taylor_expansion R a p i" i] 
        poly_lift_hom_comm_taylor_expansion[of S φ p a]
        poly_lift_hom_comm_taylor_expansion_cf[of S φ p a i]
        assms UP_cring.taylor_term_def[of S]
    unfolding taylor_term_def UP_cring_def P_def
  by (simp add: UP_cring.taylor_expansion_cf_closed UP_cring_axioms) 

lemma(in UP_cring) poly_lift_hom_degree_bound:
  assumes "cring S"
  assumes "h  ring_hom R S"
  assumes "f  carrier (UP R)"
  shows "deg S (poly_lift_hom R S h f)  deg R f"
    using poly_lift_hom_closed[of S h f]  UP_cring.deg_leqI[of S "poly_lift_hom R S h f" "deg R f"] assms ring_hom_zero[of h R S] deg_aboveD[of f] coeff_simp[of f]
    unfolding P_def UP_cring_def
    by (simp add: P_def R.ring_axioms cring.axioms(1) poly_lift_hom_cf)

lemma(in UP_cring) deg_eqI:
  assumes "f  carrier (UP R)"
  assumes "deg R f  n"
  assumes "f n  𝟬"
  shows "deg R f = n"
  using assms coeff_simp[of f] P_def deg_leE le_neq_implies_less by blast

lemma(in UP_cring) poly_lift_hom_degree_eq:
  assumes "cring S"
  assumes "h  ring_hom R S"
  assumes "h (lcf f)  𝟬S⇙"
  assumes "f  carrier (UP R)"
  shows "deg S (poly_lift_hom R S h f) = deg R f"
  apply(rule UP_cring.deg_eqI)
  using assms unfolding UP_cring_def apply blast
  using poly_lift_hom_closed[of S h f] assms apply blast 
  using poly_lift_hom_degree_bound[of S h f] assms apply blast 
  using assms poly_lift_hom_cf[of S h f] 
  by (metis P_def)

lemma(in UP_cring) poly_lift_hom_lcoeff:
  assumes "cring S"
  assumes "h  ring_hom R S"
  assumes "h (lcf f)  𝟬S⇙"
  assumes "f  carrier (UP R)"
  shows "UP_ring.lcf S (poly_lift_hom R S h f) = h (lcf f)"
  using poly_lift_hom_degree_eq[of S h f] assms 
  by (simp add: P_def poly_lift_hom_cf)

end 

(**************************************************************************************************)
(**************************************************************************************************)
section‹Coefficient List Constructor for Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)

definition list_to_poly where
"list_to_poly R as n = (if n < length as then as!n else 𝟬R)"

context UP_ring
begin

lemma(in UP_ring) list_to_poly_closed:
  assumes "set as  carrier R"
  shows "list_to_poly R as  carrier P"
  apply(rule UP_car_memI[of "length as"])
  apply (simp add: list_to_poly_def)
  by (metis R.zero_closed assms in_mono list_to_poly_def nth_mem)

lemma(in UP_ring) list_to_poly_zero[simp]:
"list_to_poly R [] = 𝟬UP R⇙"
  unfolding list_to_poly_def 
  apply auto 
  by(simp add: UP_def)   

lemma(in UP_domain) list_to_poly_singleton:
  assumes "a  carrier R"
  shows "list_to_poly R [a] = monom P a 0"
  apply(rule ext)
  unfolding list_to_poly_def using assms 
  by (simp add: cfs_monom)
end 

definition cf_list where 
"cf_list R p = map p [(0::nat)..< Suc (deg R p)]"

lemma cf_list_length:
"length (cf_list R p) = Suc (deg R p)"
  unfolding cf_list_def 
  by simp

lemma cf_list_entries:
  assumes "i  deg R p"
  shows "(cf_list R p)!i = p i"
  unfolding cf_list_def 
  by (metis add.left_neutral assms diff_zero less_Suc_eq_le map_eq_map_tailrec nth_map_upt)
  
lemma(in UP_ring) list_to_poly_cf_list_inv:
  assumes "p  carrier (UP R)"
  shows "list_to_poly R (cf_list R p) = p"
proof
  fix x 
  show "list_to_poly R (cf_list R p) x = p x"
    apply(cases "x < degree p")
    unfolding list_to_poly_def 
    using assms cf_list_length[of R p] cf_list_entries[of _ R p]
     apply simp
    by (metis P_def UP_ring.coeff_simp UP_ring_axioms i. i  deg R p  cf_list R p ! i = p i length (cf_list R p) = Suc (deg R p) assms deg_belowI less_Suc_eq_le)
qed

section‹Polynomial Rings over a Subring›

subsection‹Characterizing the Carrier of a Polynomial Ring over a Subring›
lemma(in ring) carrier_update: 
"carrier (Rcarrier := S) = S"
"𝟬(Rcarrier := S)= 𝟬"
"𝟭(Rcarrier := S)= 𝟭"
"(⊕(Rcarrier := S)) = (⊕)"
"(⊗(Rcarrier := S)) = (⊗)"
by auto 


lemma(in UP_cring) poly_cfs_subring:
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "n. g n  S"
  shows "g  carrier (UP (R  carrier := S ))"
  apply(rule UP_cring.UP_car_memI')
  using R.subcringI' R.subcring_iff UP_cring.intro assms(1) subringE(1) apply blast
proof-
  have "carrier (Rcarrier := S) = S"
    using ring.carrier_update  by simp
  then show 0: "x. g x  carrier (Rcarrier := S)"
    using assms by blast
  have 0: "𝟬Rcarrier := S= 𝟬"
    using R.carrier_update(2) by blast
  then show "x. (deg R g) < x  g x = 𝟬Rcarrier := S⇙"
    using UP_car_memE assms(2) by presburger
qed

lemma(in UP_cring) UP_ring_subring:
  assumes "subring S R"
  shows "UP_cring (R  carrier := S )"  "UP_ring (R  carrier := S )"
  using assms unfolding UP_cring_def 
  using R.subcringI' R.subcring_iff subringE(1) apply blast  
  using assms unfolding UP_ring_def 
  using R.subcringI' R.subcring_iff subringE(1)  
  by (simp add: R.subring_is_ring)
  
lemma(in UP_cring) UP_ring_subring_is_ring:
  assumes "subring S R"
  shows "cring (UP (R  carrier := S ))"
  using assms UP_ring_subring[of S] UP_cring.UP_cring[of "Rcarrier := S"] 
  by blast
  
lemma(in UP_cring) UP_ring_subring_add_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "f  carrier (UP (R  carrier := S ))"
  shows "f UP (R  carrier := S )g  carrier (UP (R  carrier := S ))"
  using assms UP_ring_subring_is_ring[of S] 
  by (meson cring.cring_simprules(1))

lemma(in UP_cring) UP_ring_subring_mult_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "f  carrier (UP (R  carrier := S ))"
  shows "f UP (R  carrier := S )g  carrier (UP (R  carrier := S ))"
  using assms UP_ring_subring_is_ring[of S] 
  by (meson cring.carrier_is_subcring subcringE(6))

lemma(in UP_cring) UP_ring_subring_car:
  assumes "subring S R"
  shows "carrier (UP (R  carrier := S )) = {h  carrier (UP R). n. h n  S}"
proof
  show "carrier (UP (Rcarrier := S))  {h  carrier (UP R). n. h n  S}"
  proof
    fix h assume A: "h  carrier (UP (Rcarrier := S))"
    have "h  carrier P"
      apply(rule UP_car_memI[of "deg (Rcarrier := S) h"]) unfolding P_def 
      using UP_cring.UP_car_memE[of "Rcarrier := S" h] R.carrier_update[of S]
            assms UP_ring_subring A apply presburger
      using UP_cring.UP_car_memE[of "Rcarrier := S" h] assms 
      by (metis A R.ring_axioms UP_cring_def carrier (Rcarrier := S) = S cring.subcringI' is_UP_cring ring.subcring_iff subringE(1) subsetD)
    then show "h  {h  carrier (UP R). n. h n  S}"
      unfolding P_def 
      using assms A  UP_cring.UP_car_memE[of "Rcarrier := S" h] R.carrier_update[of S]
      UP_ring_subring by blast
  qed
  show "{h  carrier (UP R). n. h n  S}  carrier (UP (Rcarrier := S))"
  proof fix h assume A: "h  {h  carrier (UP R). n. h n  S}"
    have 0: "h  carrier (UP R)"
      using A by blast 
    have 1: "n. h n  S"
      using A by blast 
    show "h  carrier (UP (Rcarrier := S))"
      apply(rule UP_ring.UP_car_memI[of _ "deg R h"])
      using assms UP_ring_subring[of S] UP_cring.axioms UP_ring.intro cring.axioms(1) apply blast
      using UP_car_memE[of h] carrier_update 0 R.carrier_update(2) apply presburger
      using assms 1  R.carrier_update(1) by blast
  qed
qed

lemma(in UP_cring) UP_ring_subring_car_subset:
  assumes "subring S R"
  shows "carrier (UP (R  carrier := S ))  carrier (UP R)"
proof fix h assume "h  carrier (UP (R  carrier := S ))"
  then show "h  carrier (UP R)"
  using assms UP_ring_subring_car[of S] by blast
qed

lemma(in UP_cring) UP_ring_subring_car_subset':
  assumes "subring S R"
  assumes "h  carrier (UP (R  carrier := S ))"
  shows "h   carrier (UP R)"
  using assms UP_ring_subring_car_subset[of S] by blast 

lemma(in UP_cring) UP_ring_subring_add:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "f  carrier (UP (R  carrier := S ))"
  shows "g UP Rf = g UP (R  carrier := S )f"
proof(rule ext) fix x show "(g UP Rf) x = (g UP (Rcarrier := S)f) x"
  proof-
    have 0: " (g Pf) x = g x  f x"
      using assms cfs_add[of g f x] unfolding P_def 
      using UP_ring_subring_car_subset' by blast
    have 1: "(g UP (Rcarrier := S)f) x = g x Rcarrier := Sf x"
      using   UP_ring.cfs_add[of "R  carrier := S " g f x] UP_ring_subring[of S] assms 
      unfolding UP_ring_def UP_cring_def 
      using R.subring_is_ring by blast
    show ?thesis using 0 1  R.carrier_update(4)[of S]  
    by (simp add: P_def)
  qed
qed

lemma(in UP_cring) UP_ring_subring_deg:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  shows "deg R g = deg (R  carrier := S ) g"
proof-
  have 0: "g  carrier (UP R)"
    using assms UP_ring_subring_car[of S] by blast
  have 1: "deg R g  deg (R  carrier := S ) g"
    using 0 assms UP_cring.UP_car_memE[of "R  carrier := S " g]
             UP_car_memE[of g] P_def R.carrier_update(2) UP_ring_subring deg_leqI by presburger
  have 2: "deg (R  carrier := S ) g  deg R g"
    using 0 assms UP_cring.UP_car_memE[of "R  carrier := S " g]
             UP_car_memE[of g] P_def R.carrier_update(2) UP_ring_subring UP_cring.deg_leqI 
    by metis 

  show ?thesis using 1 2 by presburger 
qed


lemma(in UP_cring) UP_subring_monom:
  assumes "subring S R"
  assumes "a  S"
  shows "up_ring.monom (UP R) a n = up_ring.monom (UP (R  carrier := S )) a n"
proof fix x
  have 0: "a  carrier R"
    using assms subringE(1)  by blast
  have 1: "a  carrier (Rcarrier := S)"
    using assms  by (simp add: assms(2))
  have 2: " up_ring.monom (UP (Rcarrier := S)) a n x = (if n = x then a else 𝟬Rcarrier := S)"
    using 1  assms UP_ring_subring[of S] UP_ring.cfs_monom[of "Rcarrier := S" a n x] UP_cring.axioms UP_ring.intro cring.axioms(1) 
    by blast
  show "up_ring.monom (UP R) a n x = up_ring.monom (UP (Rcarrier := S)) a n x"
    using 0 1 2 cfs_monom[of a n x]  R.carrier_update(2)[of S] unfolding P_def by presburger
qed

lemma(in UP_cring) UP_ring_subring_mult:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "f  carrier (UP (R  carrier := S ))"
  shows "g UP Rf = g UP (R  carrier := S )f"
proof(rule UP_ring.poly_induct3[of "R  carrier := S " f])
  show "UP_ring (Rcarrier := S)"
    by (simp add: UP_ring_subring(2) assms(1))
  show " f  carrier (UP (Rcarrier := S))"
    by (simp add: assms(3))
  show " p q. q  carrier (UP (Rcarrier := S)) 
           p  carrier (UP (Rcarrier := S)) 
           g UP Rp = g UP (Rcarrier := S)p 
           g UP Rq = g UP (Rcarrier := S)q  g UP R(p UP (Rcarrier := S)q) = g UP (Rcarrier := S)(p UP (Rcarrier := S)q)"
  proof- fix p q 
    assume A: " q  carrier (UP (Rcarrier := S))"
           "p  carrier (UP (Rcarrier := S))"
           "g UP Rp = g UP (Rcarrier := S)p"
           "g UP Rq = g UP (Rcarrier := S)q"
    have 0: "p UP (Rcarrier := S)q = p UP Rq"
      using A UP_ring_subring_add[of S p q] 
      by (simp add: assms(1))
    have 1: "g UP R(p UP Rq) = g UP Rp UP Rg UP Rq"
      using 0 A assms  P.r_distr P_def UP_ring_subring_car_subset' by auto
    hence 2:"g UP R(p UP (Rcarrier := S)q) = g UP Rp UP Rg UP Rq"
      using 0 by simp
    have 3: "g UP (Rcarrier := S)(p UP (Rcarrier := S)q) =
          g UP (Rcarrier := S)p UP (Rcarrier := S)g UP (Rcarrier := S)q"
      using 0 A assms semiring.r_distr[of "UP (Rcarrier := S)"]  UP_ring_subring_car_subset' 
      using UP_ring.UP_r_distr UP_ring (Rcarrier := S) by blast
    hence 4: "g UP (Rcarrier := S)(p UP (Rcarrier := S)q) =
          g UP Rp UP (Rcarrier := S)g UP Rq"
      using A by simp 
    hence 5: "g UP (Rcarrier := S)(p UP (Rcarrier := S)q) =
          g UP Rp UP Rg UP Rq"
      using UP_ring_subring_add[of S]
      by (simp add: A(1) A(2) A(3) A(4) UP_ring.UP_mult_closed UP_ring (Rcarrier := S) assms(1) assms(2))
    show "g UP R(p UP (Rcarrier := S)q) = g UP (Rcarrier := S)(p UP (Rcarrier := S)q)"
      by (simp add: "2" "5")
  qed
  show "a n. a  carrier (Rcarrier := S)  g UP Rmonom (UP (Rcarrier := S)) a n = g UP (Rcarrier := S)monom (UP (Rcarrier := S)) a n"
  proof fix a n x assume A: "a  carrier (Rcarrier := S)"
    have 0: "monom (UP (Rcarrier := S)) a n = monom (UP R) a n"
      using A UP_subring_monom assms(1) by auto
    have 1: "g  carrier (UP R)"
      using assms UP_ring_subring_car_subset' by blast
    have 2: "a  carrier R"
      using A assms subringE(1)[of S R] R.carrier_update[of S] by blast       
    show "(g UP Rmonom (UP (Rcarrier := S)) a n) x = (g UP (Rcarrier := S)monom (UP (Rcarrier := S)) a n) x"
    proof(cases "x < n")
      case True
      have T0: "(g UP Rmonom (UP R) a n) x = 𝟬" 
        using 1 2 True cfs_monom_mult[of g a x n] A assms unfolding P_def  by blast
      then show ?thesis using UP_cring.cfs_monom_mult[of "Rcarrier := S" g a x n] 0 A True 
          UP_ring_subring(1) assms(1) assms(2) by auto
    next
      case False
      have F0: "(g UP Rmonom (UP R) a n) x = a  (g (x - n))" 
        using 1 2 False cfs_monom_mult_l[of g a n "x - n"] A assms unfolding P_def by simp
      have F1: "(g UP (Rcarrier := S)monom (UP (Rcarrier := S)) a n) (x - n + n) = a Rcarrier := Sg (x - n)" 
        using 1 2 False UP_cring.cfs_monom_mult_l[of "Rcarrier := S" g a n "x - n"] A assms 
              UP_ring_subring(1) by blast
      hence F2: "(g UP (Rcarrier := S)monom (UP R) a n) (x - n + n) = a  g (x - n)" 
        using UP_subring_monom[of S a n] R.carrier_update[of S] assms 0 by metis         
      show ?thesis using F0 F1 1 2 assms 
        by (simp add: "0" False add.commute add_diff_inverse_nat)
    qed
  qed
qed

lemma(in UP_cring) UP_ring_subring_one:
  assumes "subring S R"
  shows "𝟭UP R= 𝟭UP (R  carrier := S )⇙"
  using UP_subring_monom[of S 𝟭 0] assms P_def R.subcringI' UP_ring.monom_one UP_ring_subring(2) monom_one subcringE(3) by force

lemma(in UP_cring) UP_ring_subring_zero:
  assumes "subring S R"
  shows "𝟬UP R= 𝟬UP (R  carrier := S )⇙"
  using UP_subring_monom[of S 𝟬 0] UP_ring.monom_zero[of "R  carrier := S " 0]  assms monom_zero[of 0]
  UP_ring_subring[of S] subringE(2)[of S R]
  unfolding P_def  
  by (simp add: P_def R.carrier_update(2))

lemma(in UP_cring) UP_ring_subring_nat_pow:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  shows "g[^]UP Rn = g[^]UP (R  carrier := S )(n::nat)"
  apply(induction n)
  using assms apply (simp add: UP_ring_subring_one)
proof-
  fix n::nat
  assume A: "g [^]UP Rn = g [^]UP (Rcarrier := S)n"
  have "Group.monoid (UP (Rcarrier := S)) "
    using assms UP_ring_subring[of S] UP_ring.UP_ring[of "Rcarrier := S"] ring.is_monoid by blast
  hence 0 : " g [^]UP (Rcarrier := S)n  carrier (UP (Rcarrier := S))"
    using monoid.nat_pow_closed[of "UP (R  carrier := S )" g n] assms UP_ring_subring
    unfolding UP_ring_def ring_def by blast
  have 1: "g [^]UP Rn  carrier (UP R)"
    using 0 assms UP_ring_subring_car_subset'[of S] by (simp add: A)
  then have 2: "g [^]UP Rn UP Rg = g [^]UP (Rcarrier := S)n UP (Rcarrier := S)g"
    using assms UP_ring_subring_mult[of S "g [^]UP Rn" g] 
    by (simp add: "0" A)
  then show "g [^]UP RSuc n = g [^]UP (Rcarrier := S)Suc n" 
    by simp
qed

lemma(in UP_cring) UP_subring_compose_monom:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows "compose R (up_ring.monom (UP R) a n) g = compose (R  carrier := S ) (up_ring.monom (UP (R  carrier := S )) a n) g"
proof-
  have g_closed: "g  carrier (UP R)"
    using assms UP_ring_subring_car by blast
  have 0: "a  carrier R"
    using assms subringE(1) by blast  
  have 1: "compose R (up_ring.monom (UP R) a n) g = a UP R(g[^]UP Rn)"
    using monom_sub[of a g n] unfolding P_def 
    using "0" assms(2) g_closed by blast
  have 2: "compose (Rcarrier := S) (up_ring.monom (UP (Rcarrier := S)) a n) g = a UP (Rcarrier := S)g [^]UP (Rcarrier := S)n"
    using assms UP_cring.monom_sub[of "R  carrier := S " a g n] UP_ring_subring[of S] R.carrier_update[of S]
    by blast
  have 3: " g [^]UP (Rcarrier := S)n = g[^]UP Rn"
    using UP_ring_subring_nat_pow[of S g n] 
    by (simp add: assms(1) assms(2))
  have 4: "a UP R(g[^]UP Rn) = a UP (Rcarrier := S)g [^]UP (Rcarrier := S)n"
  proof fix x  
    show "(a UP Rg [^]UP Rn) x = (a UP (Rcarrier := S)g [^]UP (Rcarrier := S)n) x"
    proof-
      have LHS: "(a UP Rg [^]UP Rn) x = a  ((g [^]UP Rn) x)"
        using "0" P.nat_pow_closed P_def cfs_smult g_closed by auto
      have RHS: "(a UP (Rcarrier := S)g [^]UP (Rcarrier := S)n) x = a Rcarrier := S((g [^]UP (Rcarrier := S)n) x)"
      proof-
        have "Group.monoid (UP (Rcarrier := S)) "
          using assms UP_ring_subring[of S] UP_ring.UP_ring[of "Rcarrier := S"] ring.is_monoid by blast
        hence 0 : " g [^]UP (Rcarrier := S)n  carrier (UP (Rcarrier := S))"
          using monoid.nat_pow_closed[of "UP (R  carrier := S )" g n] assms UP_ring_subring
          unfolding UP_ring_def ring_def by blast
        have 1: "g [^]UP (Rcarrier := S)n  carrier (UP (Rcarrier := S))"
          using assms UP_ring_subring[of S] R.carrier_update[of S] 0 by blast
        then show ?thesis using UP_ring.cfs_smult UP_ring_subring assms 
          by (simp add: UP_ring.cfs_smult)
      qed
      show ?thesis using R.carrier_update RHS LHS 3 assms  
        by simp
    qed
  qed
  show ?thesis using 0 1 2 3 4 
    by simp
qed

lemma(in UP_cring) UP_subring_compose:
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "f  carrier (UP R)"
  assumes "n. g n  S"
  assumes "n. f n  S"
  shows "compose R f g = compose (R  carrier := S ) f g"
proof-
  have g_closed: "g  carrier (UP (R  carrier := S ))"
    using assms poly_cfs_subring by blast
  have 0: "n. ( h. h  carrier (UP R)  deg R h  n  h  carrier (UP (R  carrier := S ))  compose R h g = compose (R  carrier := S ) h g)"
  proof-  fix n show "( h. h  carrier (UP R)  deg R h  n  h  carrier (UP (R  carrier := S ))    compose R h g = compose (R  carrier := S ) h g)"
    proof(induction n)
      show "h. h  carrier (UP R)  deg R h  0  h  carrier (UP (Rcarrier := S))  Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
      proof fix h 
        show "h  carrier (UP R)  deg R h  0  h  carrier (UP (Rcarrier := S))  Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
        proof
          assume A: "h  carrier (UP R)  deg R h  0  h  carrier (UP (Rcarrier := S))"
          then have 0: "deg R h = 0"
            by linarith 
          then have 1: "deg (R  carrier := S ) h = 0"
            using A assms UP_ring_subring_deg[of S h] 
            by linarith
          show "Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
            using 0 1 g_closed assms sub_const[of g h] UP_cring.sub_const[of "Rcarrier := S" g h] A P_def UP_ring_subring
            by presburger
        qed
      qed
      show "n. h. h  carrier (UP R)  deg R h  n  h  carrier (UP (Rcarrier := S)) 
             Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g 
         h. h  carrier (UP R)  deg R h  Suc n  h  carrier (UP (Rcarrier := S)) 
             Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
      proof fix n h 
        assume IH: "h. h  carrier (UP R)  deg R h  n  h  carrier (UP (Rcarrier := S)) 
               Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
        show "h  carrier (UP R)  deg R h  Suc n  h  carrier (UP (Rcarrier := S)) 
           Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
        proof assume A: "h  carrier (UP R)  deg R h  Suc n  h  carrier (UP (Rcarrier := S))"
          show "Cring_Poly.compose R h g = Cring_Poly.compose (Rcarrier := S) h g"
          proof(cases "deg R h  n")
            case True
            then show ?thesis using A IH by blast
          next
            case False
            then have F0: "deg R h = Suc n"
              using A by (simp add: A le_Suc_eq)
            then have F1: "deg (Rcarrier := S) h = Suc n"
              using UP_ring_subring_deg[of S h] A  
              by (simp add: h  carrier (UP R)  deg R h  Suc n  h  carrier (UP (Rcarrier := S)) assms(1))
            obtain j where j_def: "j  carrier (UP (Rcarrier := S)) 
          h = j UP (Rcarrier := S)up_ring.monom (UP (Rcarrier := S)) (h (deg (Rcarrier := S) h)) (deg (Rcarrier := S) h) 
          deg (Rcarrier := S) j < deg (Rcarrier := S) h"
              using A UP_ring.ltrm_decomp[of "Rcarrier := S" h] assms UP_ring_subring[of S] 
              F1 by (metis (mono_tags, lifting) F0 False zero_less_Suc)
            have j_closed: "j  carrier (UP R)"  
              using j_def assms UP_ring_subring_car_subset by blast 
            have F2: "deg R j < deg R h"
              using j_def  assms  
              by (metis (no_types, lifting) F0 F1 UP_ring_subring_deg)
            have F3: "(deg (Rcarrier := S) h) = deg R h"
              by (simp add: F0 F1)
            have F30: "h (deg (Rcarrier := S) h)  S "
              using A UP_cring.UP_car_memE[of "Rcarrier := S" h "deg (Rcarrier := S) h"]
              by (metis R.carrier_update(1) UP_ring_subring(1) assms(1))
            hence F4: "up_ring.monom P (h (deg R h)) (deg R h) = 
                  up_ring.monom (UP (Rcarrier := S)) (h (deg (Rcarrier := S) h)) (deg (Rcarrier := S) h)"
              using F3 g_closed j_def UP_subring_monom[of S "h (deg (Rcarrier := S) h)"] assms 
              unfolding  P_def by metis 
            have F5: "compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g = 
                      compose (R  carrier := S ) (up_ring.monom (UP (R  carrier := S )) (h (deg (Rcarrier := S) h)) (deg (Rcarrier := S) h)) g"
              using F0 F1 F2 F3 F4 UP_subring_compose_monom[of S] assms  P_def h (deg (Rcarrier := S) h)  S 
              by (metis g_closed)
            have F5: "compose R j g =  compose (R  carrier := S ) j g"
              using F0 F2 IH UP_ring_subring_car_subset' assms(1) j_def by auto
            have F6: "h = j UP Rmonom (UP R) (h (deg R h)) (deg R h)"
              using j_def F4 UP_ring_subring_add[of S j "up_ring.monom (UP (Rcarrier := S)) (h (deg (Rcarrier := S) h)) (deg (Rcarrier := S) h)"]
                    UP_ring.monom_closed[of "Rcarrier := S" "h (deg (Rcarrier := S) h)" "deg (Rcarrier := S) h"]          
              using P_def UP_ring_subring(2) h (deg (Rcarrier := S) h)  S assms(1) by auto
            have F7: "compose R h g =compose R  j g UP Rcompose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g"
            proof-
              show ?thesis 
              using assms(2) j_closed  F5 sub_add[of g j "up_ring.monom P (h (deg R h)) (deg R h)" ] 
                     F4 F3 F2 F1 g_closed  unfolding P_def              
              by (metis A F6 ltrm_closed P_def)
            qed
            have F8: "compose (R  carrier := S ) h g = compose (R  carrier := S )  j g UP (R  carrier := S )compose (R  carrier := S ) (up_ring.monom (UP (R  carrier := S )) (h (deg (R  carrier := S ) h)) (deg (R  carrier := S ) h)) g"
            proof-
              have 0: " UP_cring (Rcarrier := S)"
                by (simp add: UP_ring_subring(1) assms(1))
              have 1: "monom (UP (Rcarrier := S)) (h (deg R h)) (deg R h)  carrier (UP (Rcarrier := S))"
                using assms 0 F30 UP_ring.monom_closed[of "Rcarrier := S" "h (deg R h)" "deg R h"] R.carrier_update[of S]
                unfolding UP_ring_def UP_cring_def 
                by (simp add: F3 cring.axioms(1))
              show ?thesis
                using 0 1 g_closed j_def UP_cring.sub_add[of "R  carrier := S " g j "monom (UP (Rcarrier := S)) (h (deg R h)) (deg R h)" ]
                using F3 by auto
            qed
            have F9: "compose R  j g  carrier (UP R)"
              by (simp add: UP_cring.sub_closed assms(2) is_UP_cring j_closed)
            have F10: "compose (R  carrier := S ) j g  carrier (UP (R  carrier := S ))"
              using assms j_def UP_cring.sub_closed[of "R  carrier := S "] UP_ring_subring(1) g_closed by blast
            have F11: " compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g  carrier (UP R)"
              using assms j_def UP_cring.sub_closed[of "R  carrier := S "] 
                    UP_ring.monom_closed[of "R  carrier := S "]
              by (simp add: A UP_car_memE(1) UP_cring.rev_sub_closed UP_ring.monom_closed is_UP_cring is_UP_ring sub_rev_sub)
            have F12: " compose (R  carrier := S ) (up_ring.monom (UP (R  carrier := S )) (h (deg (R  carrier := S ) h)) (deg (R  carrier := S ) h)) g  carrier (UP (R  carrier := S ))"
              using assms j_def UP_cring.sub_closed[of "R  carrier := S "] 
                    UP_ring.monom_closed[of "R  carrier := S "] UP_ring_subring[of S]
              using A UP_ring.ltrm_closed g_closed by fastforce
            show ?thesis using F9 F10 F11 F12 F7 F8 F5 UP_ring_subring_add[of S "compose R  j g" "compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g"]
                    assms 
              using F3 F30 UP_subring_compose_monom g_closed by auto
          qed
        qed
      qed
    qed
  qed
  show  ?thesis using 0[of "deg R f"]   
    by (simp add: assms(1) assms(3) assms(5) poly_cfs_subring)
qed   


subsection‹Evaluation over a Subring›

lemma(in UP_cring) UP_subring_eval:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows "to_function R g a = to_function (R  carrier := S ) g a"
  apply(rule UP_ring.poly_induct3[of "R  carrier := S " g] )
  apply (simp add: UP_ring_subring(2) assms(1))
    apply (simp add: assms(2))
proof-
  show "p q. q  carrier (UP (Rcarrier := S)) 
           p  carrier (UP (Rcarrier := S)) 
           to_function R p a = to_function (Rcarrier := S) p a 
           to_function R q a = to_function (Rcarrier := S) q a 
           to_function R (p UP (Rcarrier := S)q) a = to_function (Rcarrier := S) (p UP (Rcarrier := S)q) a"
  proof- fix p q assume A: "q  carrier (UP (Rcarrier := S))"
           "p  carrier (UP (Rcarrier := S))"
          " to_function R p a = to_function (Rcarrier := S) p a"
          " to_function R q a = to_function (Rcarrier := S) q a"
    have a_closed: "a  carrier R"
      using  assms R.carrier_update[of S] subringE(1) by blast
    have 0: "UP_cring (Rcarrier := S)"
      using assms by (simp add: UP_ring_subring(1))
    have 1: "to_function (Rcarrier := S) p a  S"
      using A 0 UP_cring.to_fun_closed[of "Rcarrier := S"]  
      by (simp add: UP_cring.to_fun_def assms(3))
    have 2: "to_function (Rcarrier := S) q a  S"
      using A 0 UP_cring.to_fun_closed[of "Rcarrier := S"]  
      by (simp add: UP_cring.to_fun_def assms(3))
    have 3: "p  carrier (UP R)"
      using A  assms 0 UP_ring_subring_car_subset' by blast
    have 4: "q  carrier (UP R)"
      using A  assms 0 UP_ring_subring_car_subset' by blast
    have 5: "to_fun p a  to_fun q a =  UP_cring.to_fun (Rcarrier := S) p a Rcarrier := SUP_cring.to_fun (Rcarrier := S) q a"
      using 1 2 A R.carrier_update[of S] assms by (simp add: "0" UP_cring.to_fun_def to_fun_def)
    have 6: "UP_cring.to_fun (Rcarrier := S) (p UP (Rcarrier := S)q) a =
    UP_cring.to_fun (Rcarrier := S) p a Rcarrier := SUP_cring.to_fun (Rcarrier := S) q a"
      using UP_cring.to_fun_plus[of "R  carrier := S " q p a] 
      by (simp add: "0" A(1) A(2) assms(3))
    have 7: "to_fun (p Pq) a = to_fun p a  to_fun q a"
      using to_fun_plus[of q p a] 3 4 a_closed by (simp add: P_def)
    have 8: "p UP (Rcarrier := S)q = p Pq"
      unfolding P_def using assms A R.carrier_update[of S] UP_ring_subring_add[of S p q] by simp
    show "to_function R (p UP (Rcarrier := S)q) a = to_function (Rcarrier := S) (p UP (Rcarrier := S)q) a"
      using UP_ring_subring_car_subset'[of S ] 0 1 2 3 4 5 6 7 8 A R.carrier_update[of S] 
      unfolding P_def by (simp add: UP_cring.to_fun_def to_fun_def)
  qed
  show "b n.
       b  carrier (Rcarrier := S) 
       to_function R (monom (UP (Rcarrier := S)) b n) a = to_function (Rcarrier := S) (monom (UP (Rcarrier := S)) b n) a"  
  proof- fix b n assume A: "b  carrier (Rcarrier := S)"
    have 0: "UP_cring (Rcarrier := S)"
      by (simp add: UP_ring_subring(1) assms(1))
    have a_closed: "a  carrier R"
      using assms subringE by blast 
    have 1: "UP_cring.to_fun (Rcarrier := S) (monom (UP (Rcarrier := S)) b n) a = b Rcarrier := Sa [^]Rcarrier := Sn"
      using assms A UP_cring.to_fun_monom[of "Rcarrier := S" b a n]  
      by (simp add: "0")
    have 2: "UP_cring.to_fun (Rcarrier := S) (monom (UP (Rcarrier := S)) b n)  to_function (Rcarrier := S) (monom (UP (Rcarrier := S)) b n)"
      using UP_cring.to_fun_def[of "Rcarrier := S" "monom (UP (Rcarrier := S)) b n"] 0 by linarith
    have 3: "(monom (UP (Rcarrier := S)) b n) = monom P b n"
      using A assms unfolding P_def using  UP_subring_monom by auto
    have 4:  " b  a [^] n = b Rcarrier := Sa [^]Rcarrier := Sn" 
      apply(induction n) using R.carrier_update[of S] apply simp
      using R.carrier_update[of S] R.nat_pow_consistent by auto
    hence 5: "to_function R (monom (UP (Rcarrier := S)) b n) a = b Rcarrier := Sa[^]Rcarrier := Sn"
      using 0 1 2 3 assms A UP_cring.to_fun_monom[of "Rcarrier := S" b a n] UP_cring.to_fun_def[of "Rcarrier := S" "monom (UP (Rcarrier := S)) b n"]
             R.carrier_update[of S] subringE[of S R] a_closed UP_ring.monom_closed[of "Rcarrier := S" a n]  
             to_fun_monom[of b a n]
      unfolding P_def UP_cring.to_fun_def to_fun_def by (metis subsetD)
    thus " to_function R (monom (UP (Rcarrier := S)) b n) a = to_function (Rcarrier := S) (monom (UP (Rcarrier := S)) b n) a"
      using "1" "2" by auto
  qed
qed

lemma(in UP_cring) UP_subring_eval':
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows "to_fun g a = to_function (R  carrier := S ) g a"
  unfolding to_fun_def using assms 
  by (simp add: UP_subring_eval)

lemma(in UP_cring) UP_subring_eval_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows "to_fun g a  S"
  using assms  UP_subring_eval'[of S g a] UP_cring.to_fun_closed  UP_cring.to_fun_def R.carrier_update(1) UP_ring_subring(1) by fastforce

subsection‹Derivatives and Taylor Expansions over a Subring›


lemma(in UP_cring) UP_subring_taylor:
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "n. g n  S"
  assumes "a  S"
  shows "taylor_expansion R a g = taylor_expansion (R  carrier := S ) a g"
proof-
  have a_closed: "a  carrier R"
    using assms subringE by blast
  have 0: "X_plus a  carrier (UP R)"
    using assms X_plus_closed  unfolding P_def  
    using local.a_closed by auto
  have 1: "n. X_plus a n  S"
  proof- fix n
    have "X_plus a n = (if n = 0 then a else 
                        (if n = 1 then 𝟭 else 𝟬))"
      using a_closed 
      by (simp add: cfs_X_plus)
    then show "X_plus a n  S" using subringE assms 
      by (simp add: subringE(2) subringE(3))
  qed
  have 2: "(X_poly_plus (Rcarrier := S) a) = X_plus a"
  proof-
    have 20: "(X_poly_plus (Rcarrier := S) a) = (λk. if k = (0::nat) then a else 
                        (if k = 1 then 𝟭 else 𝟬))"
      using a_closed assms UP_cring.cfs_X_plus[of "Rcarrier := S" a] R.carrier_update 
            UP_ring_subring(1) by auto
    have 21: "X_plus a = (λk. if k = (0::nat) then a else 
                        (if k = 1 then 𝟭 else 𝟬))"
       using cfs_X_plus[of a] a_closed 
       by blast
     show ?thesis apply(rule ext) using 20 21
       by auto
  qed     
  show ?thesis
    unfolding taylor_expansion_def using 0 1 2 assms UP_subring_compose[of S g "X_plus a"] 
    by (simp add: UP_subring_compose)
qed

lemma(in UP_cring) UP_subring_taylor_closed:
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "n. g n  S"
  assumes "a  S"
  shows "taylor_expansion R a g  carrier  (UP (R  carrier := S ))"
proof-
  have "g  carrier (UP (Rcarrier := S))"
    by (metis P_def R.carrier_update(1) R.carrier_update(2) UP_cring.UP_car_memI' UP_ring_subring(1) assms(1) assms(2) assms(3) deg_leE)
  then show ?thesis   
    using assms UP_cring.taylor_def[of "Rcarrier := S"] UP_subring_taylor[of S g a]
          UP_cring.taylor_closed[of "R  carrier := S " g a] UP_ring_subring(1)[of S] by simp
qed

lemma(in UP_cring) UP_subring_taylor_closed':
  assumes "subring S R"
  assumes "g  carrier  (UP (R  carrier := S ))"
  assumes "a  S"
  shows "taylor_expansion R a g  carrier  (UP (R  carrier := S ))"
  using UP_subring_taylor_closed assms UP_cring.UP_car_memE[of "R  carrier := S " g] R.carrier_update[of S]
        UP_ring_subring(1) UP_ring_subring_car_subset' by auto

lemma(in UP_cring) UP_subring_taylor':
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "n. g n  S"
  assumes "a  S"
  shows "taylor_expansion R a g n  S"
  using assms UP_subring_taylor R.carrier_update[of S] UP_cring.taylor_closed[of "R  carrier := S "]
  using UP_cring.taylor_expansion_cf_closed UP_ring_subring(1) poly_cfs_subring by metis 


lemma(in UP_cring) UP_subring_deriv:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows  "deriv g a= UP_cring.deriv (R  carrier := S ) g a"
proof-
  have 0: "(n. g n  S)"
    using assms UP_ring_subring_car by blast
  thus ?thesis 
    unfolding derivative_def using 0 UP_ring_subring_car_subset[of S] assms UP_subring_taylor[of S g a] 
    by (simp add: subset_iff)
qed

lemma(in UP_cring) UP_subring_deriv_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  shows  "deriv g a  S"
  using assms  UP_cring.deriv_closed[of "R  carrier := S " g a] UP_subring_deriv[of S g a]
        UP_ring_subring_car_subset[of S] UP_ring_subring[of S]
  by simp

lemma(in UP_cring) poly_shift_subring_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  shows "poly_shift g  carrier (UP (R  carrier := S ))"
  using UP_cring.poly_shift_closed[of "R  carrier := S " g] assms UP_ring_subring[of S]
  by simp

lemma(in UP_cring) UP_subring_taylor_appr:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  assumes "b  S"
  shows "c  S. to_fun g a= to_fun g b  (deriv g b) (a  b)   (c  (a   b)[^](2::nat))"
proof-
  have  a_closed: "a  carrier R"
    using assms subringE by blast
  have b_closed: "b  carrier R"
    using assms subringE by blast
  have g_closed:  " g  carrier (UP R)"
    using UP_ring_subring_car_subset[of S] assms by blast 
  have 0: "to_fun (shift 2 (T⇘bg)) (a  b) = to_fun (shift 2 (T⇘bg)) (a  b)"
    by simp
  have 1: "to_fun g b = to_fun g b"
    by simp
  have 2: "deriv g b = deriv g b"
    by simp 
  have 3: "to_fun g a = to_fun g b  deriv g b  (a  b)  to_fun (shift 2 (T⇘bg)) (a  b)  (a  b) [^] (2::nat)"
    using taylor_deg_1_expansion[of g b a "to_fun (shift 2 (T⇘bg)) (a  b)" "to_fun g b" "deriv g b"]
          assms a_closed b_closed g_closed 0 1 2 unfolding P_def by blast
  have 4: "to_fun (shift 2 (T⇘bg)) (a  b)  S"
  proof-
    have 0: "(2::nat) = Suc (Suc 0)"
      by simp
    have 1: "a  b  S"
      using assms unfolding a_minus_def  
      by (simp add: subringE(5) subringE(7))
    have 2: "poly_shift (T⇘bg)  carrier (UP (Rcarrier := S))"
      using poly_shift_subring_closed[of S "taylor_expansion R b g"] UP_ring_subring[of S]
            UP_subring_taylor_closed'[of S g b] assms       unfolding taylor_def
      by blast
    hence 3: "poly_shift (poly_shift (T⇘bg))  carrier (UP (Rcarrier := S))"
      using    UP_cring.poly_shift_closed[of "Rcarrier := S" "(poly_shift (T⇘bg))"]            
      unfolding taylor_def 
      using assms(1) poly_shift_subring_closed by blast
    have 4: "to_fun (poly_shift (poly_shift (T⇘bg))) (a  b)  S"
      using 1 2 3 0 UP_subring_eval_closed[of S "poly_shift (poly_shift (T⇘bg))"  "a  b"] 
              UP_cring.poly_shift_closed[of "Rcarrier := S"] assms
      by blast
    then show ?thesis 
      by (simp add: numeral_2_eq_2)
  qed
  obtain c where c_def: "c = to_fun (shift 2 (T⇘bg)) (a  b)"
    by blast 
  have 5: "c  S  to_fun g a = to_fun g b  deriv g b  (a  b)  c  (a  b) [^] (2::nat)"
    unfolding c_def using 3 4 by blast 
  thus ?thesis using c_def  4 by blast 
qed

lemma(in UP_cring) UP_subring_taylor_appr':
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  assumes "a  S"
  assumes "b  S"
  shows "c c' c''. c  S  c'  S  c''  S  to_fun g a= c  c' (a  b)   (c''  (a   b)[^](2::nat))"
  using UP_subring_taylor_appr[of S g a b] assms UP_subring_deriv_closed[of S g b] UP_subring_eval_closed[of S g b]
  by blast 

lemma (in UP_cring) pderiv_cfs: 
  assumes"g  carrier (UP R)"
  shows "pderiv g n = [Suc n](g (Suc n))"
  unfolding pderiv_def  
  using n_mult_closed[of g] assms  poly_shift_cfs[of "n_mult g" n] 
  unfolding P_def n_mult_def by blast

lemma(in ring) subring_add_pow:
  assumes "subring S R"
  assumes "a  S"
  shows "[(n::nat)] Rcarrier := Sa = [(n::nat)] a"
proof-
  have 0: "a  carrier R"
    using assms(1) assms(2) subringE(1) by blast
  have 1: "a  carrier (Rcarrier := S)"
    by (simp add: assms(2))
  show ?thesis
  apply(induction n)
  using assms 0 1 carrier_update[of S] 
  apply (simp add: add_pow_def)
  using assms 0 1 carrier_update[of S] 
  by (simp add: add_pow_def)
qed

lemma(in UP_cring) UP_subring_pderiv_equal:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  shows "pderiv g = UP_cring.pderiv (Rcarrier := S) g"
proof fix n
  show "pderiv g n = UP_cring.pderiv (Rcarrier := S) g n"
    using UP_cring.pderiv_cfs[of "R  carrier := S " g n] pderiv_cfs[of g n]
          assms R.subring_add_pow[of S "g (Suc n)" "Suc n"]
    by (simp add: UP_ring_subring(1) UP_ring_subring_car)
qed
  
lemma(in UP_cring) UP_subring_pderiv_closed:
  assumes "subring S R"
  assumes "g  carrier (UP (R  carrier := S ))"
  shows "pderiv g  carrier (UP (R  carrier := S ))"
  using assms UP_cring.pderiv_closed[of "R  carrier := S " g] R.carrier_update(1) UP_ring_subring(1)
UP_subring_pderiv_equal by auto

lemma(in UP_cring) UP_subring_pderiv_closed':
  assumes "subring S R"
  assumes "g  carrier (UP R)"
  assumes "n. g n  S"
  shows "n. pderiv g n  S"
  using assms UP_subring_pderiv_closed[of S g] poly_cfs_subring[of S g] UP_ring_subring_car
  by blast

lemma(in UP_cring) taylor_deg_one_expansion_subring:
  assumes "f  carrier (UP R)"
  assumes "subring S R"
  assumes "i. f i  S"
  assumes "a  S"
  assumes "b  S"
  shows "c  S. to_fun f b = (to_fun f a)  (deriv f a)  (b  a)  (c  (b  a)[^](2::nat))"
  apply(rule UP_subring_taylor_appr, rule assms) 
  using assms poly_cfs_subring apply blast
  by(rule assms, rule assms)
  
lemma(in UP_cring) taylor_deg_one_expansion_subring':
  assumes "f  carrier (UP R)"
  assumes "subring S R"
  assumes "i. f i  S"
  assumes "a  S"
  assumes "b  S"
  shows "c  S. to_fun f b = (to_fun f a)  (to_fun (pderiv f) a)  (b  a)  (c  (b  a)[^](2::nat))"
proof-
  have "S  carrier R"
    using assms subringE(1) by blast
  hence 0: "deriv f a = to_fun (pderiv f) a"
    using assms  pderiv_eval_deriv[of f a] unfolding P_def  by blast 
  show ?thesis 
    using assms taylor_deg_one_expansion_subring[of f S a b]
    unfolding 0 by blast 
qed

end