Theory Padic_Semialgebraic_Function_Ring

theory Padic_Semialgebraic_Function_Ring
  imports Padic_Field_Powers
begin

section‹Rings of Semialgebraic Functions›

text‹
  In order to efficiently formalize Denef's proof of Macintyre's theorem, it is necessary to be
  able to reason about semialgebraic functions algebraically. For example, we need to consider
  polynomials in one variable whose coefficients are semialgebraic functions, and take their
  Taylor expansions centered at a semialgebraic function. To facilitate this kind of reasoning, it
  is necessary to construct, for each arity $m$, a ring \texttt{SA(m)} of semialgebraic functions in
  $m$ variables. These functions must be extensional functions which are undefined outside of the
  carrier set of $\mathbb{Q}_p^m$.

  The developments in this theory are mainly lemmas and defintitions which build the necessary
  theory to prove the cell decomposition theorems of cite"denef1986", and finally Macintyre's
  Theorem, which says that semi-algebraic sets are closed under projections.
›
(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Some eint Arithmetic›
(**************************************************************************************************)
(**************************************************************************************************)

context padic_fields
begin

lemma eint_minus_ineq':
  assumes "a  eint N"
  assumes "b - a  c"
  shows "b - eint N  c"
  using assms by(induction c, induction b, induction a, auto )

lemma eint_minus_plus:
"a - (eint b + eint c) = a - eint b - eint c"
  apply(induction a)
  apply (metis diff_add_eq_diff_diff_swap idiff_eint_eint plus_eint_simps(1) semiring_normalization_rules(24))
  using idiff_infinity by presburger

lemma eint_minus_plus':
"a - (eint b + eint c) = a - eint c - eint b"
  by (metis add.commute eint_minus_plus)

lemma eint_minus_plus'':
  assumes "a - eint c - eint b = eint f"
  shows "a - eint c - eint f =  eint b"
  using assms  apply(induction a )
  apply (metis add.commute add_diff_cancel_eint eint.distinct(2) eint_add_cancel_fact)
  by simp

lemma uminus_involutive[simp]:
"-(-x::eint) = x"
  apply(induction x)
  unfolding uminus_eint_def by auto

lemma eint_minus:
"(a::eint) - (b::eint) = a + (-b)"
  apply(induction a)
   apply(induction b)
proof -
  fix int :: int and inta :: int
  have "e ea. (ea::eint) + (e + - ea) = ea - ea + e"
    by (simp add: eint_uminus_eq)
  then have "i ia. eint (ia + i) + - eint ia = eint i"
    by (metis ab_group_add_class.ab_diff_conv_add_uminus add.assoc add_minus_cancel idiff_eint_eint plus_eint_simps(1))
  then show "eint inta - eint int = eint inta + - eint int"
    by (metis ab_group_add_class.ab_diff_conv_add_uminus add.commute add_minus_cancel idiff_eint_eint)
next
  show "int. eint int -  = eint int + - "
  by (metis eint_uminus_eq i0_ne_infinity idiff_infinity_right idiff_self plus_eq_infty_iff_eint uminus_involutive)
  show "  - b =  + - b"
    apply(induction b)
  apply simp
    by auto
qed

lemma eint_mult_Suc:
   "eint (Suc k) * a = eint k * a + a"
  apply(induction a)
  apply (metis add.commute eSuc_eint mult_eSuc' of_nat_Suc)
  using plus_eint_simps(3) times_eint_simps(4)
  by presburger

lemma eint_mult_Suc_mono:
assumes "a  eint b  eint (int k) * a  eint (int k) * eint b"
shows "a  eint b  eint (int (Suc k)) * a  eint (int (Suc k)) * eint b"
  using assms eint_mult_Suc
  by (metis add_mono_thms_linordered_semiring(1))

lemma eint_nat_mult_mono:
  assumes "(a::eint)  b"
  shows "eint (k::nat)*a  eint k*b"
proof-
  have "(a::eint)  b  eint (k::nat)*a  eint k*b"
    apply(induction k) apply(induction b)
  apply (metis eint_ile eq_iff mult_not_zero of_nat_0 times_eint_simps(1))
  apply simp
 apply(induction b)
    using eint_mult_Suc_mono apply blast
  using eint_ord_simps(3) times_eint_simps(4) by presburger
  thus ?thesis using assms by blast
qed

lemma eint_Suc_zero:
"eint (int (Suc 0)) * a = a"
  apply(induction a)
  apply simp
  by simp

lemma eint_add_mono:
  assumes "(a::eint)  b"
  assumes "(c::eint)  d"
  shows "a + c  b + d"
  using assms
  by (simp add: add_mono)

lemma eint_nat_mult_mono_rev:
  assumes "k > 0"
  assumes "eint (k::nat)*a  eint k*b"
  shows "(a::eint)  b"
proof(rule ccontr)
  assume "¬ a  b"
  then have A: "b < a"
    using leI by blast
  have "b < a  eint (k::nat)*b < eint k*a"
    apply(induction b) apply(induction a)
      using A assms eint_ord_simps(2) times_eint_simps(1) zmult_zless_mono2_lemma apply presburger
    using eint_ord_simps(4) nat_mult_not_infty times_eint_simps(4) apply presburger
    using eint_ord_simps(6) by blast
  then have "eint (k::nat)*b < eint k*a"
    using A by blast
  hence "¬ eint (k::nat)*a  eint k*b"
    by (metis ¬ a  b antisym eint_nat_mult_mono linear neq_iff)
  then show False using assms by blast
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Lemmas on Function Ring Operations›
(**************************************************************************************************)
(**************************************************************************************************)

lemma Qp_funs_is_cring:
"cring (Fun⇘nQp)"
  using F.M_cring by blast

lemma Qp_funs_is_monoid:
"monoid (Fun⇘nQp)"
  using F.is_monoid by blast

lemma Qp_funs_car_memE:
  assumes "f  carrier (Fun⇘nQp)"
  shows "f  (carrier (Qpn))  (carrier Qp)"
  by (simp add: assms ring_pow_function_ring_car_memE(2))

lemma Qp_funs_car_memI:
  assumes "g  carrier (Qpn)  carrier Qp"
  assumes "x. x  (carrier (Qpn))  g x = undefined"
  shows "g  carrier (Fun⇘nQp)"
apply(rule Qp.function_ring_car_memI)
  using assms  apply blast
  using assms by blast

lemma Qp_funs_car_memI':
  assumes "g  carrier (Qpn)  carrier Qp"
  assumes "restrict g (carrier (Qpn)) = g"
  shows "g  carrier (Fun⇘nQp)"
  apply(intro Qp_funs_car_memI assms)
  using assms  unfolding restrict_def
  by (metis (mono_tags, lifting))

lemma Qp_funs_car_memI'':
  assumes "f  carrier (Qpn)  carrier Qp"
  assumes "g = (λ x  (carrier (Qpn)). f x)"
  shows "g  carrier (Fun⇘nQp)"
  apply(rule Qp_funs_car_memI)
  using assms
  apply (meson restrict_Pi_cancel)
  by (metis assms(2) restrict_def)

lemma Qp_funs_one:
"𝟭⇘Fun⇘nQp= (λ x  carrier (Qpn). 𝟭)"
  unfolding function_ring_def function_one_def
  by (meson monoid.select_convs(2))

lemma Qp_funs_zero:
"𝟬⇘Fun⇘nQp= (λ x  carrier (Qpn). 𝟬Qp)"
  unfolding function_ring_def function_zero_def
  by (meson ring_record_simps(11))

lemma Qp_funs_add:
  assumes "x  carrier (Qpn)"
  assumes "f  (carrier (Qpn))  carrier Qp"
  assumes "g  (carrier (Qpn))  carrier Qp"
  shows "(f ⇘Fun⇘nQpg) x = f x Qpg x"
  using assms function_ring_def[of "carrier (Qpn)" "Qp"]
  unfolding function_add_def
  by (metis (mono_tags, lifting) restrict_apply' ring_record_simps(12))

lemma Qp_funs_add':
  assumes "x  carrier (Qpn)"
  assumes "f  (carrier (Fun⇘nQp))"
  assumes "g  (carrier (Fun⇘nQp))"
  shows "(f ⇘Fun⇘nQpg) x = f x Qpg x"
  using assms Qp_funs_add Qp_funs_car_memE
  by blast

lemma Qp_funs_add'':
  assumes "f  (carrier (Fun⇘nQp))"
  assumes "g  (carrier (Fun⇘nQp))"
  shows "(f ⇘Fun⇘nQpg) = (λ x  carrier (Qpn). f x Qpg x)"
  unfolding function_ring_def function_add_def using ring_record_simps(12)
  by metis

lemma Qp_funs_add''':
  assumes "x  carrier (Qpn)"
  shows "(f ⇘Fun⇘nQpg) x = f x Qpg x"
  using assms function_ring_def[of "carrier (Qpn)" "Qp"]
  unfolding function_add_def
  by (metis (mono_tags, lifting) restrict_apply' ring_record_simps(12))

lemma Qp_funs_mult:
  assumes "x  carrier (Qpn)"
  assumes "f  (carrier (Qpn))  carrier Qp"
  assumes "g  (carrier (Qpn))  carrier Qp"
  shows "(f ⇘Fun⇘nQpg) x = f x  g x"
  using assms function_ring_def[of "carrier (Qpn)" "Qp"]
  unfolding function_mult_def
  by (metis (no_types, lifting) monoid.select_convs(1) restrict_apply')

lemma Qp_funs_mult':
  assumes "x  carrier (Qpn)"
  assumes "f  (carrier (Fun⇘nQp))"
  assumes "g  (carrier (Fun⇘nQp))"
  shows "(f ⇘Fun⇘nQpg) x = f x  g x"
  using assms Qp_funs_mult Qp_funs_car_memE
  by blast

lemma Qp_funs_mult'':
  assumes "f  (carrier (Fun⇘nQp))"
  assumes "g  (carrier (Fun⇘nQp))"
  shows "(f ⇘Fun⇘nQpg) = (λ x  carrier (Qpn). f x  g x)"
  unfolding function_ring_def function_mult_def using ring_record_simps(5)
  by metis

lemma Qp_funs_mult''':
  assumes "x  carrier (Qpn)"
  shows "(f ⇘Fun⇘nQpg) x = f x  g x"
  using assms function_ring_def[of "carrier (Qpn)" "Qp"]
  unfolding function_mult_def
  by (metis (mono_tags, lifting) monoid.select_convs(1) restrict_apply')

lemma Qp_funs_a_inv:
  assumes "x  carrier (Qpn)"
  assumes "f  (carrier (Fun⇘nQp))"
  shows "(⇘Fun⇘nQpf) x =  (f x)"
  using assms local.function_uminus_eval
  by (simp add: local.function_uminus_eval'')

lemma Qp_funs_a_inv':
  assumes "f  (carrier (Fun⇘nQp))"
  shows "(⇘Fun⇘nQpf) = (λ x  carrier (Qpn).  (f x))"
proof fix x
  show "(⇘Fun⇘nQpf) x = (λxcarrier (Qpn).  f x) x"
    apply(cases "x  carrier (Qpn)")
     apply (metis (no_types, lifting) Qp_funs_a_inv assms restrict_apply')
  by (simp add: assms local.function_ring_not_car)
qed

abbreviation(input) Qp_const ("𝔠⇘__") where
"Qp_const n c  constant_function (carrier (Qpn)) c"

lemma Qp_constE:
  assumes "c  carrier Qp"
  assumes "x  carrier (Qpn)"
  shows "Qp_const n c x = c"
  using assms unfolding constant_function_def
  by (meson restrict_apply')

lemma Qp_funs_Units_memI:
  assumes "f  (carrier (Fun⇘nQp))"
  assumes " x. x  carrier (Qpn)   f x  𝟬Qp⇙"
  shows "f  (Units (Fun⇘nQp))"
        "inv⇘Fun⇘nQpf = (λ x  carrier (Qpn). invQp(f x))"
proof-
  obtain g where g_def: "g = (λ x  carrier (Qpn). invQp(f x))"
    by blast
  have g_closed: "g  (carrier (Fun⇘nQp))"
    by(rule Qp_funs_car_memI, unfold  g_def, auto,
            intro field_inv(3) assms Qp.function_ring_car_memE[of _ n], auto )
  have "x. x  carrier (Qpn)  f x  g x = 𝟭"
    using assms g_def
    by (metis (no_types, lifting) Qp.function_ring_car_memE field_inv(2) restrict_apply)
  then have 0: "f ⇘Fun⇘nQpg = 𝟭⇘Fun⇘nQp⇙"
    using assms g_def Qp_funs_mult''[of f n g] Qp_funs_one[of n] g_closed
    by (metis (no_types, lifting) restrict_ext)
  then show "f  (Units (Fun⇘nQp))"
    using comm_monoid.UnitsI[of "Fun⇘nQp"] assms(1) g_closed local.F.comm_monoid_axioms by presburger
  have "inv⇘Fun⇘nQpf = g"
    using g_def g_closed 0 cring.invI[of "Fun⇘nQp"] Qp_funs_is_cring assms(1)
    by presburger
  show "inv⇘Fun⇘nQpf = (λ x  carrier (Qpn). invQp(f x))"
    using assms g_def 0  inv⇘Fun⇘nQpf = g
    by blast
qed

lemma Qp_funs_Units_memE:
  assumes "f  (Units (Fun⇘nQp))"
  shows "f ⇘Fun⇘nQpinv⇘Fun⇘nQpf = 𝟭⇘Fun⇘nQp⇙"
        "inv⇘Fun⇘nQpf ⇘Fun⇘nQpf = 𝟭⇘Fun⇘nQp⇙"
        " x. x  carrier (Qpn)   f x  𝟬Qp⇙"
  using monoid.Units_r_inv[of "Fun⇘nQp" f ] assms  Qp_funs_is_monoid
    apply blast
  using monoid.Units_l_inv[of "Fun⇘nQp" f ] assms  Qp_funs_is_monoid
    apply blast
proof-
  obtain g where g_def: "g = inv⇘Fun⇘nQpf"
    by blast
  show " x. x  carrier (Qpn)   f x  𝟬Qp⇙"
  proof- fix x assume A: "x  carrier (Qpn)"
    have "f ⇘Fun⇘nQpg = 𝟭⇘Fun⇘nQp⇙"
      using assms g_def Qp_funs_is_monoid
        Group.monoid (Fun⇘nQp); f  Units (Fun⇘nQp)  f ⇘Fun⇘nQpinv⇘Fun⇘nQpf = 𝟭⇘Fun⇘nQp⇙›
      by blast
    then have 0: "f x  g x = 𝟭"
      using A assms g_def Qp_funs_mult'[of x n f g] Qp_funs_one[of n]
      by (metis Qp_funs_is_monoid monoid.Units_closed monoid.Units_inv_closed restrict_apply)
    have 1: "g x  carrier Qp"
      using g_def A assms local.function_ring_car_closed by auto
    then show "f x  𝟬Qp⇙"
      using 0
      by (metis Qp.l_null local.one_neq_zero)
  qed
qed

lemma Qp_funs_m_inv:
  assumes "x  carrier (Qpn)"
  assumes "f  (Units (Fun⇘nQp))"
  shows "(inv⇘Fun⇘nQpf) x = invQp(f x)"
  using Qp_funs_Units_memI(2)  Qp_funs_Units_memE(3) assms
  by (metis (no_types, lifting) Qp_funs_is_monoid monoid.Units_closed restrict_apply)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Defining the Rings of Semialgebraic Functions›
(**************************************************************************************************)
(**************************************************************************************************)


definition semialg_functions  where
"semialg_functions n  = {f  (carrier (Qpn))  carrier Qp. is_semialg_function n f  f = restrict f (carrier (Qpn))}"

lemma semialg_functions_memE:
  assumes "f  semialg_functions n"
  shows "is_semialg_function n f"
        "f  carrier (Fun⇘nQp)"
        "f  carrier (Qpn)  carrier Qp"
  using semialg_functions_def assms apply blast
  apply(rule Qp_funs_car_memI')
  using assms
  apply (metis (no_types, lifting) mem_Collect_eq semialg_functions_def)
  using assms unfolding semialg_functions_def
   apply (metis (mono_tags, lifting) mem_Collect_eq)
  by (metis (no_types, lifting) assms mem_Collect_eq semialg_functions_def)

lemma semialg_functions_in_Qp_funs:
"semialg_functions n  carrier (Fun⇘nQp)"
  using semialg_functions_memE
  by blast

lemma semialg_functions_memI:
  assumes "f  carrier (Fun⇘nQp)"
  assumes "is_semialg_function n f"
  shows "f  semialg_functions n"
  using assms unfolding semialg_functions_def
  by (metis (mono_tags, lifting) Qp_funs_car_memI function_ring_car_eqI
      is_semialg_function_closed mem_Collect_eq restrict_Pi_cancel restrict_apply)

lemma restrict_is_semialg:
  assumes "is_semialg_function n f"
  shows "is_semialg_function n (restrict f (carrier (Qpn)))"
proof(rule is_semialg_functionI)
  show 0: "restrict f (carrier (Qpn))  carrier (Qpn)  carrier Qp"
    using assms is_semialg_function_closed by blast
  show "k S. S  semialg_sets (1 + k)  is_semialgebraic (n + k) (partial_pullback n (restrict f (carrier (Qpn))) k S)"
  proof- fix k S assume A: "S  semialg_sets (1 + k)"
    have "(partial_pullback n (restrict f (carrier (Qpn))) k S) = partial_pullback n f k S"
      apply(intro equalityI'  partial_pullback_memI, meson partial_pullback_memE)
      unfolding partial_pullback_def partial_image_def evimage_eq restrict_def
        apply (metis (mono_tags, lifting) le_add1 local.take_closed)
       apply blast
      by (metis (mono_tags, lifting) le_add1 local.take_closed)
    then show " is_semialgebraic (n + k) (partial_pullback n (restrict f (carrier (Qpn))) k S)"
      using assms A is_semialg_functionE is_semialgebraicI
      by presburger
  qed
qed

lemma restrict_in_semialg_functions:
  assumes "is_semialg_function n f"
  shows "(restrict f (carrier (Qpn)))  semialg_functions n"
  using assms restrict_is_semialg
  unfolding semialg_functions_def
  by (metis (mono_tags, lifting) is_semialg_function_closed mem_Collect_eq restrict_apply' restrict_ext)

lemma constant_function_is_semialg:
  assumes "a  carrier Qp"
  shows "is_semialg_function n (constant_function (carrier (Qpn)) a)"
proof-
  have "(constant_function (carrier (Qpn)) a) = restrict (Qp_ev (Qp.indexed_const a)) (carrier (Qpn))"
    apply(rule ext)
    unfolding constant_function_def
    using eval_at_point_const assms by simp
  then show ?thesis using restrict_in_semialg_functions poly_is_semialg[of  "Qp.indexed_const a"]
    using assms(1)  Qp_to_IP_car restrict_is_semialg by presburger
qed

lemma constant_function_in_semialg_functions:
  assumes "a  carrier Qp"
  shows "Qp_const n a  semialg_functions n"
  apply(unfold semialg_functions_def constant_function_def mem_Collect_eq, intro conjI, auto simp: assms)
  using assms constant_function_is_semialg[of a n] unfolding constant_function_def by auto

lemma function_one_as_constant:
"𝟭⇘Fun⇘nQp= Qp_const n 𝟭"
  unfolding constant_function_def  function_ring_def[of "carrier (Qpn)" Qp] function_one_def
  by simp

lemma function_zero_as_constant:
"𝟬⇘Fun⇘nQp= Qp_const n 𝟬Qp⇙"
  unfolding constant_function_def function_ring_def[of "carrier (Qpn)" Qp] function_zero_def
  by simp

lemma sum_in_semialg_functions:
  assumes "f  semialg_functions n"
  assumes "g  semialg_functions n"
  shows "f ⇘Fun⇘nQpg  semialg_functions n"
proof-
  have 0:"f ⇘Fun⇘nQpg = restrict (function_tuple_comp Qp [f,g] Qp_add_fun) (carrier (Qpn))"
  proof fix x
    show "(f ⇘Fun⇘nQpg) x = restrict (function_tuple_comp Qp [f, g] Qp_add_fun) (carrier (Qpn)) x"
    proof(cases "x  carrier (Qpn)")
      case True
      have " restrict (function_tuple_comp Qp [f, g] Qp_add_fun) (carrier (Qpn)) x = Qp_add_fun [f x, g x]"
        unfolding function_tuple_comp_def function_tuple_eval_def restrict_def
        using comp_apply[of "Qp_add_fun" "(λx. map (λf. f x) [f, g])" x]
        by (metis (no_types, lifting) True list.simps(8) list.simps(9))
      then have "restrict (function_tuple_comp Qp [f, g] Qp_add_fun) (carrier (Qpn)) x = f x Qpg x"
        unfolding Qp_add_fun_def
        by (metis One_nat_def nth_Cons_0 nth_Cons_Suc)
      then show ?thesis  using True function_ring_def[of "carrier (Qpn)" Qp]
        unfolding function_add_def
        by (metis (no_types, lifting) Qp_funs_add assms(1) assms(2) mem_Collect_eq semialg_functions_def)
    next
      case False
      have "(f ⇘Fun⇘nQpg) x = undefined"
        using function_ring_def[of "carrier (Qpn)" Qp] unfolding function_add_def
        by (metis (mono_tags, lifting) False restrict_apply ring_record_simps(12))
      then show ?thesis
        by (metis False restrict_def)
    qed
  qed
  have 1: "is_semialg_function_tuple n [f, g]"
    using assms is_semialg_function_tupleI[of "[f, g]" n] semialg_functions_memE
    by (metis list.distinct(1) list.set_cases set_ConsD)
  have 2: "is_semialg_function n (function_tuple_comp Qp [f,g] Qp_add_fun)"
    apply(rule semialg_function_tuple_comp[of _ _  2])
        apply (simp add: "1")
       apply simp
    by (simp add: addition_is_semialg)
  show ?thesis
    apply(rule semialg_functions_memI)
  apply (meson Qp_funs_is_cring assms(1) assms(2) cring.cring_simprules(1) semialg_functions_memE(2))
    using 0 2 restrict_is_semialg by presburger
qed

lemma prod_in_semialg_functions:
  assumes "f  semialg_functions n"
  assumes "g  semialg_functions n"
  shows "f ⇘Fun⇘nQpg  semialg_functions n"
proof-
  have 0:"f ⇘Fun⇘nQpg = restrict (function_tuple_comp Qp [f,g] Qp_mult_fun) (carrier (Qpn))"
  proof fix x
    show "(f ⇘Fun⇘nQpg) x = restrict (function_tuple_comp Qp [f, g] Qp_mult_fun) (carrier (Qpn)) x"
    proof(cases "x  carrier (Qpn)")
      case True
      have " restrict (function_tuple_comp Qp [f, g] Qp_mult_fun) (carrier (Qpn)) x = Qp_mult_fun [f x, g x]"
        unfolding function_tuple_comp_def function_tuple_eval_def restrict_def
        using comp_apply[of "Qp_mult_fun" "(λx. map (λf. f x) [f, g])" x]
        by (metis (no_types, lifting) True list.simps(8) list.simps(9))
      then have "restrict (function_tuple_comp Qp [f, g] Qp_mult_fun) (carrier (Qpn)) x = f x  g x"
        unfolding Qp_mult_fun_def
        by (metis One_nat_def nth_Cons_0 nth_Cons_Suc)
      then show ?thesis  using True function_ring_def[of "carrier (Qpn)" Qp]
        unfolding function_mult_def
        by (metis (no_types, lifting) Qp_funs_mult  assms(1) assms(2) mem_Collect_eq semialg_functions_def)
    next
      case False
      have "(f ⇘Fun⇘nQpg) x = undefined"
        using function_ring_def[of "carrier (Qpn)" Qp] unfolding function_mult_def
        by (metis (mono_tags, lifting) False restrict_apply ring_record_simps(5))
      then show ?thesis
        by (metis False restrict_def)
    qed
  qed
  have 1: "is_semialg_function_tuple n [f, g]"
    using assms is_semialg_function_tupleI[of "[f, g]" n] semialg_functions_memE
    by (metis list.distinct(1) list.set_cases set_ConsD)
  have 2: "is_semialg_function n (function_tuple_comp Qp [f,g] Qp_mult_fun)"
    apply(rule semialg_function_tuple_comp[of _ _  2])
        apply (simp add: "1")
       apply simp
    by (simp add: multiplication_is_semialg)
  show ?thesis
    apply(rule semialg_functions_memI)
    apply (meson Qp_funs_is_cring assms(1) assms(2) cring.cring_simprules(5) semialg_functions_memE(2))
    using 0 2 restrict_is_semialg by presburger
qed

lemma inv_in_semialg_functions:
  assumes "f  semialg_functions n"
  assumes " x. x  carrier (Qpn)  f x  𝟬Qp⇙"
  shows "inv⇘Fun⇘nQpf  semialg_functions n "
proof-
  have 0: "inv⇘Fun⇘nQpf = restrict (function_tuple_comp Qp [f] Qp_invert) (carrier (Qpn))"
  proof fix x
    show "(inv⇘Fun⇘nQpf) x = restrict (function_tuple_comp Qp [f] Qp_invert) (carrier (Qpn)) x"
    proof(cases "x  carrier (Qpn)")
      case True
      have "(function_tuple_comp Qp [f] Qp_invert) x = Qp_invert [f x]"
        unfolding function_tuple_comp_def function_tuple_eval_def
        using comp_apply  by (metis (no_types, lifting) list.simps(8) list.simps(9))
      then have "(function_tuple_comp Qp [f] Qp_invert) x = invQp(f x)"
        unfolding Qp_invert_def
        using True assms(2) Qp.to_R_to_R1 by presburger
      then show ?thesis
        using True restrict_apply
        by (metis (mono_tags, opaque_lifting) Qp_funs_Units_memI(1)
            Qp_funs_m_inv assms(1) assms(2) semialg_functions_memE(2))
    next
      case False
      have "inv⇘Fun⇘nQpf  carrier (Fun⇘nQp)"
        using assms
        by (meson Qp_funs_Units_memI(1) Qp_funs_is_monoid monoid.Units_inv_closed semialg_functions_memE(2))
      then show ?thesis using False restrict_apply function_ring_not_car
        by auto
    qed
  qed
  have "is_semialg_function n (function_tuple_comp Qp [f] Qp_invert)"
    apply(rule semialg_function_tuple_comp[of _ _  1])
        apply (simp add: assms(1) is_semialg_function_tuple_def semialg_functions_memE(1))
        apply simp
        using Qp_invert_is_semialg by blast
  then show ?thesis
    using "0" restrict_in_semialg_functions
    by presburger
qed

lemma a_inv_in_semialg_functions:
  assumes "f  semialg_functions n"
  shows "⇘Fun⇘nQpf  semialg_functions n"
proof-
  have "⇘Fun⇘nQpf = (Qp_const n ( 𝟭)) ⇘Fun⇘nQpf"
  proof fix x
    show "(⇘Fun⇘nQpf) x = (Qp_const n ( 𝟭) ⇘Fun⇘nQpf) x"
    proof(cases "x  carrier (Qpn)")
      case True
      have 0: "(⇘Fun⇘nQpf) x =  (f x)"
        using Qp_funs_a_inv semialg_functions_memE True assms(1) by blast
      have 1: "(Qp_const n ( 𝟭) ⇘Fun⇘nQpf) x = ( 𝟭)(f x)"
        using Qp_funs_mult[of x n "Qp_const n ( 𝟭)" f] assms Qp_constE[of " 𝟭" x n]
          Qp_funs_mult'  Qp.add.inv_closed Qp.one_closed Qp_funs_mult''' True by presburger
      have 2: "f x  carrier Qp"
        using True semialg_functions_memE[of f n]  assms by blast
      show ?thesis
        using True assms 0 1 2 Qp.l_minus Qp.l_one Qp.one_closed by presburger
    next
      case False
      have "(function_ring (carrier (Qpn)) Qpf) x = undefined"
        using  Qp_funs_a_inv'[of f n] False assms semialg_functions_memE
        by (metis (no_types, lifting) restrict_apply)
      then show ?thesis
        using False function_ring_defs(2)[of n] Qp_funs_a_inv'[of f n]
        unfolding function_mult_def restrict_def
        by presburger
    qed
  qed
  then show ?thesis
    using prod_in_semialg_functions[of "Qp_const n ( 𝟭)" n f] assms
      constant_function_in_semialg_functions[of " 𝟭" n]  Qp.add.inv_closed Qp.one_closed
    by presburger
qed

lemma semialg_functions_subring:
  shows "subring (semialg_functions n) (Fun⇘nQp)"
  apply(rule ring.subringI)
  using Qp_funs_is_cring cring.axioms(1) apply blast
  apply (simp add: semialg_functions_in_Qp_funs)
  using Qp.one_closed  constant_function_in_semialg_functions function_one_as_constant apply presburger
  using a_inv_in_semialg_functions  apply blast
  using  prod_in_semialg_functions apply blast
  using  sum_in_semialg_functions by blast

lemma semialg_functions_subcring:
  shows "subcring (semialg_functions n) (Fun⇘nQp)"
  using semialg_functions_subring cring.subcringI'
  using Qp_funs_is_cring by blast

definition SA where
"SA n = (Fun⇘nQp) carrier := semialg_functions n"

lemma SA_is_ring:
  shows "ring (SA n)"
proof-
  have "ring (Fun⇘nQp)"
    by (simp add: Qp_funs_is_cring cring.axioms(1))
  then show ?thesis
    unfolding SA_def
  using  ring.subring_is_ring[of "Fun⇘nQp" "semialg_functions n"] semialg_functions_subring[of n]
  by blast
qed

lemma SA_is_cring:
  shows "cring (SA n)"
  using ring.subcring_iff[of "Fun⇘nQp" "semialg_functions n"] semialg_functions_subcring[of n]
         Qp_funs_is_cring cring.axioms(1) semialg_functions_in_Qp_funs
    unfolding SA_def
  by blast

lemma SA_is_monoid:
  shows "monoid (SA n)"
  using SA_is_ring[of n]  unfolding ring_def
  by linarith

lemma SA_is_abelian_monoid:
  shows "abelian_monoid (SA n)"
  using SA_is_ring[of n]  unfolding ring_def abelian_group_def by blast

lemma SA_car:
"carrier (SA n) = semialg_functions n"
    unfolding SA_def
  by simp

lemma SA_car_in_Qp_funs_car:
"carrier (SA n)  carrier (Fun⇘nQp)"
  by (simp add: SA_car semialg_functions_in_Qp_funs)

lemma SA_car_memI:
  assumes "f  carrier (Fun⇘nQp)"
  assumes "is_semialg_function n f"
  shows "f  carrier (SA n)"
  using assms semialg_functions_memI[of f n] SA_car
  by blast

lemma SA_car_memE:
  assumes "f  carrier (SA n)"
  shows "is_semialg_function n f"
        "f  carrier (Fun⇘nQp)"
        "f  carrier (Qpn)  carrier Qp"
  using SA_car assms semialg_functions_memE(1) apply blast
  using SA_car assms semialg_functions_memE(2) apply blast
  using SA_car assms semialg_functions_memE(3) by blast

lemma SA_plus:
"(⊕SA n) = (⊕⇘Fun⇘nQp)"
  unfolding SA_def
  by simp

lemma SA_times:
"(⊗SA n) = (⊗⇘Fun⇘nQp)"
  unfolding SA_def
  by simp

lemma SA_one:
"(𝟭SA n) = (𝟭⇘Fun⇘nQp)"
  unfolding SA_def
  by simp

lemma SA_zero:
"(𝟬SA n) = (𝟬⇘Fun⇘nQp)"
  unfolding SA_def
  by simp

lemma SA_zero_is_function_ring:
"(Fun⇘0Qp) = SA 0"
proof-
  have 0: "carrier (Fun⇘0Qp) = carrier (SA 0)"
  proof
    show "carrier (function_ring (carrier (Qp0)) Qp)  carrier (SA 0)"
    proof fix f assume A0: "f  carrier (function_ring (carrier (Qp0)) Qp)"
      show "f  carrier (SA 0)"
      proof(rule SA_car_memI)
        show "f  carrier (function_ring (carrier (Qp0)) Qp)"
          using A0 by blast
        show "is_semialg_function 0 f"
        proof(rule is_semialg_functionI)
          show "f  carrier (Qp0)  carrier Qp"
            using A0 Qp.function_ring_car_memE by blast
          show "k S. S  semialg_sets (1 + k)  is_semialgebraic (0 + k) (partial_pullback 0 f k S)"
          proof- fix k S assume A: "S  semialg_sets (1+k)"
            obtain a where a_def: "a = f []"
              by blast
            have 0: "carrier (Qp0) = {[]}"
              using Qp_zero_carrier by blast
            have 1: "(partial_pullback 0 f k S) = {x  carrier (Qpk). a#x  S}"
            proof
              show "partial_pullback 0 f k S  {x  carrier (Qpk). a # x  S}"
                apply(rule subsetI)
              unfolding partial_pullback_def partial_image_def using a_def
              by (metis (no_types, lifting) add.left_neutral drop0 evimage_eq mem_Collect_eq take_eq_Nil)
              show "{x  carrier (Qpk). a # x  S}  partial_pullback 0 f k S"
                apply(rule subsetI)
                unfolding partial_pullback_def partial_image_def  a_def
                by (metis (no_types, lifting) add.left_neutral drop0 evimageI2 mem_Collect_eq take0)
            qed
            have 2: "cartesian_product {[a]} (partial_pullback 0 f k S) = (cartesian_product {[a]} (carrier (Qpk)))  S"
            proof
              show "cartesian_product {[a]} (partial_pullback 0 f k S)  cartesian_product {[a]} (carrier (Qpk))  S"
              proof(rule subsetI) fix x assume A: "x  cartesian_product {[a]} (partial_pullback 0 f k S)"
                then obtain y where y_def: "y  (partial_pullback 0 f k S) x = a#y"
                  using cartesian_product_memE'
                  by (metis (no_types, lifting) Cons_eq_appendI self_append_conv2 singletonD)
                hence 20: "x  S"
                  unfolding 1 by blast
                have 21: "x = [a]@y"
                  using y_def by (simp add: y_def)
                have "x  cartesian_product {[a]} (carrier (Qpk))"
                  unfolding 21 apply(rule cartesian_product_memI'[of _ Qp 1 _ k])
                  using a_def  apply (metis function_ring_car_closed Qp_zero_carrier f  carrier (function_ring (carrier (Qp0)) Qp) empty_subsetI insert_subset singletonI Qp.to_R1_closed)
                    apply blast
                   apply blast
                  using y_def unfolding partial_pullback_def  evimage_def
                  by (metis IntD2 add_cancel_right_left)
                thus "x  cartesian_product {[a]} (carrier (Qpk))  S"
                  using 20 by blast
              qed
              show " cartesian_product {[a]} (carrier (Qpk))  S  cartesian_product {[a]} (partial_pullback 0 f k S)"
              proof fix x assume A: "x  cartesian_product {[a]} (carrier (Qpk))  S"
                then obtain y where y_def: "x = a#y  y  carrier (Qpk)"
                  using cartesian_product_memE'
                  by (metis (no_types, lifting) Cons_eq_appendI IntD1 append_Nil singletonD)
                have 00: "y  partial_pullback 0 f k S"
                  using y_def unfolding 1  using A by blast
                have 01: "x = [a]@y"
                  using y_def by (simp add: y_def)
                have 02: "partial_pullback 0 f k S  carrier (Qpk)"
                  unfolding partial_pullback_def
                  by (simp add: extensional_vimage_closed)
                show "x  cartesian_product {[a]} (partial_pullback 0 f k S)"
                  unfolding 01 apply(rule cartesian_product_memI'[of "{[a]}" Qp 1 "partial_pullback 0 f k S" k "[a]" y ])
                  apply (metis function_ring_car_closed Qp_zero_carrier f  carrier (function_ring (carrier (Qp0)) Qp) a_def empty_subsetI insert_subset singletonI Qp.to_R1_closed)
                  using "02" apply blast
                   apply blast
                  using 00 by blast
              qed
            qed
            have 3:"is_semialgebraic 1 {[a]}"
            proof-
              have "a  carrier Qp"
                using a_def Qp.function_ring_car_memE 0  f  carrier (function_ring (carrier (Qp0)) Qp) by blast
              hence "[a]  carrier (Qp1)"
                using Qp.to_R1_closed by blast
              thus ?thesis
                using is_algebraic_imp_is_semialg singleton_is_algebraic by blast
            qed
            have  4: "is_semialgebraic (1+k)  (cartesian_product {[a]} (carrier (Qpk)))"
              using 3 carrier_is_semialgebraic cartesian_product_is_semialgebraic less_one by blast
            have  5: "is_semialgebraic (1+k) (cartesian_product {[a]} (partial_pullback 0 f k S))"
              unfolding 2 using 3 4 A intersection_is_semialg padic_fields.is_semialgebraicI padic_fields_axioms by blast
            have 6: "{[a]}  carrier (Qp1)"
              using a_def A0 0 by (metis Qp.function_ring_car_memE empty_subsetI insert_subset singletonI Qp.to_R1_closed)
            have  7: "is_semialgebraic (k+1) (cartesian_product (partial_pullback 0 f k S) {[a]})"
              apply(rule cartesian_product_swap)
                using "6" apply blast
               apply (metis add_cancel_right_left partial_pullback_closed)
              using "5" by auto
            have 8: "is_semialgebraic k (partial_pullback 0 f k S)"
              apply(rule cartesian_product_singleton_factor_projection_is_semialg'[of _ _ "[a]" 1])
                  apply (metis add_cancel_right_left partial_pullback_closed)
                 apply (metis A0 Qp.function_ring_car_memE Qp_zero_carrier a_def singletonI Qp.to_R1_closed)
                using "7" by blast
            thus "is_semialgebraic (0 + k) (partial_pullback 0 f k S)"
                by simp
          qed
        qed
      qed
    qed
    show "carrier (SA 0)  carrier (function_ring (carrier (Qp0)) Qp)"
      using SA_car_in_Qp_funs_car by blast
  qed
  then have 1: "semialg_functions 0 = carrier (Fun⇘0Qp)"
    unfolding 0 SA_def by auto
  show ?thesis unfolding SA_def 1 by auto
qed

lemma constant_fun_closed:
  assumes "c  carrier Qp"
  shows "constant_function (carrier (Qpm)) c  carrier (SA m)"
  using constant_function_in_semialg_functions SA_car assms by blast

lemma SA_0_car_memI:
  assumes "ξ  carrier (Qp0)  carrier Qp"
  assumes "x. x  carrier (Qp0)  ξ x = undefined"
  shows "ξ  carrier (SA 0)"
proof-
  have 0: "carrier (Qp0) = {[]}"
    by (simp add: Qp_zero_carrier)
  obtain c where c_def: "ξ [] = c"
    by blast
  have 1: "ξ = constant_function (carrier (Qp0)) c"
    unfolding constant_function_def restrict_def
    using assms c_def  unfolding 0
    by (metis empty_iff insert_iff)
  have 2: "c  carrier Qp"
    using assms(1) c_def unfolding 0 by blast
  show ?thesis unfolding 1
    using 2 constant_fun_closed by blast
qed

lemma car_SA_0_mem_imp_const:
  assumes "a  carrier (SA 0)"
  shows " c  carrier Qp. a = Qp_const 0 c"
proof-
  obtain c where c_def: "c =  a []"
    by blast
  have car_zero: "carrier (Qp0) = {[]}"
    using Qp_zero_carrier by blast
  have 0: "a = constant_function (carrier (Qp0)) c"
  proof fix x
    show " a x = constant_function (carrier (Qp0)) c x"
      apply(cases "x  carrier (Qp0)")
      using assms SA_car_memE[of a 0] c_def
      unfolding constant_function_def restrict_def  car_zero
       apply (metis empty_iff insert_iff)
      using assms SA_car_memE(2)[of a 0] c_def
      unfolding constant_function_def restrict_def car_zero
      by (metis car_zero function_ring_not_car)
  qed
  have c_closed: "c  carrier Qp"
    using assms  SA_car_memE(3)[of  a 0] unfolding c_def car_zero
    by blast
  thus ?thesis using 0 by blast
qed

lemma SA_zeroE:
  assumes "a  carrier (Qpn)"
  shows "𝟬SA na = 𝟬"
  using function_zero_eval SA_zero assms by presburger

lemma SA_oneE:
  assumes "a  carrier (Qpn)"
  shows "𝟭SA na = 𝟭"
  using function_one_eval  SA_one assms by presburger
end

sublocale padic_fields < UPSA?: UP_cring "SA m" "UP (SA m)"
  unfolding UP_cring_def using SA_is_cring[of m] by auto

context padic_fields
begin

lemma SA_add:
  assumes "x  carrier (Qpn)"
  shows "(f SA ng) x = f x Qpg x"
  using Qp_funs_add''' SA_plus assms by presburger

lemma SA_add':
  assumes "x  carrier (Qpn)"
  shows "(f SA ng) x = undefined"
proof-
  have "(f SA ng) x = function_add (carrier (Qpn)) Qp f g x"
    using SA_plus[of n] unfolding function_ring_def
    by (metis ring_record_simps(12))
  then show ?thesis
    unfolding function_add_def using restrict_apply assms
    by (metis (no_types, lifting))
qed

lemma SA_mult:
  assumes "x  carrier (Qpn)"
  shows "(f SA ng) x = f x  g x"
  using Qp_funs_mult''' SA_times assms by presburger

lemma SA_mult':
  assumes "x  carrier (Qpn)"
  shows "(f SA ng) x = undefined"
proof-
  have "(f SA ng) x = function_mult (carrier (Qpn)) Qp f g x"
    using SA_times[of n] unfolding function_ring_def
    by (metis ring_record_simps(5))
  then show ?thesis
    unfolding function_mult_def using restrict_apply assms
    by (metis (no_types, lifting))
qed

lemma SA_u_minus_eval:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  shows "(SA nf) x =  (f x)"
proof-
  have  "f SA n(SA nf) = 𝟬SA n⇙"
    using assms SA_is_cring cring.cring_simprules(17) by metis
  have  "(f SA n(SA nf)) x = 𝟬SA nx"
    using assms f SA nSA nf = 𝟬SA n⇙› by presburger
  then have "(f x)  (SA nf) x = 𝟬"
    using assms function_zero_eval SA_add unfolding   SA_zero by blast
  then show ?thesis
    using assms SA_is_ring
    by (meson Qp.add.inv_closed Qp.add.inv_comm Qp.function_ring_car_memE Qp.minus_unique Qp.r_neg SA_car_memE(2) ring.ring_simprules(3))
qed

lemma SA_a_inv_eval:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  shows "(SA nf) x =  (f x)"
proof-
  have  "f SA n(SA nf) = 𝟬SA n⇙"
    using assms SA_is_cring cring.cring_simprules(17) by metis
  have  "(f SA n(SA nf)) x = 𝟬SA nx"
    using assms f SA nSA nf = 𝟬SA n⇙› by presburger
  then have "(f x)  (SA nf) x = 𝟬"
    by (metis function_zero_eval SA_add SA_zero assms)
  then show ?thesis
    by (metis (no_types, lifting) PiE Qp_def Qp.add.m_comm Qp.minus_equality SA_is_cring Zp_def assms(1) assms(2) cring.cring_simprules(3) padic_fields.SA_car_memE(3) padic_fields_axioms)
qed

lemma SA_nat_pow:
  assumes "x  carrier (Qpn)"
  shows "(f [^]SA n(k::nat)) x = (f x) [^]Qpk"
  apply(induction k)
  using assms  nat_pow_def
  apply (metis function_one_eval SA_one old.nat.simps(6))
  using assms SA_mult
  by (metis Group.nat_pow_Suc)

lemma SA_nat_pow':
  assumes "x  carrier (Qpn)"
  shows "(f [^]SA n(k::nat)) x = undefined"
  apply(induction k)
  using assms nat_pow_def[of "SA n" f]
  apply (metis (no_types, lifting) Group.nat_pow_0 Qp_funs_one SA_one restrict_apply)
  by (metis Group.nat_pow_Suc SA_mult' assms)

lemma SA_add_closed_id:
  assumes "is_semialg_function n f"
  assumes "is_semialg_function n g"
  shows "restrict f (carrier (Qpn)) SA nrestrict g (carrier (Qpn)) = f SA ng "
proof fix x
  show "(restrict f (carrier (Qpn)) SA nrestrict g (carrier (Qpn))) x = (f SA ng) x"
    apply(cases "x  carrier (Qpn)")
    using assms restrict_apply
    apply (metis SA_add)
    using assms
    by (metis SA_add')
qed

lemma SA_mult_closed_id:
  assumes "is_semialg_function n f"
  assumes "is_semialg_function n g"
  shows "restrict f (carrier (Qpn)) SA nrestrict g (carrier (Qpn)) = f SA ng "
proof fix x
  show "(restrict f (carrier (Qpn)) SA nrestrict g (carrier (Qpn))) x = (f SA ng) x"
    apply(cases "x  carrier (Qpn)")
    using assms restrict_apply
    apply (metis SA_mult)
    using assms
    by (metis SA_mult')
qed

lemma SA_add_closed:
  assumes "is_semialg_function n f"
  assumes "is_semialg_function n g"
  shows "f SA ng  carrier (SA n)"
    using assms SA_add_closed_id
    by (metis SA_car SA_plus restrict_in_semialg_functions sum_in_semialg_functions)

lemma SA_mult_closed:
  assumes "is_semialg_function n f"
  assumes "is_semialg_function n g"
  shows "f SA ng  carrier (SA n)"
    using assms SA_mult_closed_id
  by (metis SA_car SA_is_cring cring.cring_simprules(5) restrict_in_semialg_functions)

lemma SA_add_closed_right:
  assumes "is_semialg_function n f"
  assumes "g  carrier (SA n)"
  shows "f SA ng  carrier (SA n)"
  using SA_add_closed SA_car_memE(1) assms(1) assms(2)   by blast

lemma SA_mult_closed_right:
  assumes "is_semialg_function n f"
  assumes "g  carrier (SA n)"
  shows "f SA ng  carrier (SA n)"
  using SA_car_memE(1) SA_mult_closed assms(1) assms(2)   by blast

lemma SA_add_closed_left:
  assumes "f  carrier (SA n)"
  assumes "is_semialg_function n g"
  shows "f SA ng  carrier (SA n)"
  using SA_add_closed SA_car_memE(1) assms(1) assms(2)   by blast

lemma SA_mult_closed_left:
  assumes "f  carrier (SA n)"
  assumes "is_semialg_function n g"
  shows "f SA ng  carrier (SA n)"
  using SA_car_memE(1) SA_mult_closed assms(1) assms(2)  by blast

lemma SA_nat_pow_closed:
  assumes "is_semialg_function n f"
  shows "f [^]SA n(k::nat)  carrier (SA n)"
  apply(induction k)
  using nat_pow_def[of "SA n" f ]
  apply (metis Group.nat_pow_0  monoid.one_closed SA_is_monoid)
  by (metis Group.nat_pow_Suc SA_car assms(1) assms SA_mult_closed semialg_functions_memE(1))

lemma SA_imp_semialg:
  assumes "f  carrier (SA n)"
  shows "is_semialg_function n f"
  using SA_car assms semialg_functions_memE(1) by blast

lemma SA_minus_closed:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  shows "(f  SA ng)  carrier (SA n)"
  using assms unfolding a_minus_def
  by (meson SA_add_closed_left SA_imp_semialg SA_is_ring ring.ring_simprules(3))

lemma(in ring) add_pow_closed :
  assumes "b  carrier R"
  shows "([(m::nat)]Rb)  carrier R"
  by(rule add.nat_pow_closed, rule assms)

lemma(in ring) add_pow_Suc:
  assumes "b  carrier R"
  shows "[(Suc m)]b = [m]b  b"
  using assms add.nat_pow_Suc by blast

lemma(in ring) add_pow_zero:
  assumes "b  carrier R"
  shows "[(0::nat)]b = 𝟬"
  using assms nat_mult_zero
  by blast

lemma Fun_add_pow_apply:
  assumes "b  carrier (Fun⇘nQp)"
  assumes "a  carrier (Qpn)"
  shows "([(m::nat)]⇘Fun⇘nQpb) a = [m]( b a)"
proof-
  have 0: "b a  carrier Qp"
    using Qp.function_ring_car_mem_closed assms by fastforce
  have 1: "ring (Fun⇘nQp)"
    using function_ring_is_ring by blast
  show ?thesis
  proof(induction m)
    case 0
    have "([(0::nat)] function_ring (carrier (Qpn)) Qpb) = 𝟬⇘Fun⇘nQp⇙"
      using 1 ring.add_pow_zero[of "Fun⇘nQp" b ] assms  by blast
    then show ?case
      using  function_zero_eval Qp.nat_mult_zero assms  by presburger
  next
    case (Suc m)
    then show ?case using Suc ring.add_pow_Suc[of "SA n" b m] assms
      by (metis (no_types, lifting) "0" "1" Qp.ring_axioms SA_add SA_plus ring.add_pow_Suc)
  qed
qed

lemma SA_add_pow_apply:
  assumes "b  carrier (SA n)"
  assumes "a  carrier (Qpn)"
  shows "([(m::nat)]SA nb) a = [m]( b a)"
  apply(induction m)
  using assms SA_is_ring[of n] Fun_add_pow_apply
  apply (metis function_zero_eval Qp.nat_mult_zero SA_zero ring.add_pow_zero)
  using assms SA_is_ring[of n] ring.add_pow_Suc[of "Fun⇘nQp" b]  ring.add_pow_Suc[of "SA n" b] SA_plus[of n]
  using Fun_add_pow_apply
  by (metis Qp.add.nat_pow_Suc SA_add)

lemma Qp_funs_Units_SA_Units:
  assumes "f  Units (Fun⇘nQp)"
  assumes "is_semialg_function n f"
  shows "f  Units (SA n)"
proof-
  have 0: "f  carrier (Fun⇘nQp)"
    by (meson Qp_funs_is_monoid assms(1) monoid.Units_closed)
  have 1: "inv⇘Fun⇘nQpf  semialg_functions n"
    using monoid.Units_closed[of "Fun⇘nQp" f]
      assms inv_in_semialg_functions[of f n] Qp_funs_Units_memE(3)[of f n]
      semialg_functions_memI[of f n] Qp_funs_is_monoid by blast
  then have 2: "f SA n(inv⇘Fun⇘nQpf ) = 𝟭SA n⇙"
    using Qp_funs_Units_memE(1)[of f n] SA_one SA_times assms(1)
    by presburger
  then have 3: "(inv⇘Fun⇘nQpf ) SA nf  = 𝟭SA n⇙"
    using Qp_funs_Units_memE(2)[of f n] SA_one SA_times assms(1)
    by presburger
  have 4: "f  carrier (SA n)"
    using "0" SA_car assms(2) semialg_functions_memI by blast
  have 5: "inv⇘Fun⇘nQpf  carrier (SA n)"
    using SA_car_memI "1" SA_car by blast
  show ?thesis
    using 5 4 3 2 unfolding Units_def
    by blast
qed

lemma SA_Units_memE:
  assumes "f  (Units (SA n))"
  shows "f SA ninvSA nf = 𝟭SA n⇙"
        "invSA nf SA nf = 𝟭SA n⇙"
  using assms SA_is_monoid[of n] monoid.Units_r_inv[of "SA n" f]
  apply blast
  using assms SA_is_monoid[of n] monoid.Units_l_inv[of "SA n" f]
  by blast

lemma SA_Units_closed:
  assumes "f  (Units (SA n))"
  shows "f  carrier (SA n)"
  using assms unfolding Units_def by blast

lemma SA_Units_inv_closed:
  assumes "f  (Units (SA n))"
  shows "invSA nf  carrier (SA n)"
  using assms SA_is_monoid[of n] monoid.Units_inv_closed[of "SA n" f]
  by blast

lemma SA_Units_Qp_funs_Units:
  assumes "f  (Units (SA n))"
  shows "f  (Units (Fun⇘nQp))"
proof-
  have 0:  "f SA ninvSA nf = 𝟭SA n⇙"
        "invSA nf SA nf = 𝟭SA n⇙"
     using R.Units_r_inv assms apply blast
     using R.Units_l_inv assms by blast
  have 1: "f  carrier (Fun⇘nQp)"
    using assms
    by (metis SA_car SA_is_monoid monoid.Units_closed semialg_functions_memE(2))
  have 2: "invSA nf  carrier (Fun⇘nQp)"
    using SA_Units_inv_closed SA_car assms semialg_functions_memE(2) by blast
  then show ?thesis
   using 0 1 2  SA_one SA_times local.F.UnitsI(1) by auto
qed

lemma SA_Units_Qp_funs_inv:
  assumes "f  (Units (SA n))"
  shows "invSA nf = inv⇘Fun⇘nQpf"
  using assms SA_Units_Qp_funs_Units
  by (metis (no_types, opaque_lifting) Qp_funs_is_cring Qp_funs_is_monoid SA_Units_memE(1)
      SA_is_monoid SA_one SA_times cring.invI(1) monoid.Units_closed monoid.Units_inv_Units)

lemma SA_Units_memI:
  assumes "f  (carrier (SA n))"
  assumes " x. x  carrier (Qpn)   f x  𝟬Qp⇙"
  shows "f  (Units (SA n))"
  using assms Qp_funs_Units_memI[of f n] Qp_funs_Units_SA_Units SA_car SA_imp_semialg
      semialg_functions_memE(2) by blast

lemma SA_Units_memE':
  assumes "f  (Units (SA n))"
  shows  " x. x  carrier (Qpn)   f x  𝟬Qp⇙"
  using assms Qp_funs_Units_memE[of f n] SA_Units_Qp_funs_Units
  by blast

lemma Qp_n_nonempty:
  shows "carrier (Qpn)  {}"
  apply(induction n)
  apply (simp add: Qp_zero_carrier)
  using cartesian_power_cons[of _ Qp _ 𝟭] Qp.one_closed
  by (metis Suc_eq_plus1 all_not_in_conv cartesian_power_cons empty_iff)

lemma SA_one_not_zero:
  shows "𝟭SA n 𝟬SA n⇙"
proof-
  obtain a where a_def: "a  carrier (Qpn)"
    using Qp_n_nonempty by blast
  have "𝟭SA na   𝟬SA na"
    using function_one_eval function_zero_eval SA_one SA_zero a_def local.one_neq_zero by presburger
  then show ?thesis
    by metis
qed

lemma SA_units_not_zero:
  assumes "f  Units (SA n)"
  shows "f  𝟬SA n⇙"
  using SA_one_not_zero
  by (metis assms padic_fields.SA_is_ring padic_fields_axioms ring.ring_in_Units_imp_not_zero)

lemma SA_Units_nonzero:
  assumes "f  Units (SA m)"
  assumes "x  carrier (Qpm)"
  shows "f x  nonzero Qp"
  unfolding nonzero_def  mem_Collect_eq
  apply(rule conjI)
  using assms  SA_Units_closed SA_car_memE(3)[of f m]  apply blast
  using assms SA_Units_memE' by blast

lemma SA_car_closed:
  assumes "f  carrier (SA m)"
  assumes "x  carrier (Qpm)"
  shows "f x  carrier Qp"
  using assms SA_car_memE(3) by blast

lemma SA_Units_closed_fun:
  assumes "f  Units (SA m)"
  assumes "x  carrier (Qpm)"
  shows "f x  carrier Qp"
  using SA_Units_closed SA_car_closed assms by blast

lemma SA_inv_eval:
  assumes "f  Units (SA n)"
  assumes "x  carrier (Qpn)"
  shows "(invSA nf) x = inv (f x)"
proof-
  have  "f SA n(invSA nf) = 𝟭SA n⇙"
    using assms SA_is_cring  SA_Units_memE(1) by blast
  hence  "(f SA n(invSA nf)) x = 𝟭SA nx"
    using assms by presburger
  then have "(f x)  (invSA nf) x = 𝟭"
    by (metis function_one_eval SA_mult SA_one assms)
  then show ?thesis
    by (metis Qp_def Qp_funs_m_inv Zp_def assms(1) assms(2) padic_fields.SA_Units_Qp_funs_Units padic_fields.SA_Units_Qp_funs_inv padic_fields_axioms)
qed

lemma SA_div_eval:
  assumes "f  Units (SA n)"
  assumes "h  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  shows "(h SA n(invSA nf)) x = h x  inv (f x)"
  using assms SA_inv_eval  SA_mult by presburger

lemma SA_unit_int_pow:
  assumes "f  Units (SA m)"
  assumes "x  carrier (Qpm)"
  shows "(f[^]SA m(i::int)) x = (f x)[^]i"
proof(induction i)
  case (nonneg n)
  have  0: "(f [^]SA mint n) = (f [^]SA mn)"
    using assms by (meson int_pow_int)
  then show ?case using SA_Units_closed[of f m] assms
    by (metis SA_nat_pow int_pow_int)
next
  case (neg n)
  have 0: "(f [^]SA m- int (Suc n)) = invSA m(f [^]SA m(Suc n))"
    using assms by (metis R.int_pow_inv' int_pow_int)
  then show ?case unfolding 0 using assms
    by (metis Qp.int_pow_inv' R.Units_pow_closed SA_Units_nonzero SA_inv_eval SA_nat_pow Units_eq_nonzero int_pow_int)
qed

lemma restrict_in_SA_car:
  assumes "is_semialg_function n f"
  shows "restrict f (carrier (Qpn))  carrier (SA n)"
  using assms  SA_car restrict_in_semialg_functions
  by blast

lemma SA_smult:
"(⊙SA n) = (⊙⇘Fun⇘nQp)"
  unfolding SA_def by auto

lemma SA_smult_formula:
  assumes "h  carrier (SA n)"
  assumes "q  carrier Qp"
  assumes "a  carrier (Qpn)"
  shows "(q SA nh) a = q (h a)"
  using SA_smult assms function_smult_eval SA_car_memE(2) by presburger

lemma SA_smult_closed:
  assumes "h  carrier (SA n)"
  assumes "q  carrier Qp"
  shows "q SA nh  carrier (SA n)"
proof-
  obtain g where g_def: "g = 𝔠⇘nq"
    by blast
  have g_closed: "g  carrier (SA n)"
    using g_def assms constant_function_is_semialg[of q n] constant_function_closed SA_car_memI
    by blast
  have "q SA nh = g SA nh"
    apply(rule function_ring_car_eqI[of _ n])
    using function_smult_closed SA_car_memE(2) SA_smult assms apply presburger
    using SA_car_memE(2) assms(1) assms(2) g_closed padic_fields.SA_imp_semialg padic_fields.SA_mult_closed_right padic_fields_axioms apply blast
    using Qp_constE SA_mult SA_smult_formula assms   g_def by presburger
  thus ?thesis
    using SA_imp_semialg SA_mult_closed_right assms(1) assms(2) g_closed by presburger
qed

lemma p_mult_function_val:
  assumes "f  carrier (SA m)"
  assumes "x  carrier (Qpm)"
  shows "val ((𝔭 SA mf) x) = val (f x) + 1"
proof-
  have 0: "(𝔭SA mf) x = 𝔭(f x)"
    using Qp.int_inc_closed SA_smult_formula assms(1) assms(2) by blast
  show ?thesis  unfolding 0 using assms
    by (metis Qp.function_ring_car_memE Qp.int_inc_closed Qp.m_comm SA_car semialg_functions_memE(2) val_mult val_p)
qed

lemma Qp_char_0'':
  assumes "a  carrier Qp"
  assumes "a  𝟬"
  assumes "(k::nat) > 0"
  shows "[k]a  𝟬"
proof-
  have 0: "[k]𝟭 𝟬"
    using Qp_char_0 assms(3) by blast
  have "[k]a = [k]𝟭  a"
    using Qp.add_pow_ldistr Qp.l_one Qp.one_closed assms(1) by presburger
  thus ?thesis using 0 assms
    using Qp.integral by blast
qed

lemma SA_char_zero:
  assumes "f  carrier (SA m)"
  assumes "f  𝟬SA m⇙"
  assumes "n > 0"
  shows "[(n::nat)]SA mf  𝟬SA m⇙"
proof assume A: "[n] SA mf = 𝟬SA m⇙"
  obtain x where x_def: "x  carrier (Qpm)  f x  𝟬"
    using assms
    by (metis function_ring_car_eqI R.cring_simprules(2) SA_car_memE(2) SA_zeroE)
  have 0: "([(n::nat)]SA mf) x = [n](f x)"
    using SA_add_pow_apply assms(1) x_def by blast
  have 1: "[n](f x) = 𝟬"
    using 0 unfolding A  using SA_zeroE x_def by blast
  have 2: "f x  nonzero Qp"
    using x_def assms
    by (metis Qp.function_ring_car_memE SA_car not_nonzero_Qp semialg_functions_memE(2))
  then show False using x_def
    using "1" Qp.nonzero_memE(1) Qp_char_0'' assms(3) by blast
qed



(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Defining Semialgebraic Maps›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  We can define a semialgebraic map in essentially the same way that Denef defines
  semialgebraic functions. As for functions, we can define the partial pullback of a set
  $S \subseteq \mathbb{Q}_p^{n+l}$ by a map $g: \mathbb{Q}_p^m \to \mathbb{Q}_p^n$ to be the set
  \[
  \{(x,y) \in \mathbb{Q}_p^m \times \mathbb{Q}_p^l \mid (f(x), y) \in S \}
  \]
  and say that $g$ is a semialgebraic map if for every $l$, and every semialgebraic
  $S \subseteq \mathbb{Q}_p^{n+l}$, the partial pullback of $S$ by $g$ is also semialgebraic.
  On this definition, it is immediate that the composition $f \circ g$ of a semialgebraic
  function $f: \mathbb{Q}_p^n \to \mathbb{Q}$ and a semialgebraic map
  $g: \mathbb{Q}_p^m \to \mathbb{Q}_p^n$ is semialgebraic. It is also not hard to show that a map
  is semialgebraic if and only if all of its coordinate functions are semialgebraic functions.
  This allows us to build new semialgebraic functions out of old ones via composition.
›


text‹Generalizing the notion of partial image partial pullbacks from functions to maps:›

definition map_partial_image where
"map_partial_image m f xs = (f (take m xs))@(drop m xs)"

definition map_partial_pullback where
"map_partial_pullback m f l S = (map_partial_image m f)  ¯m+lS"

lemma map_partial_pullback_memE:
  assumes "as  map_partial_pullback m f l S"
  shows "as  carrier (Qpm+l)" "map_partial_image m f as  S"
  using assms unfolding map_partial_pullback_def  evimage_def
  apply (metis (no_types, opaque_lifting) Int_iff add.commute)
  using assms unfolding map_partial_pullback_def
  by blast

lemma map_partial_pullback_closed:
"map_partial_pullback m f l S  carrier (Qpm+l)"
  using map_partial_pullback_memE(1) by blast

lemma map_partial_pullback_memI:
  assumes "as  carrier (Qpm+k)"
  assumes "(f (take m as))@(drop m as)  S"
  shows "as  map_partial_pullback m f k S"
  using assms unfolding map_partial_pullback_def map_partial_image_def
  by blast

lemma map_partial_image_eq:
  assumes "as  carrier (Qpn)"
  assumes "bs  carrier (Qpk)"
  assumes "x = as @ bs"
  shows "map_partial_image n f x = (f as)@bs"
proof-
  have 0: "(take n x) = as"
    by (metis append_eq_conv_conj assms(1) assms(3) cartesian_power_car_memE)
  have 1: "drop n x = bs"
    by (metis "0" append_take_drop_id assms(3) same_append_eq)
  show ?thesis using 0 1 unfolding map_partial_image_def
    by blast
qed

lemma map_partial_pullback_memE':
  assumes "as  carrier (Qpn)"
  assumes "bs  carrier (Qpk)"
  assumes "x = as @ bs"
  assumes "x  map_partial_pullback n f k S"
  shows "(f as)@bs  S"
  using map_partial_pullback_memE[of x n f k S] map_partial_image_def[of n f x]
  by (metis assms(1) assms(2) assms(3) assms(4) map_partial_image_eq)

text‹Partial pullbacks have the same algebraic properties as pullbacks.›

lemma map_partial_pullback_intersect:
"map_partial_pullback m f l (S1  S2) = (map_partial_pullback m f l S1)  (map_partial_pullback m f l S2)"
  unfolding map_partial_pullback_def
  by simp

lemma map_partial_pullback_union:
"map_partial_pullback m f l (S1  S2) = (map_partial_pullback m f l S1)  (map_partial_pullback m f l S2)"
  unfolding map_partial_pullback_def
  by simp

lemma map_partial_pullback_complement:
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  shows "map_partial_pullback m f l (carrier (Qpn+l) - S) = carrier (Qpm+l) - (map_partial_pullback m f l S) "
  apply(rule equalityI)
  using map_partial_pullback_def[of m f l "(carrier (Qpn+l) - S)"]
        map_partial_pullback_def[of m f l S]
   apply (metis (no_types, lifting) DiffD2 DiffI evimage_Diff map_partial_pullback_memE(1) subsetI)
proof fix x assume A: " x  carrier (Qpm+l) - map_partial_pullback m f l S"
  show " x  map_partial_pullback m f l (carrier (Qpn+l) - S) "
    apply(rule map_partial_pullback_memI)
    using A
     apply blast
  proof
    have 0: "drop m x  carrier (Qpl)"
      by (meson A DiffD1 cartesian_power_drop)
    have 1: "take m x  carrier (Qpm)"
      using A
      by (meson DiffD1 take_closed le_add1)
    show "f (take m x) @ drop m x  carrier (Qpn+l) "
      using 0 1 assms
      by (meson Pi_iff cartesian_power_concat(1))
    show "f (take m x) @ drop m x  S"
      using A unfolding map_partial_pullback_def map_partial_image_def
      by blast
  qed
qed

lemma map_partial_pullback_carrier:
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  shows "map_partial_pullback m f l (carrier (Qpn+l)) = carrier (Qpm+l)"
  apply(rule equalityI)
  using map_partial_pullback_memE(1) apply blast
proof fix x assume A: "x  carrier (Qpm+l)"
  show "x  map_partial_pullback m f l (carrier (Qpn+l))"
    apply(rule map_partial_pullback_memI)
  using A cartesian_power_drop[of x m l] take_closed assms
   apply blast
proof-
  have "f (take m x)  carrier (Qpn)"
    using A take_closed assms
    by (meson Pi_mem le_add1)
  then show "f (take m x) @ drop m x  carrier (Qpn+l)"
    using cartesian_power_drop[of x m l] A cartesian_power_concat(1)[of _ Qp n _ l]
    by blast
qed
qed

definition is_semialg_map where
"is_semialg_map m n f = (f  carrier (Qpm)  carrier (Qpn)  
                          (l  0. S  semialg_sets (n + l). is_semialgebraic (m + l) (map_partial_pullback m f l S)))"

lemma is_semialg_map_closed:
  assumes "is_semialg_map m n f"
  shows "f  carrier (Qpm)  carrier (Qpn)"
  using is_semialg_map_def assms by blast

lemma is_semialg_map_closed':
  assumes "is_semialg_map m n f" "x  carrier (Qpm)"
  shows "f x   carrier (Qpn)"
  using is_semialg_map_def assms by blast

lemma is_semialg_mapE:
  assumes "is_semialg_map m n f"
  assumes "is_semialgebraic (n + k) S"
  shows " is_semialgebraic (m + k) (map_partial_pullback m f k S)"
  using is_semialg_map_def assms
  by (meson is_semialgebraicE le0)

lemma is_semialg_mapE':
  assumes "is_semialg_map m n f"
  assumes "is_semialgebraic (n + k) S"
  shows " is_semialgebraic (m + k) (map_partial_image m f ¯m+kS)"
  using assms is_semialg_mapE unfolding map_partial_pullback_def
  by blast

lemma is_semialg_mapI:
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  assumes "k  S. S  semialg_sets (n + k)  is_semialgebraic (m + k) (map_partial_pullback m f k S)"
  shows "is_semialg_map m n f"
  using assms unfolding is_semialg_map_def
  by blast

lemma is_semialg_mapI':
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  assumes "k  S. S  semialg_sets (n + k)  is_semialgebraic (m + k) (map_partial_image m f ¯m+kS)"
  shows "is_semialg_map m n f"
  using assms is_semialg_mapI unfolding map_partial_pullback_def
  by blast

text‹Semialgebraicity for functions can be verified on basic semialgebraic sets.›

lemma is_semialg_mapI'':
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  assumes "k  S. S  basic_semialgs (n + k)  is_semialgebraic (m + k) (map_partial_pullback m f k S)"
  shows "is_semialg_map m n f"
  apply(rule is_semialg_mapI)
  using assms(1) apply blast
proof-
  show "k S. S  semialg_sets (n + k)  is_semialgebraic (m + k) (map_partial_pullback m f k S)"
  proof- fix k S assume A: "S  semialg_sets (n + k)"
    show "is_semialgebraic (m + k) (map_partial_pullback m f k S)"
      apply(rule gen_boolean_algebra.induct[of S "carrier (Qpn+k)" "basic_semialgs (n + k)"])
        using A unfolding semialg_sets_def
            apply blast
        using map_partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger
        apply (simp add: assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg map_partial_pullback_carrier map_partial_pullback_intersect)
        using map_partial_pullback_union union_is_semialgebraic apply presburger
        using assms(1) complement_is_semialg map_partial_pullback_complement plus_1_eq_Suc by presburger
  qed
qed

lemma is_semialg_mapI''':
  assumes "f  carrier (Qpm)  carrier (Qpn)"
  assumes "k  S. S  basic_semialgs (n + k)  is_semialgebraic (m + k) (map_partial_image m f ¯m+kS)"
  shows "is_semialg_map m n f"
  using is_semialg_mapI'' assms unfolding map_partial_pullback_def
  by blast

lemma id_is_semialg_map:
"is_semialg_map n n (λ x. x)"
proof-
  have 0: "k S. S  semialg_sets (n + k)  (λxs. take n xs @ drop n xs) -` S  carrier (Qpn + k) =
            S"
    apply(rule equalityI')
     apply (metis (no_types, lifting) Int_iff append_take_drop_id vimage_eq)
    by (metis (no_types, lifting) IntI append_take_drop_id in_mono is_semialgebraicI is_semialgebraic_closed vimageI)
  show ?thesis
    by(intro is_semialg_mapI,
          unfold  map_partial_pullback_def map_partial_image_def evimage_def is_semialgebraic_def  0,
          auto)
qed

lemma map_partial_pullback_comp:
  assumes "is_semialg_map m n f"
  assumes "is_semialg_map k m g"
  shows  "(map_partial_pullback k (f  g) l S) = (map_partial_pullback k g l (map_partial_pullback m f l S))"
proof
  show "map_partial_pullback k (f  g) l S  map_partial_pullback k g l (map_partial_pullback m f l S)"
  proof fix x assume A: " x  map_partial_pullback k (f  g) l S"
    show " x  map_partial_pullback k g l (map_partial_pullback m f l S)"
    proof(rule map_partial_pullback_memI)
      show 0: "x  carrier (Qpk+l)"
        using A map_partial_pullback_memE(1) by blast
      show "g (take k x) @ drop k x  map_partial_pullback m f l S"
      proof(rule map_partial_pullback_memI)
        show "g (take k x) @ drop k x  carrier (Qpm+l)"
        proof-
          have 1: "g (take k x)  carrier (Qpm)"
            using 0 assms(2) is_semialg_map_closed[of k m g]
            by (meson Pi_iff le_add1 take_closed)
          then show ?thesis
            by (metis "0" add.commute cartesian_power_concat(2) cartesian_power_drop)
        qed
        show "f (take m (g (take k x) @ drop k x)) @ drop m (g (take k x) @ drop k x)  S"
        using map_partial_pullback_memE[of x k "f  g" l S]
          comp_apply[of f g] map_partial_image_eq[of "take k x" k "drop k x" l x "f  g"]
        by (metis (no_types, lifting) A g (take k x) @ drop k x  carrier (Qpm + l)
            append_eq_append_conv append_take_drop_id cartesian_power_car_memE
            cartesian_power_drop map_partial_image_def)
      qed
    qed
  qed
  show "map_partial_pullback k g l (map_partial_pullback m f l S)  map_partial_pullback k (f  g) l S"
  proof fix x assume A: "x  map_partial_pullback k g l (map_partial_pullback m f l S)"
    have 0: "(take m (map_partial_image k g x)) = g (take k x)"
    proof-
      have "take k x  carrier (Qpk)"
        using map_partial_pullback_memE[of x k g l] A le_add1 take_closed
        by blast
      then have "length (g (take k x)) = m"
        using assms is_semialg_map_closed[of k m g] cartesian_power_car_memE
        by blast
      then show ?thesis
        using assms unfolding map_partial_image_def
        by (metis append_eq_conv_conj)
    qed
    show " x  map_partial_pullback k (f  g) l S"
      apply(rule map_partial_pullback_memI)
      using A map_partial_pullback_memE
       apply blast
      using 0 assms A comp_apply map_partial_pullback_memE[of x k g l "map_partial_pullback m f l S"]
            map_partial_pullback_memE[of "map_partial_image k g x" m f l S]
  map_partial_image_eq[of "take k x" k "drop k x" l x g]
   map_partial_image_eq[of "take m (map_partial_image k g x)" m "drop m (map_partial_image k g x)" l "(map_partial_image k g x)" f ]
  by (metis (no_types, lifting) cartesian_power_drop le_add1 map_partial_image_def map_partial_pullback_memE' take_closed)
qed
qed

lemma semialg_map_comp_closed:
  assumes "is_semialg_map m n f"
  assumes "is_semialg_map k m g"
  shows "is_semialg_map k n (f  g)"
  apply(intro is_semialg_mapI , unfold Pi_iff comp_def, intro ballI,
        intro is_semialg_map_closed'[of m n f] is_semialg_map_closed'[of k m g]  assms, blast)
proof- fix l S assume A: "S  semialg_sets (n + l)"
  have " is_semialgebraic (k + l) (map_partial_pullback k (f  g) l S)"
    using map_partial_pullback_comp is_semialg_mapE A assms(1) assms(2) is_semialgebraicI
    by presburger
  thus "is_semialgebraic (k + l) (map_partial_pullback k (λx. f (g x)) l S)"
    unfolding comp_def by auto
qed

lemma partial_pullback_comp:
  assumes "is_semialg_function m f"
  assumes "is_semialg_map k m g"
  shows  "(partial_pullback k (f  g) l S) = (map_partial_pullback k g l (partial_pullback m f l S))"
proof
  show "partial_pullback k (f  g) l S  map_partial_pullback k g l (partial_pullback m f l S)"
  proof fix x assume A: "x  partial_pullback k (f  g) l S"
    show "x  map_partial_pullback k g l (partial_pullback m f l S)"
    proof(rule map_partial_pullback_memI)
      show 0: "x  carrier (Qpk+l)"
        using A partial_pullback_memE(1) by blast
      show "g (take k x) @ drop k x  partial_pullback m f l S"
      proof(rule partial_pullback_memI)
        show "g (take k x) @ drop k x  carrier (Qpm+l)"
        proof-
          have 1: "g (take k x)  carrier (Qpm)"
            using 0 assms(2) is_semialg_map_closed[of k m g]
            by (meson Pi_iff le_add1 take_closed)
          then show ?thesis
            by (metis "0" add.commute cartesian_power_concat(2) cartesian_power_drop)
        qed
        show "f (take m (g (take k x) @ drop k x)) # drop m (g (take k x) @ drop k x)  S"
        using partial_pullback_memE[of x k "f  g" l S]
          comp_apply[of f g] partial_image_eq[of "take k x" k "drop k x" l x "f  g"]
        by (metis (no_types, lifting) A g (take k x) @ drop k x  carrier (Qpm+l)
            add_diff_cancel_left' append_eq_append_conv append_take_drop_id
            cartesian_power_car_memE length_drop local.partial_image_def)
      qed
    qed
  qed
  show "map_partial_pullback k g l (partial_pullback m f l S)  partial_pullback k (f  g) l S"
  proof fix x assume A: "x  map_partial_pullback k g l (partial_pullback m f l S)"
    have 0: "(take m (map_partial_image k g x)) = g (take k x)"
    proof-
      have "take k x  carrier (Qpk)"
        using map_partial_pullback_memE[of x k g l] A le_add1 take_closed
        by blast
      then have "length (g (take k x)) = m"
        using assms is_semialg_map_closed[of k m g] cartesian_power_car_memE
        by blast
      then show ?thesis
        using assms unfolding map_partial_image_def
        by (metis append_eq_conv_conj)
    qed
    show "x  partial_pullback k (f  g) l S"
      apply(rule partial_pullback_memI)
      using A map_partial_pullback_memE
       apply blast
      using 0 assms A comp_apply map_partial_pullback_memE[of x k g l "partial_pullback m f l S"]
            partial_pullback_memE[of "map_partial_image k g x" m f l S]
            map_partial_image_eq[of "take k x" k "drop k x" l x g]
   partial_image_eq[of "take m (map_partial_image k g x)" m "drop m (map_partial_image k g x)" l "(map_partial_image k g x)" f ]
  by (metis (no_types, lifting) cartesian_power_drop le_add1 map_partial_image_def partial_pullback_memE' take_closed)

qed
qed

lemma semialg_function_comp_closed:
  assumes "is_semialg_function m f"
  assumes "is_semialg_map k m g"
  shows "is_semialg_function k (f  g)"
proof(rule is_semialg_functionI)
  show "f  g  carrier (Qpk)  carrier Qp"
  proof fix x assume A: "x  carrier (Qpk)"
    show " (f  g) x  carrier Qp"
      using A assms is_semialg_map_closed[of k m g ] is_semialg_function_closed[of m f] comp_apply[of f g x]
      by (metis (no_types, lifting) Pi_mem)
  qed
  show " ka S. S  semialg_sets (1 + ka)  is_semialgebraic (k + ka) (partial_pullback k (f  g) ka S)"
  proof- fix n S assume A: "S  semialg_sets (1 + n)"
    show "is_semialgebraic (k + n) (partial_pullback k (f  g) n S)"
      using A assms partial_pullback_comp is_semialg_functionE is_semialg_mapE
        is_semialgebraicI  by presburger
  qed
qed

lemma semialg_map_evimage_is_semialg:
  assumes "is_semialg_map k m g"
  assumes "is_semialgebraic m S"
  shows "is_semialgebraic k (g  ¯kS)"
proof-
  have "g  ¯kS = map_partial_pullback k g 0 S"
  proof
    show "g ¯kS  map_partial_pullback k g 0 S"
    proof fix x assume A: "x  g  ¯kS"
      then have 0: "x  carrier (Qpk)  g x  S"
        by (meson evimage_eq)
      have "x = take k x @ drop k x"
        using 0  by (simp add: "0")
      then show "x  map_partial_pullback k g 0 S"
        unfolding map_partial_pullback_def map_partial_image_def
        by (metis (no_types, lifting) "0" Nat.add_0_right add.commute append_Nil2 append_same_eq
        append_take_drop_id evimageI2 map_partial_image_def map_partial_image_eq take0 take_drop)
    qed
    show "map_partial_pullback k g 0 S  g ¯kS"
    proof fix x assume A: "x  map_partial_pullback k g 0 S "
      then have 0: " g (take k x) @ (drop k x)  S"
        unfolding map_partial_pullback_def map_partial_image_def
        by blast
      have 1: "x  carrier (Qpk)"
        using A unfolding map_partial_pullback_def map_partial_image_def
        by (metis (no_types, lifting) Nat.add_0_right evimage_eq)
      hence "take k x = x"
        by (metis cartesian_power_car_memE le_eq_less_or_eq take_all)
      then show " x  g ¯kS"
        using 0 1 unfolding evimage_def
        by (metis (no_types, lifting) IntI append.assoc append_same_eq append_take_drop_id same_append_eq vimageI)
    qed
  qed
  then show ?thesis using assms
    by (metis add.right_neutral is_semialg_mapE' map_partial_pullback_def)
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection ‹Examples of Semialgebraic Maps›
(**************************************************************************************************)
(**************************************************************************************************)


lemma semialg_map_on_carrier:
  assumes "is_semialg_map n m f"
  assumes "restrict f (carrier (Qpn)) = restrict g (carrier (Qpn))"
  shows "is_semialg_map n m g"
proof(rule is_semialg_mapI)
  have 0: "f  carrier (Qpn)  carrier (Qpm)"
    using assms(1) is_semialg_map_closed
    by blast
  show "g  carrier (Qpn)  carrier (Qpm)"
  proof fix x assume A: "x  carrier (Qpn)" then show " g x  carrier (Qpm)"
      using assms(2) 0
      by (metis (no_types, lifting) PiE restrict_Pi_cancel)
  qed
  show "k S. S  semialg_sets (m + k)  is_semialgebraic (n + k) (map_partial_pullback n g k S)"
  proof- fix k S
    assume A: "S  semialg_sets (m + k)"
    have 1: "is_semialgebraic (n + k) (map_partial_pullback n f k S)"
      using A assms(1) is_semialg_mapE is_semialgebraicI
      by blast
    have 2: "(map_partial_pullback n f k S) = (map_partial_pullback n g k S)"
      unfolding map_partial_pullback_def map_partial_image_def evimage_def
    proof
      show "(λxs. f (take n xs) @ drop n xs) -` S  carrier (Qpn+k)  (λxs. g (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
      proof fix x assume A: "x  (λxs. f (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
        have "(take n x)  carrier (Qpn)"
          using assms A
          by (meson Int_iff le_add1 take_closed)
        then have "f (take n x) = g (take n x)"
          using assms unfolding  restrict_def
          by meson
        then show "x  (λxs. g (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
          using assms
          by (metis (no_types, lifting) A Int_iff vimageE vimageI)
      qed
      show " (λxs. g (take n xs) @ drop n xs) -` S  carrier (Qpn+k)  (λxs. f (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
      proof fix x assume A: "x  (λxs. g (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
        have "(take n x)  carrier (Qpn)"
          using assms
          by (meson A  inf_le2 le_add1 subset_iff take_closed)
        then have "f (take n x) = g (take n x)"
          using assms unfolding  restrict_def
          by meson
        then show " x  (λxs. f (take n xs) @ drop n xs) -` S  carrier (Qpn+k)"
          using A by blast
      qed
    qed
    then show "is_semialgebraic (n + k) (map_partial_pullback n g k S)"
      using 1 by auto
  qed
qed

lemma semialg_function_tuple_is_semialg_map:
  assumes "is_semialg_function_tuple m fs"
  assumes "length fs = n"
  shows "is_semialg_map m n (function_tuple_eval Qp m fs)"
  apply(rule is_semialg_mapI)
  using function_tuple_eval_closed[of m fs]
  apply (metis Pi_I assms(1) assms(2) semialg_function_tuple_is_function_tuple)
proof- fix k S assume A: "S  semialg_sets (n + k)"
  show "is_semialgebraic (m + k) (map_partial_pullback m (function_tuple_eval Qp m fs) k S)"
  using is_semialg_map_tupleE[of m fs k S] assms A tuple_partial_pullback_is_semialg_map_tuple[of m fs]
  unfolding tuple_partial_pullback_def map_partial_pullback_def
        map_partial_image_def tuple_partial_image_def is_semialgebraic_def
  by (metis evimage_def)
qed

lemma index_is_semialg_function:
  assumes "n > k"
  shows "is_semialg_function n (λas. as!k)"
proof-
  have 0: "restrict (λas. as!k) (carrier (Qpn)) = restrict (Qp_ev (pvar Qp k)) (carrier (Qpn))"
    using assms by (metis (no_types, lifting) eval_pvar restrict_ext)
  have 1: "is_semialg_function n (Qp_ev (pvar Qp k))"
    using pvar_closed   assms poly_is_semialg[of  "pvar Qp k"] by blast
  show ?thesis
    using 0 1 semialg_function_on_carrier[of n "Qp_ev (pvar Qp k)" "(λas. as!k)"]
    by presburger
qed

definition Qp_ith where
"Qp_ith m i = (λ x  carrier (Qpm). x!i)"

lemma Qp_ith_closed:
  assumes "i < m"
  shows "Qp_ith m i  carrier (SA m)"
proof(rule SA_car_memI)
  show "Qp_ith m i  carrier (function_ring (carrier (Qpm)) Qp)"
    apply(rule Qp.function_ring_car_memI[of "carrier(Qpm)"])
    using assms cartesian_power_car_memE'[of _ Qp m i] unfolding Qp_ith_def
    apply (metis restrict_apply)
    unfolding restrict_def by meson
  have 0: "x. x  carrier (Qpm)  Qp_ith m i x = eval_at_point Qp  x (pvar Qp i)"
    using assms eval_pvar[of i m] unfolding Qp_ith_def restrict_def by presburger
  have 1: "is_semialg_function m (eval_at_poly Qp  (pvar Qp i))"
    using assms pvar_closed[of i m] poly_is_semialg by blast
  show "is_semialg_function m (local.Qp_ith m i)"
    apply(rule semialg_function_on_carrier'[of m "eval_at_poly Qp  (pvar Qp i)"])
    using 1 apply blast
    using 0 by blast
qed

lemma take_is_semialg_map:
  assumes "n  k"
  shows "is_semialg_map n k (take k)"
proof-
  obtain fs where fs_def: "fs = map (λi::nat. (λas  carrier (Qpn). as!i)) [0::nat..< k]"
    by blast
  have 0: "is_semialg_function_tuple n fs"
  proof(rule is_semialg_function_tupleI)
    fix f assume A: "f  set fs"
    then obtain i where i_def: "i  set [0::nat..< k]  f = (λas  carrier (Qpn). as!i)"
      using A fs_def
      by (metis (no_types, lifting) in_set_conv_nth length_map nth_map)
    have i_less: "i < k"
    proof-
      have "set [0::nat..< k] = {0..<k}"
        using atLeastLessThan_upt by blast
      then show ?thesis using i_def
        using atLeastLessThan_iff by blast
    qed
    show "is_semialg_function n f"
      apply(rule semialg_function_on_carrier[of n "(λas. as ! i)"],
            rule index_is_semialg_function[of i n ]  )
      using A i_def assms by auto
  qed
  have 1: "restrict (function_tuple_eval Qp n fs) (carrier (Qpn)) = restrict (take k) (carrier (Qpn))"
  unfolding function_tuple_eval_def
  proof fix x
    show " (λxcarrier (Qpn). map (λf. f x) fs) x = restrict (take k) (carrier (Qpn)) x"
  proof(cases "x  carrier (Qpn)")
    case True
    have "(map (λf. f x) fs) = take k x"
    proof-
      have "i. i < k  (map (λf. f x) fs) ! i = x ! i"
      proof- fix i assume A: "i < k"
        have "length [0::nat..< k] = k"
          using assms by simp
        then have "length fs = k"
          using fs_def
          by (metis length_map)
        then have 0: "(map (λf. f x) fs) ! i = (fs!i) x"
          using A by (meson nth_map)
        have 1: "(fs!i) x = x!i"
          using A nth_map[of i "[0..<k]" "(λi. λascarrier (Qpn). as ! i)"] True
          unfolding fs_def  restrict_def  by auto
        then show "map (λf. f x) fs ! i = x ! i"
          using 0 assms A True fs_def nth_map[of i fs] cartesian_power_car_memE[of x Qp n]
          by blast
      qed
      then have 0: "i. i < k  (map (λf. f x) fs) ! i = (take k x) ! i"
        using assms True  nth_take by blast
      have 1: "length (map (λf. f x) fs) = length (take k x)"
        using fs_def True assms
        by (metis cartesian_power_car_memE length_map map_nth take_closed)
      have 2: "length (take k x) = k"
        using assms True cartesian_power_car_memE take_closed
        by blast
      show ?thesis using 0 1 2
        by (metis nth_equalityI)
    qed
    then show ?thesis using True unfolding restrict_def
      by presburger
  next
    case False
    then show ?thesis unfolding restrict_def
      by (simp add: False)
  qed
  qed
  have 2: " is_semialg_map n k (function_tuple_eval Qp n fs)"
    using 0 semialg_function_tuple_is_semialg_map[of n fs k] assms fs_def  length_map
    by (metis (no_types, lifting) diff_zero length_upt)
  show ?thesis using 1 2
    using semialg_map_on_carrier by blast
qed

lemma drop_is_semialg_map:
  shows "is_semialg_map (k + n) n (drop k)"
proof-
  obtain fs where fs_def: "fs = map (λi::nat. (λas  carrier (Qpn+k). as!i)) [k..<n+k]"
    by blast
  have 0: "is_semialg_function_tuple (k+n) fs"
  proof(rule is_semialg_function_tupleI)
    fix f assume A: "f  set fs"
    then obtain i where i_def: "i  set [k..<n+k]  f = (λas  carrier (Qpn+k). as!i)"
      using A fs_def
      by (metis (no_types, lifting) in_set_conv_nth length_map nth_map)
    have i_less: "i  k  i < n + k"
    proof-
      have "set [k..<n+k] = {k..<n+k}"
        using atLeastLessThan_upt by blast
      then show ?thesis using i_def
        using atLeastLessThan_iff by blast
    qed
    have "is_semialg_function (n + k) f"
      using A index_is_semialg_function[of i "n+k" ]
            i_less semialg_function_on_carrier[of "n+k" "(λas. as ! i)" f] i_def
            restrict_is_semialg
      by blast
    then show "is_semialg_function (k + n) f"
      by (simp add: add.commute)
  qed
  have 1: "restrict (function_tuple_eval Qp (n+k) fs) (carrier (Qpn+k)) = restrict (drop k) (carrier (Qpn+k))"
  unfolding function_tuple_eval_def
  proof fix x
    show " (λxcarrier (Qpn+k). map (λf. f x) fs) x = restrict (drop k) (carrier (Qpn+k)) x"
  proof(cases "xcarrier (Qpn+k)")
    case True
    have "(map (λf. f x) fs) = drop k x"
    proof-
      have "i. i < n  (map (λf. f x) fs) ! i = x ! (i+k)"
      proof- fix i assume A: "i < n"
        have 00: "length [k..<n+k] = n"
           by simp
        then have "length fs = n"
          using fs_def
          by (metis length_map)
        then have 0:"(map (λf. f x) fs) ! i = (fs!i) x"
          using A  by (meson nth_map)
        have "[k..<n+k]!i = i + k"
          by (simp add: A)
        have "( map (λi. λascarrier (Qpn+k). as ! i) [k..<n + k]) ! i = ((λi. λascarrier (Qpn+k). as ! i) ([k..<n + k] ! i))"
          using A 00 nth_map[of i "[k..< n + k]" "(λi. λascarrier (Qpn+k). as ! i)"]
          by linarith
        then have 1: "fs!i =  (λas  carrier (Qpn+k). as!(i+k))"
          using fs_def A  [k..<n + k] ! i = i + k by presburger
        then show "map (λf. f x) fs ! i = x ! (i + k)"
          using True 0
          by (metis (no_types, lifting) restrict_apply)
      qed
      then have 0: "i. i < n  (map (λf. f x) fs) ! i = (drop k x) ! i"
        using True  nth_drop
        by (metis add.commute cartesian_power_car_memE le_add2)
      have 1: "length (map (λf. f x) fs) = length (drop k x)"
        using fs_def True length_drop[of k x]
        by (metis cartesian_power_car_memE length_map length_upt)
      have 2: "length (drop k x) = n"
        using True cartesian_power_car_memE
        by (metis add_diff_cancel_right' length_drop)
      show ?thesis using 0 1 2
        by (metis nth_equalityI)
    qed
    then show ?thesis using True unfolding restrict_def
      by presburger
  next
    case False
    then show ?thesis unfolding restrict_def
      by (simp add: False)
  qed
  qed
  have 00: "length [k..<n+k] = n"
    by simp
  then have "length fs = n"
    using fs_def
    by (metis length_map)
  then have 2: " is_semialg_map (k + n) n (function_tuple_eval Qp (k + n) fs)"
    using 0 semialg_function_tuple_is_semialg_map[of "k + n" fs n]
    by blast
  show ?thesis using 1 2
    using semialg_map_on_carrier[of "k + n" n "function_tuple_eval Qp (k + n) fs" "drop k"]
    by (metis add.commute)
qed

lemma project_at_indices_is_semialg_map:
  assumes "S  {..<n}"
  shows "is_semialg_map n (card S) π⇘S⇙"
proof-
  obtain k where k_def: "k = card S"
    by blast
  have 0: "card {..<n} = n"
    by simp
  have 1: "finite S"
    using assms finite_subset
    by blast
  have 2: "card S  n"
    using assms 0 1
    by (metis card_mono finite_lessThan)
  then have k_size: " k  n"
    using k_def assms 0 1 2
    by blast
  obtain fs where fs_def: "fs = map (λi::nat. (λas  carrier (Qpn). as!(nth_elem S i))) [0..<k]"
    by blast
  have 4: "length fs = k"
    using fs_def assms "1" k_def
    by (metis Ex_list_of_length length_map map_nth)
  have 5: "is_semialg_function_tuple n fs"
  proof(rule is_semialg_function_tupleI)
    fix f assume A: "f  set fs"
    then obtain i where i_def: "i  set [0..<k]  f = (λas  carrier (Qpn). as!(nth_elem S i))"
      using A fs_def atLeast_upt "4"  in_set_conv_nth  map_eq_conv map_nth
      by auto
    have i_le_k:"i < k"
    proof-
      have "set [0..<k] = {..<k}"
        using atLeast_upt by blast
      then show ?thesis
        using i_def
        by blast
    qed
    then have i_in_S: "nth_elem S i  S"
      using "1" k_def nth_elem_closed by blast
    then have "nth_elem S i < n"
      using assms
      by blast
    show "is_semialg_function n f"
      using A index_is_semialg_function[of "nth_elem S i" n]
            semialg_function_on_carrier[of n "(λas. as ! nth_elem S i)"] i_def restrict_is_semialg
            nth_elem S i < n by blast
  qed
  have 6: "restrict (function_tuple_eval Qp n fs) (carrier (Qpn)) = restrict (π⇘S) (carrier (Qpn))"
  unfolding function_tuple_eval_def
  proof fix x
    show " (λxcarrier (Qpn). map (λf. f x) fs) x = restrict (π⇘S) (carrier (Qpn)) x"
    proof(cases "x  carrier (Qpn)")
    case True
    have "(map (λf. f x) fs) = π⇘Sx"
    proof-
      have "i. i < k  (map (λf. f x) fs) ! i = (π⇘Sx) ! i"
      proof- fix i assume A: "i < k"
        have T0:"(map (λf. f x) fs) ! i = (fs!i) x"
          using A nth_map by (metis "4")
        have T1: "indices_of x = {..< n}"
          using True cartesian_power_car_memE indices_of_def
          by blast
        have T2: "set (set_to_list S)  indices_of x"
          using assms True  by (simp add: True "1" T1)
        have T3: "length x = n"
          using True cartesian_power_car_memE by blast
        have T4: "([0..<k] ! i) =   i"
          using A by simp
        have T5: "nth_elem S i < n"
          using assms 0 1 2 A k_def
          by (meson lessThan_iff nth_elem_closed subsetD)
        have T6: "nth_elem S ([0..<k] ! i) = nth_elem S i"
          by (simp add: T4)
        have T6: "(λascarrier (Qpn). as ! nth_elem S ([0..<k] ! i)) x = x! (nth_elem S i)"
        proof-
          have "(λascarrier (Qpn). as ! nth_elem S ([0..<k] ! i)) x = x! nth_elem S ([0..<k] ! i)"
            using True restrict_def by metis
          then show ?thesis
            using A T3 T4 0 1 2 T5 T6  True
            by metis
        qed
        have T7: " map (λf. f x) fs ! i = x! (nth_elem S i)"
          using fs_def T0 A nth_map[of i "[0..<k]" "(λi. λascarrier (Qpn). as ! nth_elem S i)"]
          by (metis "4" T6 length_map)
        show "map (λf. f x) fs ! i = π⇘Sx ! i"
          using True T0 T1 T2 fs_def 5 unfolding T7
          by (metis A assms(1) k_def project_at_indices_nth)
      qed
      have 1: "length (map (λf. f x) fs) = length (π⇘Sx)"
        using fs_def True assms proj_at_index_list_length[of S x] k_def
        by (metis "4" cartesian_power_car_memE indices_of_def length_map)
      have 2: "length (π⇘Sx) = k"
        using assms True 0 1 2
        by (metis "4" length_map)
      show ?thesis using  1 2
        by (metis i. i < k  map (λf. f x) fs ! i = π⇘Sx ! i nth_equalityI)
    qed
    then show ?thesis using True unfolding restrict_def
      by presburger
    next
    case False
    then show ?thesis unfolding restrict_def
      by (simp add: False)
    qed
  qed
  have 7: " is_semialg_map n k (function_tuple_eval Qp n fs)"
    using 0 semialg_function_tuple_is_semialg_map[of n fs k] assms fs_def  length_map "4" "5" k_size
    by blast
  show ?thesis using 6 7
    using semialg_map_on_carrier k_def
    by blast
qed

lemma tl_is_semialg_map:
  shows "is_semialg_map (Suc n) n tl"
proof-
  have 0: "(card {1..<Suc n}) = n"
  proof-
    have "Suc n - 1 = n"
      using diff_Suc_1 by blast
    then show ?thesis
      by simp
  qed
  have 3: "{1..<Suc n}  {..<Suc n}"
    using atLeastLessThan_iff by blast
  have 4: " is_semialg_map (Suc n) n (project_at_indices {1..<Suc n})"
    using 0 project_at_indices_is_semialg_map
    by (metis "3")
  show ?thesis
    using 0 3 4
        semialg_map_on_carrier[of "Suc n" n "(project_at_indices {1..<Suc n})" tl]
    unfolding restrict_def
    by (metis (no_types, lifting) tl_as_projection)
qed

text‹Coordinate functions are semialgebraic maps.›

lemma coord_fun_is_SA:
  assumes "is_semialg_map n m g"
  assumes "i < m"
  shows "coord_fun Qp n g i  carrier (SA n)"
proof(rule SA_car_memI)
  show "coord_fun Qp n g i  carrier (function_ring (carrier (Qpn)) Qp)"
    apply(rule Qp.function_ring_car_memI)
    unfolding coord_fun_def using assms
     apply (metis (no_types, lifting) Pi_iff cartesian_power_car_memE' is_semialg_map_closed restrict_apply')
      by (meson restrict_apply)
  show "is_semialg_function n (coord_fun Qp n g i)"
  proof-
    have 0: "is_semialg_function m (λ x. x ! i)"
      using assms gr_implies_not0 index_is_semialg_function by blast
    have 1: "(coord_fun Qp n g i) = restrict ((λx.  x ! i)  g) (carrier (Qpn)) "
      unfolding coord_fun_def using comp_apply
      by metis
    show ?thesis
      using semialg_function_on_carrier[of n "((λx. x ! i)  g)" "coord_fun Qp n g i"]
             assms semialg_function_comp_closed[of m "λx. x ! i" n g] assms 0 1
      unfolding coord_fun_def
      using restrict_is_semialg by presburger
  qed
qed

lemma coord_fun_map_is_semialg_tuple:
  assumes "is_semialg_map n m g"
  shows "is_semialg_function_tuple n (map (coord_fun Qp n g) [0..<m])"
proof(rule is_semialg_function_tupleI)
  have 0: "set [0..<m] = {..<m}"
    using atLeast_upt by blast
  fix f assume A: "f  set (map (coord_fun Qp n g) [0..<m])"
  then obtain i where i_def: "i < m  f = coord_fun Qp n g i"
    using set_map[of "coord_fun Qp n g" "[0..<m]"] 0
    by (metis image_iff lessThan_iff)
  show " is_semialg_function n f"
  using i_def A assms coord_fun_is_SA[of n m g i] SA_imp_semialg by blast
qed

lemma semialg_map_cons:
  assumes "is_semialg_map n m g"
  assumes "f  carrier (SA n)"
  shows "is_semialg_map n (Suc m) (λ x  carrier (Qpn). f x # g x)"
proof-
  obtain Fs where Fs_def: "Fs = f # (map (coord_fun Qp n g) [0..<m])"
    by blast
  have 0: "is_semialg_function_tuple n Fs"
    apply(rule  is_semialg_function_tupleI)
    using is_semialg_function_tupleE'[of n "map (coord_fun Qp n g) [0..<m]"]
          coord_fun_map_is_semialg_tuple[of n m g] assms SA_car_memE(1)[of f n]
          set_ConsD[of _ f "map (coord_fun Qp n g) [0..<m]"] assms
    unfolding Fs_def by blast
  have 1: "(λ x  carrier (Qpn). f x # g x) = (λ x  carrier (Qpn). function_tuple_eval Qp n Fs x)"
  proof(rule ext) fix x show "(λxcarrier (Qpn). f x # g x) x = restrict (function_tuple_eval Qp n Fs) (carrier (Qpn)) x"
    proof(cases "x  carrier (Qpn)")
      case True then have T0: "(λxcarrier (Qpn). f x # g x) x = f x # g x"
        by (meson restrict_apply')
      have T1: "restrict (function_tuple_eval Qp n Fs) (carrier (Qpn)) x = function_tuple_eval Qp n Fs x"
        using True by (meson restrict_apply')
      hence T2: "restrict (function_tuple_eval Qp n Fs) (carrier (Qpn)) x = f x # (function_tuple_eval Qp n (map (coord_fun Qp n g) [0..<m]) x)"
        unfolding function_tuple_eval_def Fs_def by (metis (no_types, lifting) list.simps(9))
      have T3: "g x  carrier (Qpm)"
        using assms True is_semialg_map_closed by blast
      have "length [0..<m] = m"
        by auto
      hence T4: "length (map (coord_fun Qp n g) [0..<m]) = m"
        using length_map[of "(coord_fun Qp n g)" "[0..<m]"] by metis
      hence T5: "length (function_tuple_eval Qp n (map (coord_fun Qp n g) [0..<m]) x) = m"
        unfolding function_tuple_eval_def using length_map by metis
      have T6: "(function_tuple_eval Qp n (map (coord_fun Qp n g) [0..<m]) x) = g x"
        apply(rule nth_equalityI) using T3 T5
          using cartesian_power_car_memE apply blast
          using cartesian_power_car_memE[of "g x" Qp m]  T5 T4  T3 True
                nth_map[of _ "(map (λi. λxcarrier (Qpn). g x ! i) [0..<m])" "(λf. f x)"]
                nth_map[of _ "[0..<m]" "(λi. λxcarrier (Qpn). g x ! i)"]
          unfolding function_tuple_eval_def  coord_fun_def restrict_def
          by (metis (no_types, lifting) length [0..<m] = m map_nth nth_map)
      hence T7: "restrict (function_tuple_eval Qp n Fs) (carrier (Qpn)) x = f x # g x"
        using T4 T2 by presburger
      thus ?thesis using T0
        by presburger
    next
      case False
      then show ?thesis unfolding restrict_def
        by metis
    qed
  qed
  have 2: "length Fs = Suc m"
    unfolding Fs_def using length_map[of "(coord_fun Qp n g)" "[0..<m]"] length_Cons[of f "map (coord_fun Qp n g) [0..<m]"]
    using length_upt by presburger
  have 3: "is_semialg_map n (Suc m) (λ x  carrier (Qpn). function_tuple_eval Qp n Fs x)"
    apply(rule semialg_map_on_carrier[of _ _  "function_tuple_eval Qp n Fs"],
          intro semialg_function_tuple_is_semialg_map[of n Fs "Suc m"] 0 2)
    by auto
  show ?thesis using 1 3
    by presburger
qed

text‹Extensional Semialgebraic Maps:›

definition semialg_maps where
"semialg_maps n m  {f. is_semialg_map n m f  f  struct_maps (Qpn) (Qpm)}"

lemma semialg_mapsE:
  assumes "f  (semialg_maps n m)"
  shows "is_semialg_map n m f"
        "f  struct_maps (Qpn) (Qpm)"
        "f  carrier (Qpn)  carrier (Qpm)"
  using assms unfolding semialg_maps_def apply blast
  using assms unfolding semialg_maps_def apply blast
  apply(rule is_semialg_map_closed)
  using assms unfolding semialg_maps_def by blast

definition to_semialg_map where
"to_semialg_map n m f = restrict f (carrier (Qpn))"

lemma to_semialg_map_is_semialg_map:
  assumes "is_semialg_map n m f"
  shows "to_semialg_map n m f  semialg_maps n m"
  using assms unfolding to_semialg_map_def semialg_maps_def struct_maps_def mem_Collect_eq
  using is_semialg_map_closed' semialg_map_on_carrier by force

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Application of Functions to Segments of Tuples›
(**************************************************************************************************)
(**************************************************************************************************)


definition take_apply where
"take_apply m n f = restrict (f  take n) (carrier (Qpm))"

definition drop_apply where
"drop_apply m n f = restrict (f  drop n) (carrier (Qpm))"

lemma take_apply_closed:
  assumes "f  carrier (Fun⇘nQp)"
  assumes "k  n"
  shows "take_apply k n f  carrier (Fun⇘kQp)"
proof(rule Qp.function_ring_car_memI)
  show  "a. a  carrier (Qpk)  take_apply k n f a  carrier Qp"
  proof-  fix a assume A: "a  carrier (Qpk)" show "take_apply k n f a  carrier Qp"
      using A assms comp_apply[of f "take n" a] Qp.function_ring_car_memE[of f n] take_closed[of n k a]
      unfolding take_apply_def restrict_def
      by metis
  qed
  show " a. a  carrier (Qpk)  take_apply k n f a = undefined"
    unfolding take_apply_def restrict_def
    by meson
qed

lemma take_apply_SA_closed:
  assumes "f  carrier (SA n)"
  assumes "k  n"
  shows "take_apply k n f  carrier (SA k)"
  apply(rule SA_car_memI)
  using SA_car_memE(2) assms(1) assms(2) take_apply_closed apply blast
  using assms take_is_semialg_map[of n k] unfolding take_apply_def
  by (metis padic_fields.SA_imp_semialg
      padic_fields_axioms restrict_is_semialg semialg_function_comp_closed)

lemma drop_apply_closed:
  assumes "f  carrier (Fun⇘k - nQp)"
  assumes "k  n"
  shows "drop_apply k n f  carrier (Fun⇘kQp)"
proof(rule Qp.function_ring_car_memI)
  show  " a. a  carrier (Qpk)  drop_apply k n f a  carrier Qp"
  proof-  fix a assume A: "a  carrier (Qpk)" show "drop_apply k n f a  carrier Qp"
      using A assms comp_apply[of f "drop n" a] Qp.function_ring_car_memE[of f ] drop_closed[of n k a Qp]
      unfolding drop_apply_def restrict_def
      by (metis (no_types, opaque_lifting) Qp.function_ring_car_memE add_diff_cancel_right'
          cartesian_power_drop dec_induct diff_diff_cancel diff_less_Suc diff_less_mono2
          infinite_descent le_neq_implies_less less_antisym linorder_neqE_nat not_less0 not_less_eq_eq)
  qed
  show " a. a  carrier (Qpk)  drop_apply k n f a = undefined"
    unfolding drop_apply_def restrict_def
    by meson
qed

lemma drop_apply_SA_closed:
  assumes "f  carrier (SA (k-n))"
  assumes "k  n"
  shows "drop_apply k n f  carrier (SA k)"
  apply(rule SA_car_memI)
  using SA_car_memE(2) assms(1) assms(2) drop_apply_closed less_imp_le_nat apply blast
  using assms drop_is_semialg_map[of n "k - n" ] semialg_function_comp_closed[of "k - n" f k "drop n"] unfolding drop_apply_def
  by (metis (no_types, lifting) SA_imp_semialg le_add_diff_inverse  restrict_is_semialg)

lemma take_apply_apply:
  assumes "f  carrier (SA n)"
  assumes "a  carrier (Qpn)"
  assumes "b  carrier (Qpk)"
  shows "take_apply (n+k) n f (a@b) = f a"
proof-
  have "a@b  carrier (Qpn+k)"
    using assms cartesian_power_concat(1) by blast
  thus ?thesis
    unfolding take_apply_def restrict_def
    using assms cartesian_power_car_memE comp_apply[of f "take n"]
    by (metis append_eq_conv_conj)
qed

lemma drop_apply_apply:
  assumes "f  carrier (SA k)"
  assumes "a  carrier (Qpn)"
  assumes "b  carrier (Qpk)"
  shows "drop_apply (n+k) n f (a@b) = f b"
proof-
  have "a@b  carrier (Qpn+k)"
    using assms cartesian_power_concat(1) by blast
  thus ?thesis
    unfolding drop_apply_def  restrict_def
    using assms cartesian_power_car_memE comp_apply[of f "drop n"]
    by (metis append_eq_conv_conj)
qed

lemma drop_apply_add:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  shows "drop_apply (n+k) k (f SA ng) = drop_apply (n+k) k f SA (n + k)drop_apply (n+k) k g"
  apply(rule function_ring_car_eqI[of _ "n + k"])
  using drop_apply_SA_closed  assms fun_add_closed SA_car_memE(2) SA_plus diff_add_inverse2 drop_apply_closed le_add2 apply presburger
  using drop_apply_SA_closed assms fun_add_closed SA_car_memE(2) SA_plus diff_add_inverse2 drop_apply_closed le_add2 apply presburger
proof-
  fix a  assume A: "a  carrier (Qpn+k)"
  then obtain  b c where bc_def: "b  carrier (Qpk)  c  carrier (Qpn)  a  =  b@c"
    by (metis (no_types, lifting) add.commute cartesian_power_decomp)
  have 0: "drop_apply (n + k) k (f SA ng) a = f c  g c"
    using assms  bc_def drop_apply_apply[of "f SA ng" n b k c ]
    by (metis SA_add SA_imp_semialg add.commute padic_fields.SA_add_closed_left padic_fields_axioms)
  then show " drop_apply (n + k) k (f SA ng) a = (drop_apply (n + k) k f SA (n + k)drop_apply (n + k) k g) a"
    using bc_def drop_apply_apply assms
    by (metis A SA_add add.commute)
qed

lemma drop_apply_mult:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  shows "drop_apply (n+k) k (f SA ng) = drop_apply (n+k) k f SA (n + k)drop_apply (n+k) k g"
  apply(rule function_ring_car_eqI[of _ "n + k"])
  using drop_apply_SA_closed  assms fun_mult_closed SA_car_memE(2) SA_times diff_add_inverse2 drop_apply_closed le_add2 apply presburger
  using drop_apply_SA_closed assms fun_mult_closed SA_car_memE(2)  SA_times diff_add_inverse2 drop_apply_closed le_add2 apply presburger
proof-
  fix a  assume A: "a  carrier (Qpn+k)"
  then obtain  b c where bc_def: "b  carrier (Qpk)  c  carrier (Qpn)  a  =  b@c"
    by (metis (no_types, lifting) add.commute cartesian_power_decomp)
  have 0: "drop_apply (n + k) k (f SA ng) a = f c  g c"
    using assms  bc_def drop_apply_apply[of "f SA ng" n b k c ]
    by (metis SA_imp_semialg SA_mult SA_mult_closed_right add.commute)
  then show " drop_apply (n + k) k (f SA ng) a = (drop_apply (n + k) k f SA (n + k)drop_apply (n + k) k g) a"
    using bc_def drop_apply_apply assms by (metis A SA_mult add.commute)
qed

lemma drop_apply_one:
  shows "drop_apply (n+k) k 𝟭SA n= 𝟭SA (n+k)⇙"
  apply(rule function_ring_car_eqI[of _  "n + k"])
  apply (metis function_one_closed SA_one add_diff_cancel_right' drop_apply_closed le_add2)
  using function_one_closed SA_one apply presburger
proof-
  fix a assume A: "a  carrier (Qpn+k)"
  show "drop_apply (n + k) k 𝟭SA na = 𝟭SA (n + k)a"
  unfolding drop_apply_def restrict_def
  using  SA_one[of "n+k"] SA_one[of n] comp_apply[of "𝟭SA n⇙" "drop k" a] drop_closed[of k "n+k" a Qp]
        function_ring_defs(4)
  unfolding function_one_def
  by (metis A function_one_eval add.commute cartesian_power_drop)
qed

lemma drop_apply_is_hom:
  shows "drop_apply (n + k) k  ring_hom (SA n) (SA (n + k))"
  apply(rule ring_hom_memI)
  using drop_apply_SA_closed[of _ "k+n" k]
  apply (metis add.commute add_diff_cancel_left' le_add1)
  using drop_apply_mult apply blast
   using drop_apply_add apply blast
     using drop_apply_one by blast

lemma take_apply_add:
  assumes "f  carrier (SA k)"
  assumes "g  carrier (SA k)"
  shows "take_apply (n+k) k (f SA kg) = take_apply (n+k) k f SA (n + k)take_apply (n+k) k g"
  apply(rule function_ring_car_eqI[of _ "n + k"])
  using take_apply_SA_closed  assms fun_add_closed SA_car_memE(2) SA_plus diff_add_inverse2 take_apply_closed le_add2 apply presburger
  using take_apply_SA_closed assms fun_add_closed SA_car_memE(2) SA_plus diff_add_inverse2 take_apply_closed le_add2 apply presburger
proof-
  fix a  assume A: "a  carrier (Qpn+k)"
  then obtain  b c where bc_def: "b  carrier (Qpk)  c  carrier (Qpn)  a  =  b@c"
    by (metis (no_types, lifting) add.commute cartesian_power_decomp)
  hence 0: "take_apply (n + k) k (f SA kg) a =  f b  g b"
    using assms bc_def take_apply_apply[of "f SA kg" k b c ]
    by (metis SA_add SA_imp_semialg add.commute padic_fields.SA_add_closed_left padic_fields_axioms)
  then show "take_apply (n + k) k (f SA kg) a = (take_apply (n + k) k f SA (n + k)take_apply (n + k) k g) a"
    using bc_def take_apply_apply assms
    by (metis A SA_add add.commute)
qed

lemma take_apply_mult:
  assumes "f  carrier (SA k)"
  assumes "g  carrier (SA k)"
  shows "take_apply (n+k) k (f SA kg) = take_apply (n+k) k f SA (n + k)take_apply (n+k) k g"
  apply(rule function_ring_car_eqI[of _ "n + k"])
  using take_apply_SA_closed  assms fun_mult_closed SA_car_memE(2) SA_times diff_add_inverse2 take_apply_closed le_add2 apply presburger
  using take_apply_SA_closed assms fun_mult_closed SA_car_memE(2) SA_times diff_add_inverse2 take_apply_closed le_add2 apply presburger
proof-
  fix a  assume A: "a  carrier (Qpn+k)"
  then obtain  b c where bc_def: "b  carrier (Qpk)  c  carrier (Qpn)  a  =  b@c"
    by (metis (no_types, lifting) add.commute cartesian_power_decomp)
  hence 0: "take_apply (n + k) k (f SA kg) a =  f b  g b"
    using assms bc_def take_apply_apply[of "f SA kg" k b c ]
    by (metis SA_mult SA_imp_semialg add.commute padic_fields.SA_mult_closed_left padic_fields_axioms)
  then show "take_apply (n + k) k (f SA kg) a = (take_apply (n + k) k f SA (n + k)take_apply (n + k) k g) a"
    using bc_def take_apply_apply assms
    by (metis A SA_mult add.commute)
qed

lemma take_apply_one:
  shows "take_apply (n+k) k 𝟭SA k= 𝟭SA (n+k)⇙"
  apply(rule function_ring_car_eqI[of _  "n + k"])
  using function_one_closed SA_one le_add2 take_apply_closed apply presburger
  using function_one_closed SA_one apply presburger
proof-
  fix a assume A: "a  carrier (Qpn+k)"
  show "take_apply (n + k) k 𝟭SA ka = 𝟭SA (n + k)a"
  unfolding take_apply_def restrict_def
  using  SA_one[of "n+k"] SA_one[of k] comp_apply[of "𝟭SA k⇙" "take k" a] take_closed[of k "n + k" a]
        function_ring_defs(4)
  unfolding function_one_def
  using A function_one_eval le_add2 by metis
qed

lemma take_apply_is_hom:
  shows "take_apply (n + k) k  ring_hom (SA k) (SA (n + k))"
  apply(rule ring_hom_memI)
  using take_apply_SA_closed[of _ k "n+k"]  le_add2 apply blast
  using   take_apply_mult apply blast
   using   take_apply_add apply blast
     using   take_apply_one by blast

lemma drop_apply_units:
  assumes "f  Units (SA n)"
  shows "drop_apply (n+k) k f  Units (SA (n+k))"
  apply(rule SA_Units_memI)
  using drop_apply_SA_closed[of f "n+k" k ] assms SA_Units_closed
  apply (metis add_diff_cancel_right' le_add2)
proof-
  show "x. x  carrier (Qpn + k)  drop_apply (n + k) k f x  𝟬"
  proof- fix x assume A: "x  carrier (Qpn+k)"
    then have "drop_apply (n + k) k f x = f (drop k x)"
      unfolding drop_apply_def restrict_def by (meson comp_def)
    then show "drop_apply (n + k) k f x  𝟬"
      using SA_Units_memE'[of f n "drop k x"]
      by (metis A add.commute assms  cartesian_power_drop)
  qed
qed

lemma take_apply_units:
  assumes "f  Units (SA k)"
  shows "take_apply (n+k) k f  Units (SA (n+k))"
  apply(rule SA_Units_memI)
  using take_apply_SA_closed[of f k "n+k" ] assms SA_Units_closed le_add2 apply blast
proof-
  show "x. x  carrier (Qpn + k)  take_apply (n + k) k f x  𝟬"
  proof- fix x assume A: "x  carrier (Qpn+k)"
    then have "take_apply (n + k) k f x = f (take k x)"
      unfolding take_apply_def restrict_def by (meson comp_def)
    then show "take_apply (n + k) k f x  𝟬"
      using SA_Units_memE'[of f k "take k x"]  A assms le_add2 local.take_closed by blast
  qed
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Level Sets of Semialgebraic Functions›
(**************************************************************************************************)
(**************************************************************************************************)


lemma evimage_is_semialg:
  assumes "h  carrier (SA n)"
  assumes "is_univ_semialgebraic S"
  shows "is_semialgebraic n (h  ¯nS)"
proof-
  have 0: "is_semialgebraic 1 (to_R1 ` S)"
    using assms is_univ_semialgebraicE by blast
  have 1: "h  ¯nS = partial_pullback n h (0::nat) (to_R1 ` S)"
  proof show "h ¯nS  partial_pullback n h 0 ((λa. [a]) ` S)"
    proof fix x assume A: "x  h ¯nS"
      then have 0: "h x  S" by blast
      have x_closed: "x  carrier (Qpn)"
        by (meson A evimage_eq)
      have 1: "drop n x = []"
        using cartesian_power_car_memE[of x Qp n] drop_all x_closed
        by blast
      have 2: "take n x = x"
        using cartesian_power_car_memE[of x Qp n] x_closed take_all
        by blast
      then show "x  partial_pullback n h 0 ((λa. [a]) ` S)"
        unfolding partial_pullback_def partial_image_def evimage_def
        using 0 1 2 x_closed
        by (metis (no_types, lifting) IntI Nat.add_0_right image_iff vimageI)
    qed
    show "partial_pullback n h 0 ((λa. [a]) ` S)  h ¯nS"
    proof fix x assume A: "x  partial_pullback n h 0 ((λa. [a]) ` S)"
      have x_closed: "x  carrier (Qpn)"
        using A unfolding partial_pullback_def evimage_def
        by (metis A Nat.add_0_right partial_pullback_memE(1))
      then have "(partial_image n h) x = [h x]"
        unfolding partial_image_def
        by (metis (no_types, opaque_lifting) One_nat_def Qp.zero_closed append.right_neutral append_Nil
            local.partial_image_def partial_image_eq segment_in_car' Qp.to_R1_closed)
      then have "h x  S"
        using A unfolding partial_pullback_def
        by (metis (no_types, lifting) A image_iff partial_pullback_memE(2) Qp.to_R_to_R1)
      thus "x  h ¯nS"
        using x_closed by blast
    qed
  qed
  then show ?thesis
    using 0 is_semialg_functionE[of n h 0 "((λa. [a]) ` S)"] assms SA_car_memE(1)[of h n]
    by (metis Nat.add_0_right SA_car)
qed

lemma semialg_val_ineq_set_is_semialg:
  assumes "g  carrier (SA k)"
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (g x)  val (f x)}"
proof-
  obtain F where F_def: "F = function_tuple_eval Qp  k [f, g]"
    by blast
  have P0: "is_semialg_function_tuple k [f, g] "
    using is_semialg_function_tupleI[of "[f, g]" k]
    by (metis assms  list.distinct(1) list.set_cases padic_fields.SA_imp_semialg padic_fields_axioms set_ConsD)
  hence P1: "is_semialg_map k 2 F"
    using  assms semialg_function_tuple_is_semialg_map[of k "[f, g]" 2]
    unfolding F_def by (simp add: f  carrier (SA k) g  carrier (SA k))
  have "{x  carrier (Qpk). val (g x)  val (f x)} = F  ¯kval_relation_set"
  proof
    show "{x  carrier (Qpk). val (g x)  val (f x)}  F ¯kval_relation_set"
    proof fix x assume A: "x  {x  carrier (Qpk). val (g x)  val (f x)}"
      then have 0: "x  carrier (Qpk)  val (g x)  val (f x)" by blast
      have 1: "F x = [f x, g x]"
        unfolding F_def using A unfolding function_tuple_eval_def
        by (metis (no_types, lifting) list.simps(8) map_eq_Cons_conv)
      have 2: "val (g x)  val (f x)"
        using A
        by blast
      have 3: "F x  carrier (Qp2)"
        using assms A 1
        by (metis (no_types, lifting) "0" Qp.function_ring_car_mem_closed Qp_2I SA_car_memE(2))
      then have 4: "x  carrier (Qpk)  F x  val_relation_set"
        unfolding val_relation_set_def F_def using 0 1 2 3
        by (metis (no_types, lifting) cartesian_power_car_memE cartesian_power_car_memE'
            list.inject local.F_def one_less_numeral_iff pair_id semiring_norm(76) val_relation_setI
            val_relation_set_def zero_less_numeral)
      then show "x  F ¯kval_relation_set"
        by blast
    qed
    show "F ¯kval_relation_set  {x  carrier (Qpk). val (g x)  val (f x)}"
    proof fix x assume A: "x  F ¯kval_relation_set"
      then have 0: "x  carrier (Qpk)  F x  val_relation_set"
        by (meson evimage_eq)
      then have 0: "x  carrier (Qpk)  [f x, g x]  val_relation_set"
        unfolding F_def function_tuple_eval_def
        by (metis (no_types, lifting) list.simps(8) list.simps(9))
      then have "x  carrier (Qpk)  val (g x)  val (f x)"
        unfolding F_def val_relation_set_def
        by (metis (no_types, lifting) "0" list.inject val_relation_setE)
      then show "x  {x  carrier (Qpk). val (g x)  val (f x)}"
        by blast
    qed
  qed
  then show ?thesis
    using assms P0 P1 val_relation_is_semialgebraic semialg_map_evimage_is_semialg[of k 2 F val_relation_set] pos2
    by presburger
qed

lemma semialg_val_ineq_set_is_semialg':
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (f x)  C}"
proof-
  obtain a where a_def: "a  carrier Qp  val a = C"
    by (meson Qp.carrier_not_empty Qp.minus_closed dist_nonempty' equals0I)
  then obtain g where g_def: "g = constant_function (carrier (Qpk)) a"
    by blast
  have 0: "g  carrier (SA k)"
    using g_def a_def SA_car assms(1) constant_function_in_semialg_functions by blast
  have 1: "{x  carrier (Qpk). val (f x)  C} = {x  carrier (Qpk). val (f x)  val (g x)}"
    using g_def by (metis (no_types, lifting) Qp_constE a_def)
  then show ?thesis using assms 0 semialg_val_ineq_set_is_semialg[of f k g]
    by presburger
qed

lemma semialg_val_ineq_set_is_semialg'':
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (f x)  C}"
proof-
  obtain a where a_def: "a  carrier Qp  val a = C"
    by (meson Qp.carrier_not_empty Qp.minus_closed dist_nonempty' equals0I)
  then obtain g where g_def: "g = constant_function (carrier (Qpk)) a"
    by blast
  have 0: "g  carrier (SA k)"
    using g_def a_def SA_car assms(1) constant_function_in_semialg_functions by blast
  have 1: "{x  carrier (Qpk). val (f x)  C} = {x  carrier (Qpk). val (f x)  val (g x)}"
    using g_def by (metis (no_types, lifting) Qp_constE a_def)
  then show ?thesis using assms 0 semialg_val_ineq_set_is_semialg[of g k f]
    by presburger
qed

lemma semialg_level_set_is_semialg:
  assumes "f  carrier (SA k)"
  assumes "c  carrier Qp"
  shows "is_semialgebraic k {x  carrier (Qpk). f x = c}"
proof-
  have 0: "is_univ_semialgebraic {c}"
    apply(rule finite_is_univ_semialgebraic) using assms apply blast by auto
  have 1: "{x  carrier (Qpk). f x = c} = f  ¯k{c}"
    unfolding evimage_def by auto
  then show ?thesis using 0 assms evimage_is_semialg by presburger
qed

lemma semialg_val_eq_set_is_semialg':
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (f x) = C}"
proof(cases "C = ")
  case True
  then have "{x  carrier (Qpk). val (f x) = C} = {x  carrier (Qpk). f x = 𝟬}"
    using assms unfolding val_def by (meson eint.distinct(1))
  then have "{x  carrier (Qpk). val (f x) = C} = f  ¯k{𝟬}"
    unfolding evimage_def by blast
  then show ?thesis
    using assms semialg_level_set_is_semialg[of f k 𝟬]  Qp.zero_closed {x  carrier (Qpk). val (f x) = C} = {x  carrier (Qpk). f x = 𝟬} by presburger
next
  case False
  then obtain N::int where N_def: "C = eint N"
    by blast
  have 0: "{x  carrier (Qpk). val (f x) = C} = {x  carrier (Qpk). val (f x)  N} - {x  carrier (Qpk). val (f x)  (eint (N-1))}"
  proof
    show "{x  carrier (Qpk). val (f x) = C}  {x  carrier (Qpk). val (f x)  N} - {x  carrier (Qpk). val (f x)  (eint (N-1))}"
    proof
      fix x assume A: "x  {x  carrier (Qpk). val (f x) = C}"
      then have 0: "x  carrier (Qpk)  val (f x) = C"
        by blast
      have 1: "¬ val (f x)  (eint (N-1))"
        using A N_def assms 0 eint_ord_simps(1) by presburger
      have 2: "val (f x)  (eint N)"
        using 0 N_def eint_ord_simps(1) by presburger
      show "x  {x  carrier (Qpk). val (f x)  N} - {x  carrier (Qpk). val (f x)  (eint (N-1))}"
        using 0 1 2
        by blast
    qed
    show  "{x  carrier (Qpk). val (f x)  N} - {x  carrier (Qpk). val (f x)  (eint (N-1))}   {x  carrier (Qpk). val (f x) = C}"
    proof fix x assume A: "x  {x  carrier (Qpk). val (f x)  N} - {x  carrier (Qpk). val (f x)  (eint (N-1))}"
      have 0:  "x  carrier (Qpk)  val (f x)  C"
        using A N_def by blast
      have 1: "¬ val (f x)  (eint (N-1))"
        using A 0  by blast
      have 2: "val (f x) = C"
      proof(rule ccontr)
        assume "val (f x)  C"
        then have "val (f x) < C"
          using 0 by auto
        then obtain M where M_def: "val (f x) = eint M"
          using N_def eint_iless by blast
        then have "M < N"
          by (metis N_def val (f x) < C eint_ord_simps(2))
        then have "val (f x)  eint (N - 1)"
          using M_def eint_ord_simps(1) by presburger
        then show False using 1 by blast
      qed
      show "x  {x  carrier (Qpk). val (f x) = C} "
        using 0 2  by blast
    qed
  qed
  have 1: "is_semialgebraic k {x  carrier (Qpk). val (f x)  N}"
    using assms semialg_val_ineq_set_is_semialg'[of f k N] by blast
  have 2: "is_semialgebraic k {x  carrier (Qpk). val (f x)  (eint (N-1))}"
    using assms semialg_val_ineq_set_is_semialg' by blast
  show ?thesis using 0 1 2
    using diff_is_semialgebraic by presburger
qed

lemma semialg_val_eq_set_is_semialg:
  assumes "g  carrier (SA k)"
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (g x) = val (f x)}"
proof-
  have 0: "is_semialgebraic k {x  carrier (Qpk). val (g x)  val (f x)}"
    using assms semialg_val_ineq_set_is_semialg[of g k f] by blast
  have 1: "is_semialgebraic k {x  carrier (Qpk). val (g x)  val (f x)}"
    using assms semialg_val_ineq_set_is_semialg[of f k g] by blast
  have 2: " {x  carrier (Qpk). val (g x) = val (f x)} =  {x  carrier (Qpk). val (g x)  val (f x)}  {x  carrier (Qpk). val (g x)  val (f x)}"
    using eq_iff by blast
  show ?thesis using 0 1 2 intersection_is_semialg by presburger
qed

lemma semialg_val_strict_ineq_set_formula:
"{x  carrier (Qpk). val (g x) < val (f x)} = {x  carrier (Qpk). val (g x)  val (f x)} - {x  carrier (Qpk). val (g x) = val (f x)}"
  using neq_iff le_less by blast

lemma semialg_val_ineq_set_complement:
"carrier (Qpk) - {x  carrier (Qpk). val (g x)  val (f x)} = {x  carrier (Qpk). val (f x) < val (g x)}"
  using not_le by blast

lemma semialg_val_strict_ineq_set_complement:
"carrier (Qpk) - {x  carrier (Qpk). val (g x) < val (f x)} = {x  carrier (Qpk). val (f x)  val (g x)}"
  using not_le by blast

lemma semialg_val_strict_ineq_set_is_semialg:
  assumes "g  carrier (SA k)"
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (g x) < val (f x)}"
  using semialg_val_ineq_set_complement[of k f g] assms diff_is_semialgebraic
        semialg_val_ineq_set_is_semialg[of f ]
  by (metis (no_types, lifting) complement_is_semialg)

lemma semialg_val_strict_ineq_set_is_semialg':
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (f x) < C}"
proof-
  obtain a where a_def: "a  carrier Qp  val a = C"
    by (meson Qp.carrier_not_empty Qp.minus_closed dist_nonempty' equals0I)
  then obtain g where g_def: "g = constant_function (carrier (Qpk)) a"
    by blast
  have 0: "g  carrier (SA k)"
    using g_def a_def SA_car assms(1) constant_function_in_semialg_functions by blast
  have 1: "{x  carrier (Qpk). val (f x) < C} = {x  carrier (Qpk). val (f x) < val (g x)}"
    using g_def by (metis (no_types, lifting) Qp_constE a_def)
  then show ?thesis using assms 0 semialg_val_strict_ineq_set_is_semialg[of f k g]
    by presburger
qed

lemma semialg_val_strict_ineq_set_is_semialg'':
  assumes "f  carrier (SA k)"
  shows "is_semialgebraic k {x  carrier (Qpk). val (f x) > C}"
proof-
  obtain a where a_def: "a  carrier Qp  val a = C"
    by (meson Qp.carrier_not_empty Qp.minus_closed dist_nonempty' equals0I)
  then obtain g where g_def: "g = constant_function (carrier (Qpk)) a"
    by blast
  have 0: "g  carrier (SA k)"
    using g_def a_def SA_car assms(1) constant_function_in_semialg_functions by blast
  have 1: "{x  carrier (Qpk). val (f x) > C} = {x  carrier (Qpk). val (f x) > val (g x)}"
    using g_def by (metis (no_types, lifting) Qp_constE a_def)
  then show ?thesis using assms 0 semialg_val_strict_ineq_set_is_semialg[of g k f]
    by presburger
qed

lemma semialg_val_ineq_set_plus:
  assumes "N > 0"
  assumes "c  carrier (SA N)"
  assumes "a  carrier (SA N)"
  shows "is_semialgebraic N  {x  carrier (QpN). val (c x)  val (a x) + eint n}"
proof-
  obtain b where b_def: "b = 𝔭[^]n SA Na"
    by blast
  have b_closed: "b  carrier (SA N)"
    unfolding b_def using assms SA_smult_closed p_intpow_closed(1) by blast
  have "x. x  carrier (QpN)  val (b x) = val (a x) + eint n"
    unfolding b_def by (metis Qp.function_ring_car_memE SA_car_memE(2) SA_smult_formula assms(3) p_intpow_closed(1) times_p_pow_val)
  hence 0: "{x  carrier (QpN). val (c x)  val (a x) + eint n} = {x  carrier (QpN). val (c x)  val (b x)}"
    by (metis (no_types, opaque_lifting) add.commute)
  show ?thesis unfolding 0 using assms b_def b_closed semialg_val_ineq_set_is_semialg[of c N b]  by blast
qed

lemma semialg_val_eq_set_plus:
  assumes "N > 0"
  assumes "c  carrier (SA N)"
  assumes "a  carrier (SA N)"
  shows "is_semialgebraic N  {x  carrier (QpN). val (c x) = val (a x) + eint n}"
proof-
  obtain b where b_def: "b = 𝔭[^]n SA Na"
    by blast
  have b_closed: "b  carrier (SA N)"
    unfolding b_def using assms SA_smult_closed p_intpow_closed(1) by blast
  have "x. x  carrier (QpN)  val (b x) = val (a x) + eint n"
    unfolding b_def by (metis Qp.function_ring_car_memE SA_car_memE(2) SA_smult_formula assms(3) p_intpow_closed(1) times_p_pow_val)
  hence 0: "{x  carrier (QpN). val (c x) = val (a x) + eint n} = {x  carrier (QpN). val (c x) = val (b x)}"
    by (metis (no_types, opaque_lifting) add.commute)
  show ?thesis unfolding 0 using assms b_def b_closed semialg_val_eq_set_is_semialg[of c N b]  by blast
qed

definition SA_zero_set where
"SA_zero_set n f = {x  carrier (Qpn). f x = 𝟬}"

lemma SA_zero_set_is_semialgebraic:
  assumes "f  carrier (SA n)"
  shows "is_semialgebraic n (SA_zero_set n f)"
  using assms semialg_level_set_is_semialg[of f n 𝟬] unfolding SA_zero_set_def
  by blast

lemma SA_zero_set_memE:
  assumes "f  carrier (SA n)"
  assumes "x  SA_zero_set n f"
  shows "f x = 𝟬"
  using assms unfolding SA_zero_set_def by blast

lemma SA_zero_set_memI:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x = 𝟬"
  shows  "x  SA_zero_set n f"
  using assms unfolding SA_zero_set_def by blast

lemma SA_zero_set_of_zero:
"SA_zero_set m (𝟬SA m) = carrier (Qpm)"
  apply(rule equalityI')
  unfolding SA_zero_set_def mem_Collect_eq
  apply blast
  using SA_zeroE by blast

definition SA_nonzero_set where
"SA_nonzero_set n f = {x  carrier (Qpn). f x  𝟬}"

lemma nonzero_evimage_closed:
  assumes "f  carrier (SA n)"
  shows "is_semialgebraic n {x  carrier (Qpn). f x  𝟬}"
proof-
  have "{x  carrier (Qpn). f x  𝟬} = f  ¯nnonzero Qp"
    unfolding nonzero_def evimage_def using SA_car_memE[of f n] assms by blast
  thus ?thesis   using assms evimage_is_semialg[of f n "nonzero Qp"] nonzero_is_univ_semialgebraic
    by presburger
qed

lemma SA_nonzero_set_is_semialgebraic:
  assumes "f  carrier (SA n)"
  shows "is_semialgebraic n (SA_nonzero_set n f)"
  using assms semialg_level_set_is_semialg[of f n 𝟬] unfolding SA_nonzero_set_def
  using nonzero_evimage_closed by blast

lemma SA_nonzero_set_memE:
  assumes "f  carrier (SA n)"
  assumes "x  SA_nonzero_set n f"
  shows "f x  𝟬"
  using assms unfolding SA_nonzero_set_def by blast

lemma SA_nonzero_set_memI:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x   𝟬"
  shows  "x  SA_nonzero_set n f"
  using assms unfolding SA_nonzero_set_def
  by blast

lemma SA_nonzero_set_of_zero:
"SA_nonzero_set m (𝟬SA m) = {}"
  apply(rule equalityI')
  unfolding SA_nonzero_set_def mem_Collect_eq
  using SA_zeroE apply blast
  by blast

lemma SA_car_memI':
  assumes "x. x  carrier (Qpm)  f x  carrier Qp"
  assumes "x. x  carrier (Qpm)  f x = undefined"
  assumes "k n P. n > 0  P  carrier (Qp [𝒳⇘1 + k⇙])  is_semialgebraic (m + k) (partial_pullback m f k (basic_semialg_set (1 + k) n P))"
  shows "f  carrier (SA m)"
  apply(rule SA_car_memI)
   apply(rule Qp.function_ring_car_memI)
  using assms(1) apply blast using assms(2) apply blast
  apply(rule is_semialg_functionI')
  using assms(1) apply blast
  using assms(3) unfolding is_basic_semialg_def
  by blast

lemma(in padic_fields) SA_zero_set_is_semialg:
  assumes "a  carrier (SA m)"
  shows "is_semialgebraic m {x  carrier (Qpm). a x = 𝟬}"
  using assms semialg_level_set_is_semialg[of a m 𝟬] Qp.zero_closed by blast

lemma(in padic_fields) SA_nonzero_set_is_semialg:
  assumes "a  carrier (SA m)"
  shows "is_semialgebraic m {x  carrier (Qpm). a x  𝟬}"
proof-
  have 0: "{x  carrier (Qpm). a x  𝟬} = carrier (Qpm) - {x  carrier (Qpm). a x = 𝟬}"
    by blast
  show ?thesis using assms  SA_zero_set_is_semialg[of a m] complement_is_semialg[of m "{x  carrier (Qpm). a x = 𝟬}"]
    unfolding 0  by blast
qed

lemma zero_set_nonzero_set_covers:
"carrier (Qpn) = SA_zero_set n f  SA_nonzero_set n f"
  unfolding SA_zero_set_def SA_nonzero_set_def
  apply(rule equalityI')
  unfolding mem_Collect_eq
  apply blast
  by blast

lemma zero_set_nonzero_set_covers':
  assumes "S  carrier (Qpn)"
  shows "S = (S  SA_zero_set n f)  (S  SA_nonzero_set n f)"
  using assms zero_set_nonzero_set_covers by blast

lemma zero_set_nonzero_set_covers_semialg_set:
  assumes "is_semialgebraic n S"
  shows "S = (S  SA_zero_set n f)  (S  SA_nonzero_set n f)"
  using assms is_semialgebraic_closed zero_set_nonzero_set_covers' by blast


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Partitioning Semialgebraic Sets According to Valuations of Functions›
(**************************************************************************************************)
(**************************************************************************************************)


text‹
  Given a semialgebraic set $A$ and a finite set of semialgebraic functions $Fs$, a common
  construction is to simplify one's understanding of the behaviour of the functions $\mathit{Fs}$ on
  $A$ by finitely paritioning $A$ into subsets where the element $f \in F$ for which $val (f x)$ is
  minimal is constant as $x$ ranges over each piece of the parititon. The function
  \texttt{Min\_set} helps construct this by picking out the subset of a set $A$ where the valuation
  of a particular element of $\mathit{Fs}$ is minimal. Such a set will always be semialgebraic.
›

lemma disjointify_semialg:
  assumes "finite As"
  assumes "As  semialg_sets n"
  shows "disjointify As  semialg_sets n"
  using assms unfolding semialg_sets_def
  by (simp add: disjointify_gen_boolean_algebra)

lemma semialgebraic_subalgebra:
  assumes "finite Xs"
  assumes "Xs  semialg_sets n"
  shows "atoms_of Xs  semialg_sets n"
  using assms unfolding semialg_sets_def
  by (simp add: atoms_of_gen_boolean_algebra)

lemma(in padic_fields) finite_intersection_is_semialg:
  assumes "finite  Xs"
  assumes "Xs  {}"
  assumes "F ` Xs  semialg_sets m"
  shows "is_semialgebraic m ( i  Xs. F i)"
proof-
  have 0: "F ` Xs  semialg_sets m  F ` Xs  {} "
    using assms by blast
  thus ?thesis
    using assms finite_intersection_is_semialgebraic[of "F ` Xs" m]
    by blast
qed


definition Min_set where
"Min_set m As a = carrier (Qpm)  ( f  As. {x  carrier (Qpm). val (a x)  val (f x) })"

lemma Min_set_memE:
  assumes "x  Min_set m As a"
  assumes "f  As"
  shows "val (a x)  val (f x)"
  using assms unfolding Min_set_def by blast

lemma Min_set_closed:
"Min_set m As a  carrier (Qpm)"
  unfolding Min_set_def by blast

lemma Min_set_semialg0:
  assumes "As  carrier (SA m)"
  assumes "finite As"
  assumes "a  As"
  assumes "As  {}"
  shows "is_semialgebraic m (Min_set m As a)"
  unfolding Min_set_def apply(rule intersection_is_semialg)
  using carrier_is_semialgebraic apply blast
  apply(rule finite_intersection_is_semialg)
  using assms apply blast
  using assms apply blast
proof(rule subsetI) fix x assume A: " x  (λi. {x  carrier (Qpm). val (a x)  val (i x)}) ` As"
  then obtain f where f_def: "f  As  x = {x  carrier (Qpm). val (a x)  val (f x)}"
    by blast
  have f_closed: "f  carrier (SA m)"
    using f_def assms by blast
  have a_closed: "a  carrier (SA m)"
    using assms by blast
  have "is_semialgebraic m  {x  carrier (Qpm). val (a x)  val (f x)}"
    using a_closed f_closed semialg_val_ineq_set_is_semialg by blast
  thus " x  semialg_sets m"
    using f_def unfolding is_semialgebraic_def by blast
qed

lemma Min_set_semialg:
  assumes "As  carrier (SA m)"
  assumes "finite As"
  assumes "a  As"
  shows "is_semialgebraic m (Min_set m As a)"
  apply(cases "As = {}")
  using Min_set_def  assms(3) apply blast
  using assms Min_set_semialg0 by blast

lemma Min_sets_cover:
  assumes "As  {}"
  assumes "finite As"
  shows "carrier (Qpm) = ( a  As. Min_set m As a)"
proof
  show "carrier (Qpm)   (Min_set m As ` As)"
  proof fix x assume A: "x  carrier (Qpm) "
    obtain v where v_def: "v = Min ((λf. val (f x)) ` As)"
      by blast
    obtain f where f_def: "f  As  v = val (f x)"
      unfolding v_def using assms Min_in[of "((λf. val (f x)) ` As)"]
      by blast
    have v_def': "v = val (f x)"
      using f_def by blast
    have 0: "x  Min_set m As f"
      unfolding Min_set_def
      apply(rule IntI)
      using A apply blast
    proof(rule InterI) fix s assume s: "s  (λfa. {x  carrier (Qpm). val (f x)  val (fa x)}) ` As"
      then obtain g where g_def: "g  As  s=  {x  carrier (Qpm). val (f x)  val (g x)}"
        by blast
      have s_def: "s=  {x  carrier (Qpm). val (f x)  val (g x)}"
        using g_def by blast
      have 00: " val (g x)  ((λf. val (f x)) ` As)"
        using g_def by blast
      show "x  s"
        unfolding s_def mem_Collect_eq using 00 A assms  MinE[of "((λf. val (f x)) ` As)" v "val (g x)"]
        unfolding v_def  by (metis f_def finite_imageI v_def)
    qed
    thus "x   (Min_set m As ` As)"
      using f_def by blast
  qed
  show " (Min_set m As ` As)  carrier (Qpm)"
    unfolding Min_set_def by blast
qed

text‹
  The sets defined by the function \texttt{Min\_set} for a fixed set of functions may not all be
   disjoint, but we can easily refine then to obtain a finite partition via the function
  "disjointify".
›

definition Min_set_partition where
"Min_set_partition m As B = disjointify ((∩)B  ` (Min_set m As ` As))"

lemma Min_set_partition_finite:
  assumes "finite As"
  shows "finite (Min_set_partition m As B)"
  unfolding Min_set_partition_def
  by (meson assms disjointify_finite finite_imageI)

lemma Min_set_partition_semialg0:
  assumes "finite As"
  assumes "As  carrier (SA m)"
  assumes "is_semialgebraic m B"
  assumes "S  ((∩)B  ` (Min_set m As ` As))"
  shows "is_semialgebraic m S"
  using Min_set_semialg[of As m] assms intersection_is_semialg[of m B]
  by blast

lemma Min_set_partition_semialg:
  assumes "finite As"
  assumes "As  carrier (SA m)"
  assumes "is_semialgebraic m B"
  assumes "S  (Min_set_partition m As B)"
  shows "is_semialgebraic m S"
proof-
  have 0: "(∩) B ` Min_set m As ` As  semialg_sets m "
    apply(rule subsetI)
    using Min_set_partition_semialg0[of As m B ] assms unfolding is_semialgebraic_def
    by blast
  thus ?thesis
  unfolding is_semialgebraic_def
  using assms  Min_set_partition_semialg0[of As m B] disjointify_semialg[of "((∩) B ` Min_set m As ` As)" m]
  unfolding Min_set_partition_def  is_semialgebraic_def by blast
qed

lemma Min_set_partition_covers0:
  assumes "finite As"
  assumes "As  {}"
  assumes "As  carrier (SA m)"
  assumes "is_semialgebraic m B"
  shows " ((∩)B  ` (Min_set m As ` As)) = B"
proof-
  have 0: " ((∩)B  ` (Min_set m As ` As)) = B   (Min_set m As ` As)"
    by blast
  have 1: "B  carrier (Qpm)"
    using assms is_semialgebraic_closed by blast
  show ?thesis unfolding 0 using 1 assms Min_sets_cover[of As m]  by blast
qed

lemma Min_set_partition_covers:
  assumes "finite As"
  assumes "As  carrier (SA m)"
  assumes "As  {}"
  assumes "is_semialgebraic m B"
  shows " (Min_set_partition m As B)  = B"
  unfolding Min_set_partition_def
  using Min_set_partition_covers0[of As m B] assms disjointify_union[of "((∩) B ` Min_set m As ` As)"]
  by (metis finite_imageI)

lemma Min_set_partition_disjoint:
  assumes "finite As"
  assumes "As  carrier (SA m)"
  assumes "As  {}"
  assumes "is_semialgebraic m B"
  assumes "s  Min_set_partition m As B"
  assumes "s'  Min_set_partition m As B"
  assumes "s   s'"
  shows  "s  s' = {}"
  apply(rule  disjointify_is_disjoint[of "((∩) B ` Min_set m As ` As)" s s'])
  using assms finite_imageI apply blast
  using assms unfolding Min_set_partition_def apply blast
  using assms unfolding Min_set_partition_def apply blast
  using assms by blast

lemma Min_set_partition_memE:
  assumes "finite As"
  assumes "As  carrier (SA m)"
  assumes "As  {}"
  assumes "is_semialgebraic m B"
  assumes "s  Min_set_partition m As B"
  shows "f  As. (x  s. (g  As. val (f x)  val (g x)))"
proof-
  obtain s' where s'_def: "s'  ((∩) B ` Min_set m As ` As)  s   s'"
    using finite_imageI assms disjointify_subset[of "((∩) B ` Min_set m As ` As)" s] unfolding Min_set_partition_def  by blast
  obtain f where f_def: "f  As  s' = B  Min_set m As f"
    using s'_def by blast
  have 0: "(x  s'. (g  As. val (f x)  val (g x)))"
    using f_def Min_set_memE[of _ m As f] by blast
  thus ?thesis
    using s'_def  by (meson f_def subset_iff)
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Valuative Congruence Sets for Semialgebraic Functions›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  The set of points $x$ where the values $\mathit{ord}\ f(x)$ satisfy a congruence are important
  basic examples of semialgebraic sets, and will be vital in the proof of Macintyre's Theorem. The
  lemma below is essentially the content of Denef's Lemma 2.1.3 from his cell decomposition paper.
›

lemma pre_SA_unit_cong_set_is_semialg:
  assumes "k  0"
  assumes "f  Units (SA n)"
  shows "is_semialgebraic n {x  carrier (Qpn). ord (f x) mod k  = a }"
proof-
  have 0: "{x  carrier (Qpn). ord (f x) mod k  = a } = f  ¯nord_congruence_set k a"
    unfolding ord_congruence_set_def
    apply(rule equalityI')
    using assms unfolding evimage_def vimage_def mem_Collect_eq
    apply (metis (mono_tags, lifting) IntI Qp.function_ring_car_memE SA_Units_closed SA_Units_memE' SA_car_memE(2) mem_Collect_eq not_nonzero_Qp)
    using assms by blast
  show ?thesis unfolding 0
    apply(rule evimage_is_semialg)
    using assms  apply blast
    using  assms ord_congruence_set_univ_semialg[of k a]
    by blast
qed

lemma SA_unit_cong_set_is_semialg:
  assumes "f  Units (SA n)"
  shows "is_semialgebraic n {x  carrier (Qpn). ord (f x) mod k = a}"
proof(cases "k  0")
  case True
  then show ?thesis
    using assms pre_SA_unit_cong_set_is_semialg by presburger
next
  case False
  show ?thesis
  proof(cases "a = 0")
    case True
    have T0: "{x  carrier (Qpn). ord (f x) mod k  = a } = {x  carrier (Qpn). ord (f x) mod (-k)  = a }"
      apply(rule equalityI')
      unfolding mem_Collect_eq using True zmod_zminus2_not_zero  apply meson
      using True zmod_zminus2_not_zero
      by (metis equation_minus_iff)
    show ?thesis unfolding T0 apply(rule pre_SA_unit_cong_set_is_semialg[of "-k" f n a])
      using False apply presburger using assms by blast
  next
    case F: False
    show ?thesis
    proof(cases "a = k")
      case True
      have 0: "{x  carrier (Qpn). ord (f x) mod k = a} = {x  carrier (Qpn). ord (f x) mod k = k}"
        using True by blast
      have 1: "{x  carrier (Qpn). ord (f x) mod k = a} = {}  k = 0"
      proof(cases "{x  carrier (Qpn). ord (f x) mod k = a}  {}")
        case T: True
        then obtain x where "ord (f x) mod k = k"
          unfolding True by blast
        then have "k = 0"
          by (metis mod_mod_trivial mod_self)
        thus ?thesis by blast
      next
        case False
        then show ?thesis by blast
      qed
      show ?thesis apply(cases "{x  carrier (Qpn). ord (f x) mod k = a} = {}")
        using empty_is_semialgebraic apply presburger
        using 1 pre_SA_unit_cong_set_is_semialg assms by blast
    next
      case F': False
      have 0: "{x  carrier (Qpn). ord (f x) mod k = a} = {x  carrier (Qpn). ord (f x) mod (-k) = a - k}"
        apply(rule equalityI')
        unfolding mem_Collect_eq using zmod_zminus2_eq_if assms apply (metis F)
        unfolding mem_Collect_eq zmod_zminus2_eq_if using False F F' assms
        by (metis (no_types, opaque_lifting) cancel_ab_semigroup_add_class.diff_right_commute group_add_class.right_minus_eq)
      show ?thesis unfolding 0 apply(rule pre_SA_unit_cong_set_is_semialg)
        using False apply presburger using assms by blast
    qed
  qed
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Gluing Functions Along Semialgebraic Sets›
(**************************************************************************************************)
(**************************************************************************************************)
text‹
  Semialgebraic functions have the useful property that they are closed under piecewise definitions.
  That is, if $f, g$ are semialgebraic and $C \subseteq \mathbb{Q}_p^m$ is a semialgebraic set,
  then the function:
  \[
   h(x) =
    \begin{cases}
                                     f(x) & \text{if $x \in C$} \\
                                     g(x) & \text{if $x \in \mathbb{Q}_p^m - C$} \\
    undefined & \text{otherwise}
    \end{cases}
  \]
  is again semialgebraic. The function $h$ can be obtained by the definition
  \[\texttt{h = fun\_glue m C f g}\] which is defined below. This closure property means that we
  can avoid having to define partial semialgebraic functions which are undefined outside of some
  proper subset of $\mathbb{Q}_p^m$, since it usually suffices to just define the function as some
  arbitrary constant outside of the desired domain. This is useful for defining partial
  multiplicative inverses of arbitrary functions. If $f$ is semialgebraic, then its nonzero set
  $\{x \in \mathbb{Q}_p^m \mid f x \neq 0\}$ is semialgebraic. By gluing $f$ to the constant
  function $1$ outside of its nonzero set, we obtain an invertible element in the ring
  \texttt{SA(m)} which evaluates to a multiplicative inverse of $f(x)$ on the largest domain
  possible.
›
(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Defining Piecewise Semialgebraic Functions›
(**************************************************************************************************)
(**************************************************************************************************)


text‹
  An important property that will be repeatedly used is that we can define piecewise semialgebraic
  functions, which will themselves be semialgebraic as long as the pieces are semialgebraic sets.
  An important application of this principle will be that a function $f$ which is always nonzero
  on some semialgebraic set $A$ can be replaced with a global unit in the ring of semialgebraic
  functions. This global unit admits a global multiplicative inverse that inverts $f$ pointwise on
  $A$, and allows us to avoid having to consider localizations of function rings to locally invert
  such functions.
›

definition fun_glue where
"fun_glue n S f g = (λx  carrier (Qpn). if x  S then f x else g x)"

lemma fun_glueE:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "S  carrier (Qpn)"
  assumes "x  S"
  shows "fun_glue n S f g x = f x"
proof-
  have "x  carrier (Qpn)"
    using assms by blast
  thus ?thesis
    unfolding fun_glue_def using assms
    by (metis (mono_tags, lifting) restrict_apply')
qed

lemma fun_glueE':
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "S  carrier (Qpn)"
  assumes "x  carrier (Qpn) - S"
  shows "fun_glue n S f g x = g x"
proof-
  have 0: "x  carrier (Qpn)"
    using assms by blast
  have 1: "x  S"
    using assms by blast
  show ?thesis
    unfolding fun_glue_def using assms 0 1
    by (metis (mono_tags, lifting) restrict_apply')
qed

lemma fun_glue_evimage:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "S  carrier (Qpn)"
  shows "fun_glue n S f g ¯nT = ((f  ¯nT)  S)  ((g  ¯nT) - S)"
proof
  show "fun_glue n S f g ¯nT  ((f  ¯nT)  S)  ((g  ¯nT) - S)"
  proof fix x assume A: "x  fun_glue n S f g ¯nT "
    then have 0: "fun_glue n S f g x  T"
      by blast
    have 1: "x  carrier (Qpn)"
      using A  by (meson evimage_eq)
    show "x  ((f  ¯nT)  S)  ((g  ¯nT) - S)"
      apply(cases "x  S")
      apply auto[1]
      using "1" apply force
      using "0" assms(1) assms(2) assms(3) fun_glueE apply force
      apply auto[1] using 1 apply blast
      using A 1  unfolding fun_glue_def evimage_def Int_iff  by auto
  qed
  show " f ¯nT  S  (g ¯nT - S)  fun_glue n S f g ¯nT"
  proof fix x assume A: "x  f ¯nT  S  (g ¯nT - S)"
    then have x_closed: "x  carrier (Qpn)"
      by (metis (no_types, opaque_lifting) Diff_iff Int_iff UnE extensional_vimage_closed subsetD)
    show "x  fun_glue n S f g ¯nT"
      apply(cases "x  S")
      using x_closed fun_glueE assms
       apply (metis A DiffD2 IntD1 UnE evimage_eq)
      using x_closed fun_glueE' assms
      by (metis A Diff_iff Int_iff Un_iff evimageD evimageI2)
  qed
qed

lemma fun_glue_partial_pullback:
  assumes "f  carrier (SA k)"
  assumes "g  carrier (SA k)"
  assumes "S  carrier (Qpk)"
  shows "partial_pullback k (fun_glue k S f g) n T =
          ((cartesian_product S (carrier (Qpn)))  partial_pullback k f n T)  ((partial_pullback k g n T)- (cartesian_product S (carrier (Qpn))))"
proof
  show "partial_pullback k (fun_glue k S f g) n T  (cartesian_product S (carrier (Qpn)))  partial_pullback k f n T  (partial_pullback k g n T - (cartesian_product S (carrier (Qpn))))"
  proof fix x assume A: "x  partial_pullback k (fun_glue k S f g) n T "
    then have x_closed: "x  carrier (Qpk+n)" unfolding partial_pullback_def partial_image_def
      by (meson evimage_eq)
    show " x  cartesian_product S (carrier (Qpn))  partial_pullback k f n T  (partial_pullback k g n T - (cartesian_product S (carrier (Qpn))))"
    proof(cases "x  cartesian_product S (carrier (Qpn))")
      case True
      then have T0: "take k x  S"
        using assms cartesian_product_memE(1) by blast
      then have "(fun_glue k S f g) (take k x) = f (take k x)"
        using assms fun_glueE[of f k g S "take k x"]
        by blast
      then have "partial_image k (fun_glue k S f g) x = partial_image k f x"
        using A x_closed unfolding partial_pullback_def partial_image_def
        by blast
      then show ?thesis using T0 A unfolding partial_pullback_def evimage_def
        by (metis IntI Int_iff True Un_iff vimageI vimage_eq x_closed)
    next
      case False
      then have F0: "take k x  carrier (Qpk) - S"
        using A x_closed assms cartesian_product_memI
        by (metis (no_types, lifting) DiffI carrier_is_semialgebraic cartesian_power_drop is_semialgebraic_closed le_add1 local.take_closed)
      then have "(fun_glue k S f g) (take k x) = g (take k x)"
        using assms fun_glueE'[of f k g S "take k x"]
        by blast
      then have "partial_image k (fun_glue k S f g) x = partial_image k g x"
        using A x_closed unfolding partial_pullback_def partial_image_def
        by blast
      then have "x  partial_pullback k g n T "
        using F0 x_closed A unfolding partial_pullback_def partial_image_def evimage_def
        by (metis (no_types, lifting) A IntI local.partial_image_def partial_pullback_memE(2) vimageI)
      then have "x  (partial_pullback k g n T - cartesian_product S (carrier (Qpn)))"
        using False by blast
      then show ?thesis by blast
    qed
  qed
  show "cartesian_product S (carrier (Qpn))  partial_pullback k f n T  (partial_pullback k g n T - cartesian_product S (carrier (Qpn)))
     partial_pullback k (fun_glue k S f g) n T"
  proof fix x assume A: "x  cartesian_product S (carrier (Qpn))  partial_pullback k f n T  (partial_pullback k g n T - cartesian_product S (carrier (Qpn)))"
    then have x_closed: "x  carrier (Qpn+k)"
      by (metis DiffD1 Int_iff Un_iff add.commute partial_pullback_memE(1))
    show "x  partial_pullback k (fun_glue k S f g) n T"
    proof(cases "x  cartesian_product S (carrier (Qpn))  partial_pullback k f n T")
      case True
      show ?thesis apply(rule partial_pullback_memI)
        using x_closed apply (metis add.commute)
        using x_closed True assms fun_glueE[of f k g S "take k x"] partial_pullback_memE[of x k f n T]
        unfolding partial_image_def by (metis Int_iff cartesian_product_memE(1))
    next
      case False
      show ?thesis apply(rule partial_pullback_memI)
        using x_closed apply (metis add.commute)
        using A x_closed False assms fun_glueE'[of f k g S "take k x"] partial_pullback_memE[of x k g n T]
        unfolding partial_image_def
        by (metis (no_types, lifting) Diff_iff Un_iff carrier_is_semialgebraic cartesian_power_drop cartesian_product_memI is_semialgebraic_closed le_add2 local.take_closed)
    qed
  qed
qed

lemma fun_glue_eval_closed:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "is_semialgebraic n S"
  assumes "x  carrier (Qpn)"
  shows "fun_glue n S f g x  carrier Qp"
  apply(cases "x  S")
  using assms fun_glueE SA_car_memE
   apply (metis Qp.function_ring_car_mem_closed is_semialgebraic_closed)
proof- assume A: "x  S"
  then have 0: "x  carrier (Qpn) - S"
    using assms by auto
  hence 1: "fun_glue n S f g x = g x"
    using assms fun_glueE' is_semialgebraic_closed by auto
  show "fun_glue n S f g x  carrier Qp"
    unfolding 1 using assms SA_car_memE by blast
qed

lemma fun_glue_closed:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "fun_glue n S f g  carrier (SA n)"
  apply(rule SA_car_memI)
   apply(rule Qp.function_ring_car_memI)
  using fun_glue_eval_closed assms apply blast
  using fun_glue_def unfolding restrict_apply apply metis
  apply(rule is_semialg_functionI, intro Pi_I fun_glue_eval_closed assms, blast)
proof-
  fix k T assume A: "T  semialg_sets (1 + k)"
  have 0: "is_semialgebraic (n+k) (partial_pullback n f k T)"
    using assms A SA_car_memE[of f n]  is_semialg_functionE[of n f k T]  padic_fields.is_semialgebraicI padic_fields_axioms by blast
  have 1: "is_semialgebraic (n+k) (partial_pullback n g k T)"
    using assms A SA_car_memE[of g n]  is_semialg_functionE[of n g k T] padic_fields.is_semialgebraicI padic_fields_axioms by blast
  have 2: "partial_pullback n (fun_glue n S f g) k T =
    cartesian_product S (carrier (Qpk))  partial_pullback n f k T  (partial_pullback n g k T - cartesian_product S (carrier (Qpk)))"
    using assms fun_glue_partial_pullback[of f n g S k T] f  carrier (SA n) g  carrier (SA n) is_semialgebraic_closed
    by blast
  show "is_semialgebraic (n + k) (partial_pullback n (fun_glue n S f g) k T)"
    using assms 0 1 2 cartesian_product_is_semialgebraic carrier_is_semialgebraic
      diff_is_semialgebraic intersection_is_semialg union_is_semialgebraic by presburger
qed

lemma fun_glue_unit:
  assumes "f  carrier (SA n)"
  assumes "is_semialgebraic n S"
  assumes "x. x  S  f x  𝟬"
  shows "fun_glue n S f 𝟭SA n Units (SA n)"
proof(rule SA_Units_memI)
  show "fun_glue n S f 𝟭SA n carrier (SA n)"
    using fun_glue_closed assms SA_is_cring cring.cring_simprules(6) by blast
  show "x. x  carrier (Qpn)  fun_glue n S f 𝟭SA nx  𝟬"
  proof- fix x assume A: "x  carrier (Qpn)"
    show "fun_glue n S f 𝟭SA nx  𝟬"
      apply(cases "x  S")
      using assms SA_is_cring cring.cring_simprules(6) assms(3)[of x] fun_glueE[of f n "𝟭SA n⇙" S x]
      apply (metis is_semialgebraic_closed)
      using assms SA_is_cring[of n] cring.cring_simprules(6)[of "SA n"]
            A  fun_glueE'[of f n "𝟭SA n⇙" S x] is_semialgebraic_closed[of n S]
      unfolding SA_one[of n] function_ring_defs(4)[of n] function_one_def
      by (metis Diff_iff function_one_eval Qp_funs_one local.one_neq_zero)
  qed
qed

definition parametric_fun_glue where
"parametric_fun_glue n Xs fs = (λx  carrier (Qpn). let S = (THE S. S  Xs  x  S) in (fs S x))"

lemma parametric_fun_glue_formula:
  assumes "Xs partitions (carrier (Qpn))"
  assumes "x  S"
  assumes "S  Xs"
  shows "parametric_fun_glue n Xs fs x = fs S x"
proof-
  have 0: "(THE S. S  Xs  x  S) = S"
    apply(rule the_equality)
    using assms apply blast
    using assms unfolding is_partition_def by (metis Int_iff empty_iff disjointE)
  have 1: "x  carrier (Qpn)"
    using assms unfolding is_partition_def by blast
  then show ?thesis using 0 unfolding parametric_fun_glue_def  restrict_def by metis
qed

definition char_fun where
"char_fun n S = (λ x  carrier (Qpn). if x  S then 𝟭 else 𝟬)"

lemma char_fun_is_semialg:
  assumes "is_semialgebraic n S"
  shows "char_fun n S  carrier (SA n)"
proof-
  have "char_fun n S = fun_glue n S 𝟭SA n𝟬SA n⇙"
    unfolding char_fun_def fun_glue_def
    by (metis (no_types, lifting) function_one_eval function_zero_eval SA_one SA_zero restrict_ext)
  then show ?thesis
    using assms fun_glue_closed
    by (metis SA_is_cring cring.cring_simprules(2) cring.cring_simprules(6))
qed

lemma SA_finsum_apply:
  assumes "finite S"
  assumes "x  carrier (Qpn)"
  shows "F  S  carrier (SA n)  finsum (SA n) F S x = (sS. F s x)"
proof(rule finite.induct[of S])
  show "finite S"
    using assms by blast
  show " F  {}  carrier (SA n)  finsum (SA n) F {} x = (s{}. F s x)"
    using assms abelian_monoid.finsum_empty[of "SA n"]   Qp.abelian_monoid_axioms SA_is_abelian_monoid
    by (simp add: SA_zeroE)
  show "A a. finite A 
           F  A  carrier (SA n)  finsum (SA n) F A x = (sA. F s x) 
           F  insert a A  carrier (SA n)  finsum (SA n) F (insert a A) x = (sinsert a A. F s x)"
  proof- fix A a assume A: "finite A" "F  A  carrier (SA n)  finsum (SA n) F A x = (sA. F s x)"
    show " F  insert a A  carrier (SA n)  finsum (SA n) F (insert a A) x = (sinsert a A. F s x)"
    proof assume A': "F  insert a A  carrier (SA n)"
      then have 0: "F  A  carrier (SA n)"
        by blast
      hence 1: "finsum (SA n) F A x = (sA. F s x)"
        using A by blast
      show "finsum (SA n) F (insert a A) x = (sinsert a A. F s x)"
      proof(cases "a  A")
        case True
        then show ?thesis
          using 1 by (metis insert_absorb)
      next
        case False
        have F00: "(λs. F s x)  A  carrier Qp"
          apply(rule Pi_I, rule SA_car_closed[of _ n] )
          using "0" assms by auto
        have F01: "F a x  carrier Qp"
          using A' assms
          by (metis (no_types, lifting) Qp.function_ring_car_mem_closed Pi_split_insert_domain SA_car_in_Qp_funs_car subsetD)
        have F0: "(sinsert a A. F s x) = F a x  (sA. F s x)"
          using F00 F01  A' False A(1) Qp.finsum_insert[of A a "λs. F s x"] by blast
        have F1: "finsum (SA n) F (insert a A) = F a SA nfinsum (SA n) F A"
          using abelian_monoid.finsum_insert[of "SA n" A a F]
          by (metis (no_types, lifting) A' A(1) False Pi_split_insert_domain SA_is_abelian_monoid assms(1))
        show ?thesis
          using Qp.finsum_closed[of "λs. F s x" A] abelian_monoid.finsum_closed[of "SA n" F A]
                SA_is_abelian_monoid[of n] assms F0 F1 "0" A(2) SA_add by presburger
      qed
    qed
  qed
qed

lemma SA_finsum_apply_zero:
  assumes "finite S"
  assumes "F  S  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "s. s  S  F s x = 𝟬"
  shows "finsum (SA n) F S x = 𝟬"
proof-
  have "finsum (SA n) F S x = (sS. F s x)"
    using SA_finsum_apply assms by blast
  then show ?thesis using assms
    by (metis Qp.add.finprod_one_eqI)
qed

lemma parametric_fun_glue_is_SA:
  assumes "finite Xs"
  assumes "Xs partitions (carrier (Qpn))"
  assumes "fs  Xs  carrier (SA n)"
  assumes " S  Xs. is_semialgebraic n S"
  shows "parametric_fun_glue n Xs fs  carrier (SA n)"
proof-
  obtain F where F_def: "F = (λS. fs S SA nchar_fun n S)"
    by blast
  have 0: "F  Xs  carrier (SA n)" proof fix S assume "S  Xs" then show "F S  carrier (SA n)"
      using  SA_mult_closed[of n "fs S" "char_fun n S"] char_fun_is_semialg[of n S] assms SA_car_memE
      unfolding F_def by blast qed
  have 1: "S x. S  Xs  x  S  F S x = fs S x"
  proof- fix S x assume A: "S  Xs" "x  S"
    then have x_closed: "x  carrier (Qpn)"
      using assms unfolding is_partition_def by blast
    then have 0: "F S x = fs S x  char_fun n S x"
      unfolding F_def using SA_mult by blast
    have 1: "char_fun n S x = 𝟭"
      using char_fun_def A x_closed by auto
    have 2: "fs S x  carrier Qp"
      apply(intro SA_car_closed[of _ n] x_closed )
      using assms A by auto
    show "F S x = fs S x"
      unfolding 0  1 using 2 Qp.cring_simprules(12) by auto
  qed
  have 2: "S x. S  Xs  x  carrier (Qpn)  x  S  F S x = 𝟬"
  proof- fix S x assume A: "S  Xs" "x  carrier (Qpn)" "x  S"
    then have x_closed: "x  carrier (Qpn)"
      using assms unfolding is_partition_def by blast
    hence 20: "F S x = fs S x  char_fun n S x"
      unfolding F_def using SA_mult by blast
    have 21: "char_fun n S x = 𝟬"
      unfolding char_fun_def using A x_closed by auto
    have 22: "fs S x  carrier Qp"
      apply(intro SA_car_closed[of _ n] x_closed )
      using assms A by auto
    show "F S x = 𝟬"
      using 22 unfolding 20 21 by auto
  qed
  obtain g where g_def: "g = finsum (SA n) F Xs"
    by blast
  have g_closed: "g  carrier (SA n)"
    using abelian_monoid.finsum_closed[of "SA n" F Xs] assms SA_is_ring 0
    unfolding g_def ring_def abelian_group_def by blast
  have "g = parametric_fun_glue n Xs fs"
  proof fix x show "g x = parametric_fun_glue n Xs fs x"
    proof(cases "x  carrier (Qpn)")
      case True
      then obtain S where S_def: "S  Xs  x  S"
        using assms is_partitionE by blast
      then have T0: "parametric_fun_glue n Xs fs x = F S x"
        using 1 assms parametric_fun_glue_formula by blast
      have T1: "g x = F S x"
      proof-
        have 00: " F S SA nfinsum (SA n) F (Xs - {S}) =  finsum (SA n) F (insert S (Xs - {S}))"
          using abelian_monoid.finsum_insert[of "SA n" "Xs - {S}" S F ]
          by (metis (no_types, lifting) "0" Diff_iff Pi_anti_mono Pi_split_insert_domain SA_is_abelian_monoid S_def Set.basic_monos(7) assms(1) finite_Diff insert_iff subsetI)
        have T10: "g  = F S SA nfinsum (SA n) F (Xs - {S})"
          using S_def unfolding 00 g_def
          by (simp add: insert_absorb)
        have T11: "finsum (SA n) F (Xs - {S})  carrier (SA n)"
          using abelian_monoid.finsum_closed[of "SA n" F "Xs - {S}"] assms SA_is_ring 0
          unfolding g_def ring_def abelian_group_def by blast
        hence T12: "g x = F S x  finsum (SA n) F (Xs - {S}) x"
          using SA_add S_def T10 assms is_semialgebraic_closed by blast
        have T13: "finsum (SA n) F (Xs - {S}) x = 𝟬"
          apply(rule SA_finsum_apply_zero[of "Xs - {S}" F n x])
              using assms apply blast
                using "0" apply blast
                using True apply blast
        proof-
           fix s assume AA: "s  Xs - {S}"
           then have "x  s"
             using True assms S_def is_partitionE[of Xs "carrier (Qpn)"]  disjointE[of Xs S s]
             by blast
           then show "F s x = 𝟬"
             using AA 2[of s x] True by blast
        qed
        have T14: "F S x  carrier Qp"
          using assms True S_def  by (metis (no_types, lifting) "0" Qp.function_ring_car_mem_closed PiE SA_car_memE(2))
        then show ?thesis using T12 T13 assms True Qp.add.l_cancel_one Qp.zero_closed by presburger
      qed
      show ?thesis using T0 T1 by blast
    next
      case False
      then show ?thesis
        using g_closed unfolding parametric_fun_glue_def
        by (metis (mono_tags, lifting) function_ring_not_car SA_car_memE(2) restrict_def)
    qed
  qed
  then show ?thesis using g_closed by blast
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Turning Functions into Units Via Gluing›
(**************************************************************************************************)
(**************************************************************************************************)

text‹
  By gluing a function to the multiplicative unit on its zero set, we can get a canonical choice of
  local multiplicative inverse of a function $f$. Denef's proof frequently reasons about functions
  of the form $\frac{f(x)}{g(x)}$ with the tacit understanding that they are meant to be defined
  on the largest domain of definition possible. This technical tool allows us to replicate this
  kind of reasoning in our formal proofs.
›

definition to_fun_unit where
"to_fun_unit n f = fun_glue n {x  carrier (Qpn). f x  𝟬} f 𝟭SA n⇙"

lemma to_fun_unit_is_unit:
  assumes "f  carrier (SA n)"
  shows "to_fun_unit n f  Units (SA n)"
  unfolding to_fun_unit_def
  apply(rule fun_glue_unit)
  apply (simp add: assms)
  using assms nonzero_evimage_closed[of f] apply blast
  by blast

lemma to_fun_unit_closed:
  assumes "f  carrier (SA n)"
  shows "to_fun_unit n f  carrier (SA n)"
  using assms to_fun_unit_is_unit SA_is_ring SA_Units_closed by blast

lemma to_fun_unit_eq:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  shows "to_fun_unit n f x = f x"
  unfolding to_fun_unit_def fun_glue_def using assms
  by simp

lemma to_fun_unit_eq':
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x = 𝟬"
  shows "to_fun_unit n f x = 𝟭"
  unfolding to_fun_unit_def fun_glue_def using assms
  by (simp add: SA_oneE)

definition one_over_fun where
"one_over_fun n f = invSA n(to_fun_unit n f)"

lemma one_over_fun_closed:
  assumes "f  carrier (SA n)"
  shows "one_over_fun n f  carrier (SA n)"
  using assms SA_is_ring[of n] to_fun_unit_is_unit[of f n]
  by (metis SA_Units_closed one_over_fun_def ring.Units_inverse)

lemma one_over_fun_eq:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  shows "one_over_fun n f x = inv (f x)"
  using assms to_fun_unit_eq unfolding one_over_fun_def
  using Qp_funs_m_inv SA_Units_Qp_funs_Units SA_Units_Qp_funs_inv to_fun_unit_is_unit by presburger

lemma one_over_fun_smult_eval:
  assumes "f  carrier (SA n)"
  assumes "a  𝟬"
  assumes "a  carrier Qp"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  shows "one_over_fun n (a SA nf) x = inv (a  (f x))"
  using one_over_fun_eq[of "a SA nf" n x] assms
  by (metis Qp.function_ring_car_memE Qp.integral SA_car_memE(2) SA_smult_closed SA_smult_formula)

lemma one_over_fun_smult_eval':
  assumes "f  carrier (SA n)"
  assumes "a  𝟬"
  assumes "a  carrier Qp"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  shows "one_over_fun n (a SA nf) x = inv a  inv (f x)"
proof-
  have 0: "one_over_fun n (a SA nf) x = inv (a  f x)"
    using assms one_over_fun_smult_eval[of f n a x]
    by fastforce
  have 1: "f x  nonzero Qp"
    by(intro nonzero_memI SA_car_closed[ of _ n] assms)
  show ?thesis
    unfolding 0 using 1 assms
    using Qp.comm_inv_char Qp.cring_simprules(11) Qp.cring_simprules(5) SA_car_closed field_inv(2) field_inv(3) local.fract_cancel_right by presburger
qed



lemma SA_add_pow_closed:
  assumes "f  carrier (SA n)"
  shows "([(k::nat)]SA nf)  carrier (SA n)"
  using assms SA_is_ring[of n]
  by (meson ring.nat_mult_closed)

lemma one_over_fun_add_pow_eval:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  assumes "(k::nat) > 0"
  shows "one_over_fun n ([k]SA nf) x = inv ([k] f x)"
proof-
  have 0: "([k] SA nf) x = [k]  f x"
    using assms SA_add_pow_apply[of f n x k] by linarith
  hence 1: "([k] SA nf) x  𝟬"
    using assms Qp_char_0'' Qp.function_ring_car_mem_closed SA_car_memE(2)
    by metis
  have 2: "one_over_fun n ([k] SA nf) x = inv ([k] SA nf) x"
    using assms one_over_fun_eq[of "[k]SA nf" n x] 1 SA_add_pow_closed by blast
  thus ?thesis using 1 0 by presburger
qed

lemma one_over_fun_pow_closed:
  assumes "f  carrier (SA n)"
  shows "one_over_fun n (f[^]SA n(k::nat))  carrier (SA n)"
  using assms
  by (meson SA_nat_pow_closed one_over_fun_closed padic_fields.SA_imp_semialg padic_fields_axioms)

lemma one_over_fun_pow_eval:
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "f x  𝟬"
  shows "one_over_fun n (f[^]SA n(k::nat)) x = inv ((f x) [^] k)"
  using one_over_fun_eq[of  "f[^]SA nk" n x] assms
  by (metis Qp.function_ring_car_memE Qp.nonzero_pow_nonzero SA_car_memE(2) SA_nat_pow SA_nat_pow_closed padic_fields.SA_car_memE(1) padic_fields_axioms)


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Inclusions of Lower Dimensional Function Rings›
(**************************************************************************************************)
(**************************************************************************************************)


definition fun_inc where
"fun_inc m n f = (λ x  carrier (Qpm). f (take n x))"

lemma fun_inc_closed:
  assumes "f  carrier (SA n)"
  assumes "m  n"
  shows  "fun_inc m n f  carrier (SA m)"
proof-
  have 0: "x. x  carrier (Qpm)  fun_inc m n f x = (f  take n) x"
    unfolding fun_inc_def by (metis comp_apply restrict_apply')
  have 1: "is_semialg_function m (f  take n)"
    using assms comp_take_is_semialg
    by (metis SA_imp_semialg le_neq_implies_less padic_fields.semialg_function_comp_closed padic_fields_axioms take_is_semialg_map)
  have 2: "is_semialg_function m (fun_inc m n f)"
    using 0 1 semialg_function_on_carrier' by blast
  show ?thesis apply(rule SA_car_memI) apply(rule  Qp.function_ring_car_memI)
    using "2" is_semialg_function_closed apply blast
    using fun_inc_def[of m n f] unfolding restrict_def apply presburger
    using 2 by blast
qed

lemma fun_inc_eval:
  assumes "x  carrier (Qpm)"
  shows "fun_inc m n f x = f (take n x)"
  unfolding fun_inc_def using assms
  by (meson restrict_apply')

lemma ord_congruence_set_univ_semialg_fixed:
  assumes "n  0"
  shows "is_univ_semialgebraic (ord_congruence_set n a)"
  using ord_congruence_set_univ_semialg assms
  by auto

lemma ord_congruence_set_SA_function:
  assumes "n  0"
  assumes "c  carrier (SA (m+l))"
  shows "is_semialgebraic (m+l) {x  carrier (Qpm+l). c x  nonzero Qp  ord (c x) mod n = a}"
proof-
  have 0: "{x  carrier (Qpm+l). c x  nonzero Qp  ord (c x) mod n = a} = c  ¯m+l(ord_congruence_set n a)"
    unfolding ord_congruence_set_def evimage_def using assms by blast
  show ?thesis unfolding 0 apply(rule evimage_is_semialg)
    using assms apply blast using assms  ord_congruence_set_univ_semialg_fixed[of n a]
    by blast
qed

lemma ac_cong_set_SA:
  assumes "n > 0"
  assumes "k  Units (Zp_res_ring n)"
  assumes "c  carrier (SA (m+l))"
  shows "is_semialgebraic (m+l) {x  carrier (Qpm+l). c x  nonzero Qp  ac n (c x) = k}"
proof-
  have  "{x  carrier (Qpm+l). c x  nonzero Qp  ac n (c x) = k}= (c ¯m + lac_cong_set n k)"
    unfolding ac_cong_set_def evimage_def nonzero_def mem_Collect_eq using assms by blast
  thus ?thesis
    using assms ac_cong_set_is_univ_semialg[of n k] evimage_is_semialg[of c "m+l" "ac_cong_set n k"]
    by presburger
qed

lemma ac_cong_set_SA':
  assumes "n >0 "
  assumes "k  Units (Zp_res_ring n)"
  assumes "c  carrier (SA m)"
  shows "is_semialgebraic m {x  carrier (Qpm). c x  nonzero Qp  ac n (c x) = k}"
  using assms ac_cong_set_SA[of n k c m 0] unfolding Nat.add_0_right by blast

lemma ac_cong_set_SA'':
  assumes "n >0 "
  assumes "m > 0"
  assumes "k  Units (Zp_res_ring n)"
  assumes "c  carrier (SA m)"
  assumes "x. x  carrier (Qpm)  c x  𝟬"
  shows "is_semialgebraic m {x  carrier (Qpm). ac n (c x) = k}"
proof-
  have "{x  carrier (Qpm). c x  nonzero Qp  ac n (c x) = k} = {x  carrier (Qpm). ac n (c x) = k}"
    apply(rule subset_antisym) apply blast
    apply(rule subsetI) using assms unfolding nonzero_def mem_Collect_eq
    using Qp.function_ring_car_memE SA_car_memE(2) by blast
  thus ?thesis using assms ac_cong_set_SA'[of n k c m] by metis
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Miscellaneous›
(**************************************************************************************************)
(**************************************************************************************************)


lemma nth_pow_wits_SA_fun_prep:
  assumes "n > 0"
  assumes "h  carrier (SA m)"
  assumes "ρ  nth_pow_wits n"
  shows "is_semialgebraic m (h ¯mpow_res n ρ)"
  by(intro evimage_is_semialg assms  pow_res_is_univ_semialgebraic nth_pow_wits_closed(1)[of n])

definition kth_rt where
"kth_rt m (k::nat) f x = (if x  carrier (Qpm) then (THE b. b  carrier Qp  b[^]k = (f x)  ac (nat (ord ([k]𝟭)) + 1) b = 1)
                                                           else undefined )"

text‹Normalizing a semialgebraic function to have a constant angular component›

lemma ac_res_Unit_inc:
  assumes "n > 0"
  assumes "t  Units (Zp_res_ring n)"
  shows "ac n ([t]𝟭) = t"
proof-
  have 0: "[t]𝟭 𝟬"
    using assms by (metis Qp_char_0_int less_one less_or_eq_imp_le nat_neq_iff zero_not_in_residue_units)
  have 1: "[t]𝟭  𝒪p"
    by (metis Zp.one_closed Zp_int_mult_closed image_eqI inc_of_int)
  hence 2: "angular_component ([t]𝟭) = ac_Zp ([t]Zp𝟭Zp)"
    using angular_component_of_inclusion[of "[t]Zp𝟭Zp⇙"]
    by (metis "0" Qp.int_inc_zero Zp.int_inc_zero Zp_int_inc_closed inc_of_int not_nonzero_Qp)
  hence 3: "ac n ([t]𝟭) = ac_Zp ([t]Zp𝟭Zp) n"
    unfolding ac_def using 0 by presburger
  hence "val_Zp ([t]Zp𝟭Zp) = 0"
  proof-
    have "coprime p t"
      using assms
      by (metis coprime_commute less_one less_or_eq_imp_le nat_neq_iff padic_integers.residue_UnitsE padic_integers_axioms)
    then show ?thesis
      by (metis Zp_int_inc_closed Zp_int_inc_res coprime_mod_right_iff coprime_power_right_iff mod_by_0 order_refl p_residues residues.m_gt_one residues.mod_in_res_units val_Zp_0_criterion val_Zp_p val_Zp_p_int_unit zero_less_one zero_neq_one_class.one_neq_zero zero_not_in_residue_units)
  qed
  hence 4: "[t]Zp𝟭Zp Units Zp"
    using val_Zp_0_imp_unit by blast
  hence 5: "ac_Zp ([t]Zp𝟭Zp)  = [t]Zp𝟭Zp⇙"  using
      ac_Zp_of_Unit  val_Zp ([t] Zp𝟭Zp) = 0 by blast
  have 6: "ac_Zp ([t]Zp𝟭Zp) n  = t"
  proof-
    have "t  carrier (Zp_res_ring n)"
      using assms monoid.Units_closed[of "Zp_res_ring n" t] cring_def padic_integers.R_cring padic_integers_axioms ring_def by blast
    hence "t < p^n  t  0 "
      using p_residue_ring_car_memE by auto
    thus ?thesis
      unfolding 5 unfolding Zp_int_inc_rep p_residue_def residue_def by auto
  qed
  show ?thesis
    unfolding 3 using 6 by blast
qed

lemma val_of_res_Unit:
  assumes "n > 0"
  assumes "t  Units (Zp_res_ring n)"
  shows "val ([t]𝟭) = 0"
proof-
  have 0: "[t]𝟭 𝟬"
    using assms by (metis Qp_char_0_int less_one less_or_eq_imp_le nat_neq_iff zero_not_in_residue_units)
  have 1: "[t]𝟭  𝒪p"
    by (metis Zp.one_closed Zp_int_mult_closed image_eqI inc_of_int)
  hence 2: "angular_component ([t]𝟭) = ac_Zp ([t]Zp𝟭Zp)"
    using angular_component_of_inclusion[of "[t]Zp𝟭Zp⇙"]
    by (metis "0" Qp.int_inc_zero Zp.int_inc_zero Zp_int_inc_closed inc_of_int not_nonzero_Qp)
  hence 3: "ac n ([t]𝟭) = ac_Zp ([t]Zp𝟭Zp) n"
    unfolding ac_def using 0 by presburger
  hence "val_Zp ([t]Zp𝟭Zp) = 0"
  proof-
    have "coprime p t"
      using assms
      by (metis coprime_commute less_one less_or_eq_imp_le nat_neq_iff padic_integers.residue_UnitsE padic_integers_axioms)
    then show ?thesis
      by (metis Zp_int_inc_closed Zp_int_inc_res coprime_mod_right_iff coprime_power_right_iff mod_by_0 order_refl p_residues residues.m_gt_one residues.mod_in_res_units val_Zp_0_criterion val_Zp_p val_Zp_p_int_unit zero_less_one zero_neq_one_class.one_neq_zero zero_not_in_residue_units)
  qed
  then show ?thesis using assms
    by (metis Zp_int_inc_closed inc_of_int val_of_inc)
qed


lemma(in padic_integers) res_map_is_hom:
  assumes "N > 0"
  shows "ring_hom_ring Zp (Zp_res_ring N) (λ x. x N)"
  apply(rule ring_hom_ringI)
  apply (simp add: R.ring_axioms)
  using assms cring.axioms(1) local.R_cring apply blast
  using residue_closed apply blast
  using residue_of_prod apply blast
  using residue_of_sum apply blast
  using assms residue_of_one(1) by blast

lemma ac_of_fraction:
  assumes "N > 0"
  assumes "a  nonzero Qp"
  assumes "b  nonzero Qp"
  shows "ac N (a ÷ b) = ac N a Zp_res_ring NinvZp_res_ring Nac N b"
  using ac_mult[of a "inv b" N] ac_inv assms Qp.nonzero_closed nonzero_inverse_Qp by presburger

lemma pow_res_eq_rel:
  assumes "n > 0"
  assumes "b  carrier Qp"
  shows "{x  carrier Qp. pow_res n x = pow_res n b} = pow_res n b"
  apply(rule equalityI',  unfold mem_Collect_eq, metis pow_res_refl,
        intro conjI)
    using pow_res_def apply auto[1]
    apply(rule equal_pow_resI)
    using pow_res_def apply auto[1]
    using  pow_res_refl assms  by (metis equal_pow_resI)

lemma pow_res_is_univ_semialgebraic':
  assumes "n > 0"
  assumes "b  carrier Qp"
  shows "is_univ_semialgebraic {x  carrier Qp. pow_res n x = pow_res n b}"
  using assms pow_res_eq_rel pow_res_is_univ_semialgebraic by presburger

lemma evimage_eqI:
  assumes "c  carrier (SA n)"
  shows "c  ¯n{x  carrier Qp. P x} = {x  carrier (Qpn). P (c x)}"
  by(rule equalityI', unfold evimage_def mem_Collect_eq Int_iff,  intro conjI, auto
        , rule SA_car_closed[of _ n], auto simp: assms)

lemma SA_pow_res_is_semialgebraic:
  assumes "n > 0"
  assumes "b  carrier Qp"
  assumes "c  carrier (SA N)"
  shows "is_semialgebraic N  {x  carrier (QpN). pow_res n (c x) = pow_res n b}"
proof-
  have " c  ¯N{x  carrier Qp. pow_res n x = pow_res n b} = {x  carrier (QpN). pow_res n (c x) = pow_res n b}"
    apply(rule evimage_eqI) using assms by blast
  thus ?thesis
    using pow_res_is_univ_semialgebraic' evimage_is_semialg assms
    by (metis (no_types, lifting))
qed

lemma eint_diff_imp_eint:
  assumes "a  nonzero Qp"
  assumes "b  carrier Qp"
  assumes "val a = val b + eint i"
  shows "b  nonzero Qp"
  using assms val_zero
  by (metis Qp.nonzero_closed Qp.not_nonzero_memE not_eint_eq plus_eint_simps(2) val_ord')

lemma SA_minus_eval:
  assumes "f  carrier (SA n)"
  assumes "g  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  shows "(f  SA ng) x = f x  g x"
  using assms unfolding a_minus_def
  using SA_a_inv_eval SA_add  by metis

lemma Qp_cong_set_evimage:
  assumes "f  carrier (SA n)"
  assumes "a  carrier Zp"
  shows  "is_semialgebraic n (f ¯n(Qp_cong_set α a))"
  using assms Qp_cong_set_is_univ_semialgebraic evimage_is_semialg by blast

lemma SA_constant_res_set_semialg:
  assumes "l  carrier (Zp_res_ring n)"
  assumes "f  carrier (SA m)"
  shows "is_semialgebraic m {x  carrier (Qpm). f x  𝒪p  Qp_res (f x) n = l}"
proof-
  have 0: "{x  carrier (Qpm). f x  𝒪p  Qp_res (f x) n = l} = f  ¯m{x  𝒪p. Qp_res x n = l}"
    unfolding evimage_def  by blast
  show ?thesis unfolding 0
    by(rule evimage_is_semialg, rule assms, rule constant_res_set_semialg, rule assms)
qed

lemma val_ring_cong_set:
  assumes "f  carrier (SA k)"
  assumes "x. x  carrier (Qpk)  f x  𝒪p"
  assumes "t  carrier (Zp_res_ring n)"
  shows "is_semialgebraic k {x  carrier (Qpk). to_Zp (f x) n = t}"
proof-
  have 0: "[t] Zp𝟭Zp carrier Zp "
    by blast
  have 1: "([t] Zp𝟭Zp) n = t"
    using assms
    unfolding Zp_int_inc_rep p_residue_def residue_def residue_ring_def by simp
  have "{x  carrier (Qpk). to_Zp (f x) n = t} = f  ¯k{x  𝒪p. (to_Zp x) n = t}"
    unfolding evimage_def using assms by auto
  then show ?thesis using 0 1 assms  Qp_cong_set_evimage[of f k "[t]Zp𝟭Zp⇙" n] unfolding Qp_cong_set_def
    by presburger
qed

lemma val_ring_pullback_SA:
  assumes "N > 0"
  assumes "c  carrier (SA N)"
  shows "is_semialgebraic N {x  carrier (QpN). c x  𝒪p}"
proof-
  have 0: "{x  carrier (QpN). c x  𝒪p} = c  ¯N𝒪p"
    unfolding evimage_def by blast
  have 1: "is_univ_semialgebraic 𝒪p"
    using Qp_val_ring_is_univ_semialgebraic by blast
  show ?thesis using assms 0 1 evimage_is_semialg by presburger
qed

lemma(in padic_fields) res_eq_set_is_semialg:
  assumes "k > 0"
  assumes "c  carrier (SA k)"
  assumes "s  carrier (Zp_res_ring n)"
  shows "is_semialgebraic k {x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s}"
proof-
  obtain a where a_def: "a = [s]𝟭"
    by blast
  have 0: "a  𝒪p"
    using a_def
    by (metis Zp.one_closed Zp_int_mult_closed image_iff inc_of_int)
  have 1: "to_Zp a = [s]Zp𝟭Zp⇙"
    using 0 unfolding a_def
    by (metis Qp_def Zp_def Zp_int_inc_closed ι_def inc_to_Zp padic_fields.inc_of_int padic_fields_axioms)
  have 2: "([s]Zp𝟭Zp) n = s"
    using assms
    by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2))
  have 3: "{x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s} = c  ¯k⇙ B⇘n⇙[a]"
  proof
    show "{x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s}  c ¯k⇙ B⇘int n⇙[a]"
    proof fix x assume A: "x  {x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s}"
      then have  30: "x  carrier (Qpk)" by blast
      have 31: "c x  𝒪p" using A by blast
      have 32: "to_Zp (c x) n = s" using A by blast
      have 33: "to_Zp (c x)  carrier Zp"
        using 31 val_ring_memE to_Zp_closed by blast
      have 34: "to_Zp (c x) n = (to_Zp a) n"
        using "1" "2" "32" by presburger
      hence "((to_Zp (c x)) Zp(to_Zp a)) n = 0"
        using "1" "33" Zp_int_inc_closed res_diff_zero_fact'' by presburger
      hence 35: "val_Zp ((to_Zp (c x)) Zp(to_Zp a))  n"
        using "1" "33" "34" Zp.one_closed Zp_int_mult_closed val_Zp_dist_def val_Zp_dist_res_eq2 by presburger
      have 36: "val (c x  a) = val_Zp ((to_Zp (c x)) Zp(to_Zp a))"
        using 31 0
        by (metis to_Zp_minus to_Zp_val val_ring_minus_closed)
      hence "val (c x  a)  n"
        using 35 by presburger
      hence "c x   B⇘int n⇙[a]"
        using 31 c_ballI  val_ring_memE by blast
      thus "x  c ¯k⇙ B⇘int n⇙[a]"
        using 30  by blast
    qed
    show "c ¯k⇙ B⇘int n⇙[a]  {x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s}"
    proof fix x assume A: "x  c ¯k⇙ B⇘int n⇙[a]"
      have x_closed: "x  carrier (Qpk)"
        using A  by (meson evimage_eq)
      have 00: "val (c x  a)  n"
        using A c_ballE(2) by blast
      have cx_closed: "c x  carrier Qp"
        using x_closed assms function_ring_car_closed SA_car_memE(2) by blast
      hence 11: "c x  a  𝒪p"
      proof-
        have "(0::eint)  n"
          by (metis eint_ord_simps(1) of_nat_0_le_iff zero_eint_def)
        thus ?thesis using 00 order_trans[of "0::eint" n] Qp_val_ringI
          by (meson "0" Qp.minus_closed val_ring_memE cx_closed)
      qed
      hence 22: "c x  𝒪p"
      proof-
        have 00: "c x = (c x  a)  a"
          using cx_closed 0
          by (metis "11" Qp.add.inv_solve_right' Qp.minus_eq val_ring_memE(2))
        have 01: "(c x  a)  a  𝒪p"
          by(intro val_ring_add_closed 0 11)
        then show ?thesis
          using 0 11 image_iff  "00" by auto
      qed
      have 33: "val (c x  a) = val_Zp (to_Zp (c x) Zpto_Zp a)"
        using 11 22 0
        by (metis to_Zp_minus to_Zp_val)
      have 44: "val_Zp (to_Zp (c x) Zpto_Zp a)  n"
        using 33 00 by presburger
      have tzpcx: "to_Zp (c x)  carrier Zp"
        using 22 by (metis image_iff inc_to_Zp)
      have tzpa: "to_Zp a  carrier Zp"
        using 0 val_ring_memE to_Zp_closed by blast
      have 55: "(to_Zp (c x) Zpto_Zp a) n = 0"
        using 44 tzpcx tzpa Zp.minus_closed zero_below_val_Zp by blast
      hence 66: "to_Zp (c x) n = s"
        using 0 1 2 tzpa tzpcx
        by (metis res_diff_zero_fact(1))
      then show "x  {x  carrier (Qpk). c x  𝒪p  to_Zp (c x) n = s}"
        using "22" x_closed by blast
    qed
  qed
  thus ?thesis
    using evimage_is_semialg[of c k] 0  val_ring_memE assms(2) ball_is_univ_semialgebraic by presburger
qed

lemma SA_constant_res_set_semialg':
  assumes "f  carrier (SA m)"
  assumes "C  Qp_res_classes n"
  shows "is_semialgebraic m (f  ¯mC)"
proof-
  obtain l where l_def: "l  carrier (Zp_res_ring n)  C = Qp_res_class n ([l]𝟭)"
    using Qp_res_classes_wits assms by blast
  have C_eq: "C = Qp_res_class n ([l]𝟭)"
    using l_def by blast
  have 0: "Qp_res ([l]  𝟭) n = l"
    using l_def
    by (metis Qp_res_int_inc mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2))
  have 1: "f  ¯mC = {x  carrier (Qpm). f x  𝒪p  Qp_res (f x) n = l}"
    apply(rule equalityI')
    unfolding evimage_def C_eq Qp_res_class_def mem_Collect_eq unfolding 0 apply blast
    by blast
  show ?thesis
    unfolding 1 apply(rule SA_constant_res_set_semialg )
    using  l_def apply blast by(rule assms)
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Semialgebraic Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)


lemma UP_SA_n_is_ring:
  shows "ring (UP (SA n))"
  using SA_is_ring
  by (simp add: UP_ring.UP_ring UP_ring.intro)

lemma UP_SA_n_is_cring:
  shows "cring (UP (SA n))"
  using SA_is_cring
  by (simp add: UP_cring.UP_cring UP_cring.intro)

text‹The evaluation homomorphism from \texttt{Qp\_funs} to \texttt{Qp}›

definition eval_hom where
"eval_hom a = (λf. f a)"

lemma eval_hom_is_hom:
  assumes "a  carrier (Qpn)"
  shows "ring_hom_ring (Fun⇘nQp) Qp (eval_hom a)"
  apply(rule ring_hom_ringI)
  using Qp_funs_is_cring cring.axioms(1) apply blast
  apply (simp add: Qp.ring_axioms)
  apply (simp add: Qp.function_ring_car_mem_closed assms eval_hom_def)
  apply (metis Qp_funs_mult' assms eval_hom_def)
  apply (metis Qp_funs_add' assms eval_hom_def)
  by (metis function_one_eval assms eval_hom_def)

text‹the homomorphism from \texttt{Fun n Qp [x]} to \texttt{Qp [x]} induced by evaluation of coefficients›

definition Qp_fpoly_to_Qp_poly where
"Qp_fpoly_to_Qp_poly n a = poly_lift_hom (Fun⇘nQp) Qp (eval_hom a)"

lemma Qp_fpoly_to_Qp_poly_is_hom:
  assumes "a  carrier (Qpn)"
  shows "(Qp_fpoly_to_Qp_poly n a)  ring_hom (UP (Fun⇘nQp)) (Qp_x) "
  unfolding Qp_fpoly_to_Qp_poly_def
  apply(rule UP_cring.poly_lift_hom_is_hom)
  unfolding UP_cring_def
  apply (simp add: Qp_funs_is_cring)
  apply (simp add: UPQ.R_cring)
  using assms eval_hom_is_hom[of a] ring_hom_ring.homh by blast

lemma Qp_fpoly_to_Qp_poly_extends_apply:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (Fun⇘nQp)"
  shows "Qp_fpoly_to_Qp_poly n a (to_polynomial (Fun⇘nQp) f) = to_polynomial Qp (f a)"
  unfolding Qp_fpoly_to_Qp_poly_def
  using assms eval_hom_is_hom[of a] UP_cring.poly_lift_hom_extends_hom[of "Fun⇘nQp" Qp "eval_hom a" ]
        Qp.function_ring_car_memE[of f n] ring_hom_ring.homh
  unfolding eval_hom_def UP_cring_def
  using Qp_funs_is_cring UPQ.R_cring by blast

lemma Qp_fpoly_to_Qp_poly_X_var:
  assumes "a  carrier (Qpn)"
  shows "Qp_fpoly_to_Qp_poly n a (X_poly (Fun⇘nQp)) = X_poly Qp"
  unfolding X_poly_def Qp_fpoly_to_Qp_poly_def
  apply(rule UP_cring.poly_lift_hom_X_var) unfolding UP_cring_def
    apply (simp add: Qp_funs_is_cring)
      apply (simp add: UPQ.R_cring)
        using assms(1) eval_hom_is_hom ring_hom_ring.homh
  by blast

lemma Qp_fpoly_to_Qp_poly_monom:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (Fun⇘nQp)"
  shows "Qp_fpoly_to_Qp_poly n a (up_ring.monom (UP (Fun⇘nQp)) f m) = up_ring.monom Qp_x (f a) m"
  unfolding Qp_fpoly_to_Qp_poly_def
  using UP_cring.poly_lift_hom_monom[of "Fun⇘nQp" Qp "eval_hom a" f m] assms ring_hom_ring.homh
        eval_hom_is_hom[of a] unfolding eval_hom_def UP_cring_def
  using Qp_funs_is_cring UPQ.R_cring by blast

lemma Qp_fpoly_to_Qp_poly_coeff:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (UP (Fun⇘nQp))"
  shows "Qp_fpoly_to_Qp_poly n a f k = (f k) a"
  using assms UP_cring.poly_lift_hom_cf[of "Fun⇘nQp" Qp "eval_hom a" f k] eval_hom_is_hom[of a]
  unfolding  Qp_fpoly_to_Qp_poly_def eval_hom_def
  using Qp_funs_is_cring ring_hom_ring.homh  ring_hom_ring.homh
  unfolding eval_hom_def UP_cring_def
  using UPQ.R_cring by blast

lemma Qp_fpoly_to_Qp_poly_eval:
  assumes "a  carrier (Qpn)"
  assumes "P  carrier (UP (Fun⇘nQp))"
  assumes "f  carrier (Fun⇘nQp)"
  shows "(UP_cring.to_fun (Fun⇘nQp) P f) a  = UP_cring.to_fun Qp (Qp_fpoly_to_Qp_poly n a P) (f a)"
  unfolding Qp_fpoly_to_Qp_poly_def
  using UP_cring.poly_lift_hom_eval[of "Fun⇘nQp" Qp "eval_hom a" P f]
        eval_hom_is_hom[of a] eval_hom_def assms ring_hom_ring.homh Qp_funs_is_cring
  unfolding eval_hom_def UP_cring_def
  using UPQ.R_cring by blast

lemma Qp_fpoly_to_Qp_poly_sub:
  assumes "f  carrier (UP (Fun⇘nQp))"
  assumes "g  carrier (UP (Fun⇘nQp))"
  assumes "a  carrier (Qpn)"
  shows "Qp_fpoly_to_Qp_poly n a (compose (Fun⇘nQp) f g) = compose Qp (Qp_fpoly_to_Qp_poly n a f) (Qp_fpoly_to_Qp_poly n a g)"
  unfolding Qp_fpoly_to_Qp_poly_def
  using assms UP_cring.poly_lift_hom_sub[of "Fun⇘nQp" Qp "eval_hom a" f g]
        eval_hom_is_hom[of a] ring_hom_ring.homh[of "Fun⇘nQp" Qp "eval_hom a"]
        Qp_funs_is_cring
  unfolding eval_hom_def UP_cring_def
  using UPQ.R_cring by blast

lemma Qp_fpoly_to_Qp_poly_taylor_poly:
  assumes "F  carrier (UP (Fun⇘nQp))"
  assumes "c  carrier (Fun⇘nQp)"
  assumes "a  carrier (Qpn)"
  shows "Qp_fpoly_to_Qp_poly n a (taylor_expansion (Fun⇘nQp) c F) =
        taylor_expansion Qp (c a) (Qp_fpoly_to_Qp_poly n a F)"
proof-
  have 0: "X_poly (Fun⇘nQp) UP (Fun⇘nQp)to_polynomial (Fun⇘nQp) c  carrier (UP (Fun⇘nQp))"
    by (metis Qp_funs_is_cring UP_cring.X_plus_closed UP_cring_def X_poly_plus_def assms(2))
  have 1: "poly_lift_hom (Fun⇘nQp) Qp (eval_hom a) (X_poly (Fun⇘nQp) UP (Fun⇘nQp)to_polynomial (Fun⇘nQp) c) = X_poly Qp Qp_xUPQ.to_poly (c a)"
  proof-
    have 10: "poly_lift_hom (Fun⇘nQp) Qp (eval_hom a)  ring_hom (UP (Fun⇘nQp)) Qp_x"
      using Qp_fpoly_to_Qp_poly_def Qp_fpoly_to_Qp_poly_is_hom assms
      by presburger
    have 11: " to_polynomial (Fun⇘nQp) c  carrier (UP (Fun⇘nQp))"
      by (meson Qp_funs_is_cring UP_cring.intro UP_cring.to_poly_closed assms)
    have 12: "X_poly (Fun⇘nQp)  carrier (UP (Fun⇘nQp)) "
      using UP_cring.X_closed[of "Fun⇘nQp"] unfolding UP_cring_def
      using Qp_funs_is_cring
      by blast
    have "Qp_fpoly_to_Qp_poly n a (X_poly (Fun⇘nQp) UP (Fun⇘nQp)to_polynomial (Fun⇘nQp) c) =
    Qp_fpoly_to_Qp_poly n a (X_poly (Fun⇘nQp)) Qp_xQp_fpoly_to_Qp_poly n a (to_polynomial (Fun⇘nQp) c)"
      using assms 0 10 11 12 Qp_fpoly_to_Qp_poly_extends_apply[of a n c] Qp_fpoly_to_Qp_poly_is_hom[of a] Qp_fpoly_to_Qp_poly_X_var[of a]
      using ring_hom_add[of "Qp_fpoly_to_Qp_poly n a"  "UP (Fun⇘nQp)" Qp_x "X_poly (Fun⇘nQp)" "to_polynomial (Fun⇘nQp) c" ]
      unfolding Qp_fpoly_to_Qp_poly_def
      by blast
     then show ?thesis
       using Qp_fpoly_to_Qp_poly_X_var Qp_fpoly_to_Qp_poly_def Qp_fpoly_to_Qp_poly_extends_apply assms
       by metis
  qed
  have 2: "poly_lift_hom (Fun⇘nQp) Qp (eval_hom a) (compose (Fun⇘nQp) F (X_poly (Fun⇘nQp) UP (Fun⇘nQp)to_polynomial (Fun⇘nQp) c)) =
   UPQ.sub (poly_lift_hom (Fun⇘nQp) Qp (eval_hom a) F)
    (poly_lift_hom (Fun⇘nQp) Qp (eval_hom a) (X_poly (Fun⇘nQp) UP (Fun⇘nQp)to_polynomial (Fun⇘nQp) c))"
    using 0 1 Qp_fpoly_to_Qp_poly_sub[of F n "X_poly_plus (Fun⇘nQp) c" a] assms
    unfolding Qp_fpoly_to_Qp_poly_def X_poly_plus_def
    by blast
  show ?thesis
    using assms 0 1
    unfolding Qp_fpoly_to_Qp_poly_def  taylor_expansion_def X_poly_plus_def
    using "2" by presburger
qed

lemma SA_is_UP_cring:
  shows "UP_cring (SA n)"
  unfolding UP_cring_def
  by (simp add: SA_is_cring)

lemma eval_hom_is_SA_hom:
  assumes "a  carrier (Qpn)"
  shows "ring_hom_ring (SA n) Qp (eval_hom a)"
  apply(rule ring_hom_ringI)
  using SA_is_cring cring.axioms(1) assms(1) apply blast
  using Qp.ring_axioms apply blast
  apply (metis (no_types, lifting) SA_car assms eval_hom_def Qp.function_ring_car_mem_closed semialg_functions_memE(2))
  apply (metis (mono_tags, lifting) Qp_funs_mult' SA_car SA_times assms eval_hom_def semialg_functions_memE(2))
  apply (metis (mono_tags, lifting) Qp_funs_add' SA_car SA_plus assms eval_hom_def semialg_functions_memE(2))
  using Qp_constE Qp.one_closed SA_one assms eval_hom_def function_one_as_constant
  by (metis function_one_eval)

text‹the homomorphism from \texttt{(SA n)[x]} to \texttt{Qp [x]} induced by evaluation of coefficients›

definition SA_poly_to_Qp_poly where
"SA_poly_to_Qp_poly n a = poly_lift_hom (SA n) Qp (eval_hom a)"

lemma SA_poly_to_Qp_poly_is_hom:
  assumes "a  carrier (Qpn)"
  shows "(SA_poly_to_Qp_poly n a)  ring_hom (UP (SA n)) (Qp_x) "
  unfolding SA_poly_to_Qp_poly_def
  apply(rule UP_cring.poly_lift_hom_is_hom)
  using SA_is_cring assms(1) UP_cring.intro apply blast
  apply (simp add: UPQ.R_cring)
  using assms eval_hom_is_SA_hom ring_hom_ring.homh by blast

lemma SA_poly_to_Qp_poly_closed:
  assumes "a  carrier (Qpn)"
  assumes "P  carrier (UP (SA n))"
  shows "SA_poly_to_Qp_poly n a P  carrier Qp_x"
  using assms SA_poly_to_Qp_poly_is_hom[of a] ring_hom_closed[of "SA_poly_to_Qp_poly n a" "UP (SA n)" Qp_x P]
  by blast

lemma SA_poly_to_Qp_poly_add:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  shows "SA_poly_to_Qp_poly n a (f UP (SA n)g) = SA_poly_to_Qp_poly n a f Qp_xSA_poly_to_Qp_poly n a g"
  using SA_poly_to_Qp_poly_is_hom ring_hom_add assms
  by (metis (no_types, opaque_lifting))

lemma SA_poly_to_Qp_poly_minus:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
   shows "SA_poly_to_Qp_poly n a (f UP (SA n)g) = SA_poly_to_Qp_poly n a f Qp_xSA_poly_to_Qp_poly n a g"
  using SA_poly_to_Qp_poly_is_hom[of a] assms SA_is_ring[of n]
        ring.ring_hom_minus[of "UP (SA n)" Qp_x "SA_poly_to_Qp_poly n a" f g] UP_SA_n_is_ring
        UPQ.UP_ring
  by blast

lemma SA_poly_to_Qp_poly_mult:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  shows "SA_poly_to_Qp_poly n a (f UP (SA n)g) = SA_poly_to_Qp_poly n a f Qp_xSA_poly_to_Qp_poly n a g"
  using SA_poly_to_Qp_poly_is_hom ring_hom_mult assms
  by (metis (no_types, opaque_lifting))

lemma SA_poly_to_Qp_poly_extends_apply:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (SA n)"
  shows "SA_poly_to_Qp_poly n a (to_polynomial (SA n) f) = to_polynomial Qp (f a)"
  unfolding SA_poly_to_Qp_poly_def
  using assms eval_hom_is_SA_hom[of a] UP_cring.poly_lift_hom_extends_hom[of "SA n" Qp "eval_hom a" f]
         eval_hom_def   SA_is_cring Qp.cring_axioms ring_hom_ring.homh
  unfolding eval_hom_def UP_cring_def
  by blast

lemma SA_poly_to_Qp_poly_X_var:
  assumes "a  carrier (Qpn)"
  shows "SA_poly_to_Qp_poly n a (X_poly (SA n)) = X_poly Qp"
  unfolding X_poly_def SA_poly_to_Qp_poly_def
  apply(rule UP_cring.poly_lift_hom_X_var)
    using SA_is_cring assms(1)
    using UP_cring.intro apply blast
    apply (simp add: Qp.cring_axioms)
    using assms eval_hom_is_SA_hom ring_hom_ring.homh by blast

lemma SA_poly_to_Qp_poly_X_plus:
  assumes "a  carrier (Qpn)"
  assumes "c  carrier (SA n)"
  shows "SA_poly_to_Qp_poly n a (X_poly_plus (SA n) c) = UPQ.X_plus (c a)"
  unfolding X_poly_plus_def
  using assms SA_poly_to_Qp_poly_add[of a n "X_poly (SA n)" "to_polynomial (SA n) c"]
        SA_poly_to_Qp_poly_extends_apply[of a n c] UP_cring.X_closed[of "SA n"] SA_is_cring[of n]
        SA_poly_to_Qp_poly_X_var[of a] UP_cring.to_poly_closed[of "SA n" c]
  unfolding UP_cring_def
  by metis

lemma SA_poly_to_Qp_poly_X_minus:
  assumes "a  carrier (Qpn)"
  assumes "c  carrier (SA n)"
  shows "SA_poly_to_Qp_poly n a (X_poly_minus (SA n) c) = UPQ.X_minus (c a)"
  unfolding X_poly_minus_def
  using assms SA_poly_to_Qp_poly_minus[of a n "X_poly (SA n)" "to_polynomial (SA n) c"]
        SA_poly_to_Qp_poly_extends_apply[of a n c] UP_cring.X_closed[of "SA n"] SA_is_cring[of n]
        SA_poly_to_Qp_poly_X_var[of a n] UP_cring.to_poly_closed[of "SA n" c]
  unfolding UP_cring_def
  by metis

lemma SA_poly_to_Qp_poly_monom:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (SA n)"
  shows "SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) f m) = up_ring.monom Qp_x (f a) m"
  unfolding SA_poly_to_Qp_poly_def
  using UP_cring.poly_lift_hom_monom[of "SA n" Qp "eval_hom a" f n] assms eval_hom_is_SA_hom eval_hom_def
   SA_is_cring Qp.cring_axioms UP_cring.poly_lift_hom_monom ring_hom_ring.homh
  by (metis UP_cring.intro)

lemma SA_poly_to_Qp_poly_coeff:
  assumes "a  carrier (Qpn)"
  assumes "f  carrier (UP (SA n))"
  shows "SA_poly_to_Qp_poly n a f k = (f k) a"
  using assms UP_cring.poly_lift_hom_cf[of "SA n" Qp "eval_hom a" f k] eval_hom_is_SA_hom[of a]
  using SA_is_cring Qp.cring_axioms ring_hom_ring.homh
  unfolding  SA_poly_to_Qp_poly_def eval_hom_def UP_cring_def
  by blast

lemma SA_poly_to_Qp_poly_eval:
  assumes "a  carrier (Qpn)"
  assumes "P  carrier (UP (SA n))"
  assumes "f  carrier (SA n)"
  shows "(UP_cring.to_fun (SA n) P f) a  = UP_cring.to_fun Qp (SA_poly_to_Qp_poly n a P) (f a)"
  unfolding SA_poly_to_Qp_poly_def
  using UP_cring.poly_lift_hom_eval[of "SA n" Qp "eval_hom a" P f]
        eval_hom_is_SA_hom[of a] eval_hom_def assms SA_is_cring Qp.cring_axioms ring_hom_ring.homh
  unfolding  SA_poly_to_Qp_poly_def eval_hom_def UP_cring_def
  by blast

lemma SA_poly_to_Qp_poly_sub:
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  assumes "a  carrier (Qpn)"
  shows "SA_poly_to_Qp_poly n a (compose (SA n) f g) = compose Qp (SA_poly_to_Qp_poly n a f) (SA_poly_to_Qp_poly n a g)"
  unfolding SA_poly_to_Qp_poly_def
  using assms UP_cring.poly_lift_hom_sub[of "SA n" Qp "eval_hom a" f g]
        eval_hom_is_SA_hom[of a] ring_hom_ring.homh[of "SA n" Qp "eval_hom a"]
        SA_is_cring Qp.cring_axioms
  unfolding  SA_poly_to_Qp_poly_def eval_hom_def UP_cring_def
  by blast

lemma SA_poly_to_Qp_poly_deg_bound:
  assumes "g  carrier (UP (SA m))"
  assumes "x  carrier (Qpm)"
  shows "deg Qp (SA_poly_to_Qp_poly m x g)  deg (SA m) g"
  apply(rule  UPQ.deg_leqI)
    using assms SA_poly_to_Qp_poly_closed[of x m g] apply blast
  proof- fix n assume A: "deg (SA m) g < n"
    then have "g n = 𝟬SA m⇙"
      using assms SA_is_UP_cring[of m] UP_cring.UP_car_memE(2) by blast
    thus "SA_poly_to_Qp_poly m x g n = 𝟬"
      using assms SA_poly_to_Qp_poly_coeff[of x m g n] function_zero_eval SA_zero by presburger
  qed

lemma SA_poly_to_Qp_poly_taylor_poly:
  assumes "F  carrier (UP (SA n))"
  assumes "c  carrier (SA n)"
  assumes "a  carrier (Qpn)"
  shows "SA_poly_to_Qp_poly n a (taylor_expansion (SA n) c F) =
        taylor_expansion Qp (c a) (SA_poly_to_Qp_poly n a F)"
  unfolding SA_poly_to_Qp_poly_def using assms Qp.cring_axioms SA_is_cring eval_hom_def
      eval_hom_is_SA_hom UP_cring.poly_lift_hom_comm_taylor_expansion[of "SA n" Qp "eval_hom a" F c] ring_hom_ring.homh
  unfolding  SA_poly_to_Qp_poly_def eval_hom_def UP_cring_def
  by metis

lemma SA_poly_to_Qp_poly_comm_taylor_term:
  assumes "F  carrier (UP (SA n))"
  assumes "c  carrier (SA n)"
  assumes "a  carrier (Qpn)"
  shows "SA_poly_to_Qp_poly n a (UP_cring.taylor_term (SA n) c F i) =
        UP_cring.taylor_term Qp (c a) (SA_poly_to_Qp_poly n a F) i"
  unfolding SA_poly_to_Qp_poly_def using assms  Qp.cring_axioms SA_is_cring eval_hom_def
      eval_hom_is_SA_hom UP_cring.poly_lift_hom_comm_taylor_term[of "SA n" Qp "eval_hom a" F c i] ring_hom_ring.homh
  unfolding  SA_poly_to_Qp_poly_def eval_hom_def UP_cring_def
  by metis

lemma SA_poly_to_Qp_poly_comm_pderiv:
  assumes "F  carrier (UP (SA n))"
  assumes "a  carrier (Qpn)"
  shows "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) F) =
        UP_cring.pderiv Qp (SA_poly_to_Qp_poly n a F)"
  apply(rule UP_ring.poly_induct3[of "SA n" F]) unfolding UP_ring_def
  apply (simp add: SA_is_ring assms(1))
  using assms apply blast
proof-
  show "p q. q  carrier (UP (SA n)) 
           p  carrier (UP (SA n)) 
           SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) p) = UPQ.pderiv (SA_poly_to_Qp_poly n a p) 
           SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) q) = UPQ.pderiv (SA_poly_to_Qp_poly n a q) 
           SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (p UP (SA n)q)) = UPQ.pderiv (SA_poly_to_Qp_poly n a (p UP (SA n)q))"
  proof- fix p q assume A: "q  carrier (UP (SA n))"
           "p  carrier (UP (SA n))"
           "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) p) = UPQ.pderiv (SA_poly_to_Qp_poly n a p)"
           "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) q) = UPQ.pderiv (SA_poly_to_Qp_poly n a q)"
    show "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (p UP (SA n)q)) = UPQ.pderiv (SA_poly_to_Qp_poly n a (p UP (SA n)q))"
    proof-
      have 0: "SA_poly_to_Qp_poly n a p  carrier (UP Qp)"
        using A assms SA_poly_to_Qp_poly_closed[of a n p]
        by blast
      have 1: "SA_poly_to_Qp_poly n a q  carrier (UP Qp)"
        using A SA_poly_to_Qp_poly_closed[of a n q] assms by blast
      have 2: "UPQ.pderiv (SA_poly_to_Qp_poly n a p)  carrier (UP Qp)"
        using  UPQ.pderiv_closed[of "SA_poly_to_Qp_poly n a p"] 0  by blast
      have 3: "UPQ.pderiv (SA_poly_to_Qp_poly n a q)  carrier (UP Qp)"
        using A assms  UPQ.pderiv_closed[of "SA_poly_to_Qp_poly n a q"] 1 by blast
      have 4: "UP_cring.pderiv (SA n) p  carrier (UP (SA n))"
        using A UP_cring.pderiv_closed[of "SA n" p] unfolding UP_cring_def
        using SA_is_cring assms(1) by blast
      have 5: "UP_cring.pderiv (SA n) q  carrier (UP (SA n))"
        using A UP_cring.pderiv_closed[of "SA n" q] unfolding UP_cring_def
        using SA_is_cring assms(1) by blast
      have 6: "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) p UP (SA n)UP_cring.pderiv (SA n) q) =
    SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) p) UP QpSA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) q)"
        using A 4 5 SA_poly_to_Qp_poly_add assms  by blast
      have 7: "UPQ.pderiv (SA_poly_to_Qp_poly n a p UP QpSA_poly_to_Qp_poly n a q) =
    UPQ.pderiv (SA_poly_to_Qp_poly n a p) UP QpUPQ.pderiv (SA_poly_to_Qp_poly n a q)"
        using "0" "1" UPQ.pderiv_add by blast
      have 8: "UP_cring.pderiv (SA n) (p UP (SA n)q) = UP_cring.pderiv (SA n) p UP (SA n)UP_cring.pderiv (SA n) q"
        using A assms UP_cring.pderiv_add[of "SA n" p q]
        unfolding UP_cring_def  using SA_is_cring by blast
      have 9: "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (p UP (SA n)q)) =
    UPQ.pderiv (SA_poly_to_Qp_poly n a p) UP QpUPQ.pderiv (SA_poly_to_Qp_poly n a q)"
        using A 6 8 by presburger
      have 10: "UPQ.pderiv (SA_poly_to_Qp_poly n a (p UP (SA n)q)) =
    UPQ.pderiv (SA_poly_to_Qp_poly n a p) UP QpUPQ.pderiv (SA_poly_to_Qp_poly n a q)"
        using "7" A(1) A(2) SA_poly_to_Qp_poly_add assms  by presburger
      show ?thesis using 9 10
        by presburger
    qed
  qed
  show "aa na.
       aa  carrier (SA n) 
       SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (up_ring.monom (UP (SA n)) aa na)) =
       UPQ.pderiv (SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) aa na))"
  proof- fix b m assume A: "b  carrier (SA n)"
    show "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (up_ring.monom (UP (SA n)) b m)) =
       UPQ.pderiv (SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) b m))"
    proof-
      have 0: "(UP_cring.pderiv (SA n) (up_ring.monom (UP (SA n)) b m)) =
              (up_ring.monom (UP (SA n)) ([m]SA nb) (m-1))"
        using UP_cring.pderiv_monom[of "SA n" b m] unfolding UP_cring_def
        using SA_is_cring b  carrier (SA n) assms(1) by blast
      have 1: "(SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) b m)) =up_ring.monom (UP Qp) (b a) m"
        using SA_poly_to_Qp_poly_monom b  carrier (SA n) assms  by blast
      have 2: "b a  carrier Qp"
        using A assms Qp.function_ring_car_mem_closed SA_car_memE(2) by metis
      hence 3: "UPQ.pderiv (SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) b m)) = up_ring.monom (UP Qp) ([m](b a)) (m-1)"
        using 1 2 A UPQ.pderiv_monom[of "b a" m]
        by presburger
      have 4: "[m] SA nb  carrier (SA n)"
        using A assms  SA_is_cring[of n] ring.add_pow_closed[of "SA n" b m] SA_is_ring
        by blast
      have 5: "SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) ([m] SA nb) (m-1)) = up_ring.monom (UP Qp) (([m] SA nb) a) (m-1)"
        using SA_poly_to_Qp_poly_monom[of a n "[m]SA nb" "m-1"] assms 4 by blast
      have 6: "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) (up_ring.monom (UP (SA n)) b m)) = up_ring.monom (UP Qp) (([m] SA nb) a) (m - 1)"
         using 5 0  by presburger
      thus ?thesis using assms A 3 6 SA_add_pow_apply[of b n a]
        by auto
    qed
  qed
qed

lemma SA_poly_to_Qp_poly_pderiv:
  assumes "g  carrier (UP (SA m))"
  assumes "x  carrier (Qpm)"
  shows "UPQ.pderiv (SA_poly_to_Qp_poly m x g) = (SA_poly_to_Qp_poly m x (pderiv m g))"
proof
  fix n
  have 0: "UPQ.pderiv (SA_poly_to_Qp_poly m x g) n = [Suc n]  SA_poly_to_Qp_poly m x g (Suc n)"
    by(rule UPQ.pderiv_cfs[of "SA_poly_to_Qp_poly m x g"n], rule SA_poly_to_Qp_poly_closed, rule assms , rule assms)
  have 1: "SA_poly_to_Qp_poly m x (UPSA.pderiv m g) n = UPSA.pderiv m g n x"
    by(rule SA_poly_to_Qp_poly_coeff[of x m "UPSA.pderiv m g" n], rule assms, rule UPSA.pderiv_closed, rule assms)
  have 2: "SA_poly_to_Qp_poly m x (UPSA.pderiv m g) n = ([Suc n] SA mg (Suc n)) x"
    using UPSA.pderiv_cfs[of g m n] assms unfolding 1 by auto
  show "UPQ.pderiv (SA_poly_to_Qp_poly m x g) n = SA_poly_to_Qp_poly m x (UPSA.pderiv m g) n"
    unfolding 0 2 using SA_poly_to_Qp_poly_coeff assms
    by (metis "0" "2" SA_poly_to_Qp_poly_comm_pderiv)
qed

lemma(in UP_cring) pderiv_deg_lt:
  assumes "f  carrier (UP R)"
  assumes "deg R f > 0"
  shows "deg R (pderiv f) < deg R f"
proof-
  obtain n where n_def: "n = deg R f"
    by blast
  have 0: "k. k  n  pderiv f k = 𝟬"
    using pderiv_cfs assms unfolding n_def
    by (simp add: UP_car_memE(2))
  obtain k where k_def: "n = Suc k"
    using n_def assms gr0_implies_Suc by presburger
  have "deg R (pderiv f)  k"
    apply(rule deg_leqI)
    using P_def assms(1) pderiv_closed apply presburger
    apply(rule 0)
    unfolding k_def by presburger
  thus ?thesis using k_def unfolding n_def by linarith
qed

lemma deg_pderiv:
  assumes "f  carrier (UP (SA m))"
  assumes "deg (SA m) f > 0"
  shows "deg (SA m) (pderiv m f) = deg (SA m) f - 1"
proof-
  obtain n where n_def: "n = deg (SA m) f"
    by blast
  have 0: "f n  𝟬SA m⇙"
    unfolding n_def using assms UPSA.deg_ltrm by fastforce
  have 1: "(pderiv m f) (n-1) = [n]SA m(f n)"
    using assms unfolding n_def using  Suc_diff_1 UPSA.pderiv_cfs by presburger
  have 2: "deg (SA m) (pderiv m f)  (n-1)"
    using 0 assms SA_char_zero
    by (metis "1" UPSA.deg_eqI UPSA.lcf_closed UPSA.pderiv_closed n_def nat_le_linear)
  have 3: "deg (SA m) (pderiv m f) < deg (SA m) f"
    using assms  pderiv_deg_lt by auto
  thus ?thesis using 2  unfolding  n_def by presburger
qed

lemma SA_poly_to_Qp_poly_smult:
  assumes "a  carrier (SA m)"
  assumes "f  carrier (UP (SA m))"
  assumes "x  carrier (Qpm)"
  shows "SA_poly_to_Qp_poly m x (a UP (SA m)f) = a x UP QpSA_poly_to_Qp_poly m x f"
proof-
  have 0: "a UP (SA m)f = to_polynomial (SA m) a UP (SA m)f"
    using assms  UPSA.to_poly_mult_simp(1) by presburger
  have 1: "SA_poly_to_Qp_poly m x (a UP (SA m)f) = SA_poly_to_Qp_poly m x (to_polynomial (SA m) a) UP QpSA_poly_to_Qp_poly m x f"
    unfolding 0 apply(rule SA_poly_to_Qp_poly_mult)
    using assms apply blast
    using assms to_poly_closed  apply blast
    using assms by blast
  have 2: "SA_poly_to_Qp_poly m x (a UP (SA m)f) = (to_polynomial Qp (a x)) UP QpSA_poly_to_Qp_poly m x f"
    unfolding 1 using assms SA_poly_to_Qp_poly_monom unfolding to_polynomial_def
    by presburger
  show ?thesis
    unfolding 2 apply(rule UP_cring.to_poly_mult_simp(1)[of Qp "a x" "SA_poly_to_Qp_poly m x f"])
    unfolding UP_cring_def
    using F.R_cring apply blast
    using assms SA_car_memE apply blast
    using assms SA_poly_to_Qp_poly_closed[of x m f] by blast
qed

lemma SA_poly_constant_res_class_semialg:
  assumes "f  carrier (UP (SA m))"
  assumes "i x. x  carrier (Qpm)  f i x  𝒪p"
  assumes "deg (SA m) f  d"
  assumes "C  poly_res_classes n d"
  shows "is_semialgebraic m {x  carrier (Qpm). SA_poly_to_Qp_poly m x f  C}"
proof-
  obtain Fs where Fs_def: "Fs = f ` {..d}"
    by blast
  obtain g where g_def: "g  val_ring_polys_grad d  C = poly_res_class n d g"
    using assms unfolding poly_res_classes_def by blast
  have C_eq: " C = poly_res_class n d g"
    using g_def by blast
  have 0: "{x  carrier (Qpm). SA_poly_to_Qp_poly m x f  C} =
            ( i  {..d}. {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n})"
    apply(rule equalityI')
    unfolding mem_Collect_eq
  proof(rule InterI)
    fix x S
    assume A: " x  carrier (Qpm)  SA_poly_to_Qp_poly m x f  C"
           "S  (λi. {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n}) ` {..d}"
    have 0: "C = poly_res_class n d (SA_poly_to_Qp_poly m x f)"
      unfolding C_eq
      apply(rule equalityI', rule poly_res_class_memI, rule poly_res_class_memE[of _ n d g], blast
          , rule poly_res_class_memE[of _ n d g], blast,  rule poly_res_class_memE[of _ n d g], blast)
      using poly_res_class_memE[of _ n d ]A
       apply (metis (no_types, lifting) C_eq)
      apply(rule poly_res_class_memI, rule poly_res_class_memE[of _ n d "SA_poly_to_Qp_poly m x f"], blast,
            rule poly_res_class_memE[of _ n d "SA_poly_to_Qp_poly m x f"], blast,
            rule poly_res_class_memE[of _ n d "SA_poly_to_Qp_poly m x f"], blast)
      using poly_res_class_memE[of _ n d ]A
       by (metis (no_types, lifting) C_eq)
    obtain i where i_def: "i  {..d} 
                            S = {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n}"
      using A by blast
    have S_eq: "S = {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n}"
      using i_def by blast
    have 1: "i. SA_poly_to_Qp_poly m x f i = f i x"
      apply(rule SA_poly_to_Qp_poly_coeff)
      using A apply blast by(rule assms)
    have 2: "deg Qp (SA_poly_to_Qp_poly m x f)  d"
      using assms SA_poly_to_Qp_poly_deg_bound[of f m x]
      using A(1) by linarith
    have 3: "Qp_res (SA_poly_to_Qp_poly m x f i) n = Qp_res (g i) n"
      apply(rule poly_res_class_memE[of _ _ d], rule poly_res_class_memI)
      using g_def val_ring_polys_grad_memE apply blast
      using g_def val_ring_polys_grad_memE apply blast
      using g_def val_ring_polys_grad_memE apply blast
      apply(rule poly_res_class_memE[of _ _ d],rule poly_res_class_memI)
         apply(rule SA_poly_to_Qp_poly_closed)
      using A apply blast
         apply(rule assms)
      apply(rule 2)
      unfolding 1 using assms A apply blast
      using A unfolding C_eq
      using poly_res_class_memE(4)[of "SA_poly_to_Qp_poly m x f" n d g]
      unfolding 1 by metis
    show "x  S"
      using A 3 assms
      unfolding S_eq mem_Collect_eq unfolding 1
      by blast
  next

    show "x. x  (id. {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n})  x  carrier (Qpm)  SA_poly_to_Qp_poly m x f  C"
    proof-
      fix x assume A: "x  (id. {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n})"
      have x_closed: "x  carrier (Qpm)"
        using A by blast
      have 0: "i. SA_poly_to_Qp_poly m x f i = f i x"
        apply(rule SA_poly_to_Qp_poly_coeff)
        using A apply blast by(rule assms)
      have 1: "deg Qp (SA_poly_to_Qp_poly m x f)  d"
        using assms SA_poly_to_Qp_poly_deg_bound[of f m x]
        using x_closed by linarith
      have 2: "i. Qp_res (f i x) n = Qp_res (g i) n"
      proof- fix i
        have 20: "i > d  i > deg Qp g"
          using g_def val_ring_polys_grad_memE(2) by fastforce
        have 21: "i > d  g i= 𝟬"
          using 20 g_def val_ring_polys_grad_memE UPQ.deg_leE by blast
        have 22: "i > d  Qp_res (g i) n = 0"
          unfolding 21 Qp_res_zero by blast
        have 23: "i > d  SA_poly_to_Qp_poly m x f i = 𝟬"
          apply(rule UPQ.deg_leE, rule SA_poly_to_Qp_poly_closed, rule x_closed, rule assms)
          by(rule le_less_trans[of _ d], rule 1, blast)
        show " Qp_res (f i x) n = Qp_res (g i) n"
          apply(cases "i  d")
          using A apply blast
          using 22 21 1 23 unfolding 0
          by (metis less_or_eq_imp_le linorder_neqE_nat)
      qed
      have 3: "SA_poly_to_Qp_poly m x f  C"
        unfolding C_eq
        apply(rule poly_res_class_memI, rule SA_poly_to_Qp_poly_closed, rule x_closed, rule assms, rule 1)
        unfolding 0
        by(rule assms, rule x_closed, rule 2)
      show " x  carrier (Qpm)  SA_poly_to_Qp_poly m x f  C"
        using x_closed 3 by blast
    qed
  qed
  have 1: "i. is_semialgebraic m {x  carrier (Qpm). f i x  𝒪p  Qp_res (f i x) n = Qp_res (g i) n}"
    apply(rule SA_constant_res_set_semialg, rule Qp_res_closed, rule val_ring_polys_grad_memE[of _ d])
    using g_def apply blast
    using assms UPSA.UP_car_memE(1) by blast
  show "is_semialgebraic m {x  carrier (Qpm). SA_poly_to_Qp_poly m x f  C}"
    unfolding 0
    apply(rule finite_intersection_is_semialg, blast, blast, rule subsetI)
    using 1 unfolding is_semialgebraic_def by blast
qed

text‹Maps a polynomial $F(t) \in UP (SA n)$ to a function sending $(t, a) \in (Q_p (n + 1) \mapsto F(a)(t) \in Q_p$ ›

definition SA_poly_to_SA_fun where
  "SA_poly_to_SA_fun n P = (λ a  carrier (QpSuc n). UP_cring.to_fun Qp (SA_poly_to_Qp_poly n (tl a) P) (hd a))"

lemma SA_poly_to_SA_fun_is_fun:
  assumes "P  carrier (UP (SA n))"
  shows "SA_poly_to_SA_fun n P  (carrier (QpSuc n)  carrier Qp)"
proof fix x assume A: "x  carrier (QpSuc n)"
  obtain t where t_def: "t = hd x" by blast
  obtain a where a_def: "a = tl x" by blast
  have t_closed: "t  carrier Qp"
    using A t_def cartesian_power_head
    by blast
  have a_closed: "a  carrier (Qpn)"
    using A a_def cartesian_power_tail
    by blast
  have 0: "SA_poly_to_SA_fun n P x = UP_cring.to_fun Qp (SA_poly_to_Qp_poly n a P) t"
    unfolding SA_poly_to_SA_fun_def using t_def a_def
    by (meson A restrict_apply')
  show "SA_poly_to_SA_fun n P x  carrier Qp"
    using assms t_closed a_def 0 UP_cring.to_fun_closed[of Qp "SA_poly_to_Qp_poly n a P" ]
    unfolding SA_poly_to_SA_fun_def
    using SA_poly_to_Qp_poly_closed a_closed UPQ.to_fun_closed by presburger
qed

lemma SA_poly_to_SA_fun_formula:
  assumes "P  carrier (UP (SA n))"
  assumes "x  carrier (Qpn)"
  assumes "t  carrier Qp"
  shows "SA_poly_to_SA_fun n P (t#x) = (SA_poly_to_Qp_poly n x P)t"
proof-
  have 0: "hd (t#x) = t"
    by simp
  have 1: "tl (t#x) = x"
    by auto
  have 2: "(t#x)  carrier (QpSuc n)"
    by (metis add.commute assms cartesian_power_cons plus_1_eq_Suc)
  show ?thesis
    unfolding SA_poly_to_SA_fun_def
    using 0 1 2 assms
    by (metis (no_types, lifting) restrict_apply')
qed

lemma semialg_map_comp_in_SA:
  assumes "f  carrier (SA n)"
  assumes "is_semialg_map m n g"
  shows "(λ a  carrier (Qpm). f (g a))  carrier (SA m)"
proof(rule SA_car_memI)
  show "(λacarrier (Qpm). f (g a))  carrier (Qp_funs m)"
  proof(rule Qp_funs_car_memI)
    show " (λacarrier (Qpm). f (g a))  carrier (Qpm)  carrier Qp"
    proof fix a assume A: "a  carrier (Qpm)"
      then have "g a  carrier (Qpn)"
        using is_semialg_map_def[of m n g] assms
        by blast
      then show  "f (g a)  carrier Qp"
        using A assms  SA_car_memE(3)[of f n]
        by blast
    qed
    show " x. x  carrier (Qpm)  (λacarrier (Qpm). f (g a)) x = undefined"
      unfolding restrict_def by metis
  qed
  have 0: "is_semialg_function m (f  g)"
    using assms semialg_function_comp_closed[of n f m g] SA_car_memE(1)[of f n]
    by blast
  have 1: " (a. a  carrier (Qpm)  (f  g) a = (λacarrier (Qpm). f (g a)) a)"
    using assms  comp_apply[of f g] unfolding restrict_def
    by metis
  then show "is_semialg_function m (λacarrier (Qpm). f (g a))"
    using 0 1 semialg_function_on_carrier'[of m "f  g" "λ a  carrier (Qpm). f (g a)" ]
    by blast
qed

lemma tl_comp_in_SA:
  assumes "f  carrier (SA n)"
  shows "(λ a  carrier (QpSuc n). f (tl a))  carrier (SA (Suc n))"
  using assms semialg_map_comp_in_SA[of f _ _  tl] tl_is_semialg_map[of "n"]
  by blast

lemma SA_poly_to_SA_fun_add_eval:
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  assumes "a  carrier (QpSuc n)"
  shows "SA_poly_to_SA_fun n (f UP (SA n)g) a = SA_poly_to_SA_fun n f a QpSA_poly_to_SA_fun n g a"
  unfolding SA_poly_to_SA_fun_def
  using assms SA_poly_to_Qp_poly_add[of "tl a" n f g]
  by (metis (no_types, lifting) SA_poly_to_Qp_poly_closed UPQ.to_fun_plus cartesian_power_head cartesian_power_tail restrict_apply')

lemma SA_poly_to_SA_fun_add:
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  shows "SA_poly_to_SA_fun n (f UP (SA n)g) = SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g"
proof fix x
  show " SA_poly_to_SA_fun n (f UP (SA n)g) x = (SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) x"
  proof(cases "x  carrier (QpSuc n)")
    case True
    then show ?thesis using SA_poly_to_SA_fun_add_eval[of f n g x] SA_add[of x n]
      using SA_mult assms(1) assms(2) SA_add
      by blast
  next
    case False
    have F0: "SA_poly_to_SA_fun n (f UP (SA n)g) x = undefined"
      unfolding SA_poly_to_SA_fun_def
      by (meson False restrict_apply)
    have F1: "(SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) x = undefined"
      using  False SA_add' by blast
    then show ?thesis
      using F0 by blast
  qed
qed

lemma SA_poly_to_SA_fun_monom:
  assumes "f  carrier (SA n)"
  assumes "a  carrier (QpSuc n)"
  shows "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = (f (tl a))(hd a)[^]Qpk "
proof-
  have "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = SA_poly_to_Qp_poly n (tl a) (up_ring.monom (UP (SA n)) f k)  lead_coeff a"
    unfolding SA_poly_to_SA_fun_def using assms
    by (meson restrict_apply)
  then have "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = up_ring.monom Qp_x (f (tl a)) k lead_coeff a"
    using SA_poly_to_Qp_poly_monom[of "tl a" n f k] assms
    by (metis cartesian_power_tail)
  then show ?thesis using assms
    by (metis (no_types, lifting) SA_car cartesian_power_head cartesian_power_tail UPQ.to_fun_monom Qp.function_ring_car_mem_closed semialg_functions_memE(2))
qed

lemma SA_poly_to_SA_fun_monom':
  assumes "f  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "t  carrier Qp"
  shows "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) (t#x) = (f x)t[^]Qpk "
proof-
  have 0: "hd (t#x) = t"
    by simp
  have 1: "tl (t#x) = x"
    by auto
  have 2: "(t#x)  carrier (QpSuc n)"
    by (metis add.commute assms  cartesian_power_cons plus_1_eq_Suc)
  show ?thesis
    using 0 1 2 SA_poly_to_SA_fun_monom[of f n "t#x" k] assms SA_poly_to_SA_fun_monom
    by presburger
qed

lemma hd_is_semialg_function:
  assumes "n > 0"
  shows "is_semialg_function n hd"
proof-
  have 0: "is_semialg_function n (λ a  carrier (Qpn). a!0)"
    using assms index_is_semialg_function restrict_is_semialg by blast
  have 1: "restrict (λacarrier (Qpn). a ! 0) (carrier (Qpn)) = restrict lead_coeff (carrier (Qpn))"
  proof fix x
    show "restrict (λacarrier (Qpn). a ! 0) (carrier (Qpn)) x = restrict lead_coeff (carrier (Qpn)) x"
      unfolding restrict_def
      apply(cases "x  carrier (Qpn)")
       apply (metis assms cartesian_power_car_memE drop_0 hd_drop_conv_nth)
      by presburger
  qed
  show ?thesis
    using 0 1 assms semialg_function_on_carrier[of n "(λ a  carrier (Qpn). a!0)" hd]
    by blast
qed

lemma SA_poly_to_SA_fun_monom_closed:
  assumes "f  carrier (SA n)"
  shows "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k)  carrier (SA (Suc n))"
proof-
  have 0: "is_semialg_function (Suc n) (f  tl)"
    using SA_imp_semialg assms(1)  semialg_function_comp_closed tl_is_semialg_map by blast
  obtain h where h_def: "h = restrict hd (carrier (QpSuc n))"
    by blast
  have h_closed: "h  carrier (SA (Suc n))"
    using hd_is_semialg_function SA_car h_def restrict_in_semialg_functions
    by blast
  have h_pow_closed: "h[^]SA (Suc n)k  carrier (SA (Suc n))"
    using assms(1) h_closed monoid.nat_pow_closed[of "SA (Suc n)" h k] SA_is_monoid[of "Suc n"]
    by blast
  have monom_term_closed: "(f  tl) SA (Suc n)h[^]SA (Suc n)k  carrier (SA (Suc n))"
    apply(rule SA_mult_closed_right)
    using "0" apply linarith
    using h_pow_closed by blast

  have 0: "a. a  carrier (QpSuc n)  SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = ((f  tl) SA (Suc n)h[^]SA (Suc n)k) a"
  proof- fix a assume "a  carrier (QpSuc n)"
    then show "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = ((f  tl) SA (Suc n)h[^]SA (Suc n)k) a"
      using assms SA_poly_to_SA_fun_monom[of f n a k] comp_apply[of f tl a] h_def restrict_apply
            SA_mult[of a "Suc n" "f  tl" "h [^]SA (Suc n)k"] SA_nat_pow[of a "Suc n" h k]
      by metis
  qed
  have 1: "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) = ((f  tl) SA (Suc n)h[^]SA (Suc n)k)"
  proof fix x
    show "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) x = ((f  tl) SA (Suc n)h [^]SA (Suc n)k) x"
      apply(cases "x  carrier (QpSuc n)")
      using "0" apply blast
      using monom_term_closed unfolding SA_poly_to_SA_fun_def
      using restrict_apply
      by (metis (no_types, lifting) SA_mult')
  qed
  show ?thesis
    using 1 monom_term_closed
    by presburger
qed

lemma SA_poly_to_SA_fun_is_SA:
  assumes "P  carrier (UP (SA n))"
  shows "SA_poly_to_SA_fun n P  carrier (SA (Suc n))"
  apply(rule UP_ring.poly_induct3[of "SA n" P])
  unfolding UP_ring_def using assms SA_is_ring apply blast
  using assms apply blast
  using assms SA_poly_to_SA_fun_add[of ]
  using SA_add_closed_right SA_imp_semialg  zero_less_Suc apply presburger
  using SA_poly_to_SA_fun_monom_closed assms(1)
  by blast

lemma SA_poly_to_SA_fun_mult:
  assumes "f  carrier (UP (SA n))"
  assumes "g  carrier (UP (SA n))"
  shows "SA_poly_to_SA_fun n (f UP (SA n)g) = SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g"
proof(rule function_ring_car_eqI[of _ "Suc n"])
  show "SA_poly_to_SA_fun n (f UP (SA n)g)  carrier (function_ring (carrier (QpSuc n)) Qp)"
  proof-
    have "f UP (SA n)g  carrier (UP (SA n))"
      using assms  SA_is_UP_cring
      by (meson cring.cring_simprules(5) padic_fields.UP_SA_n_is_cring padic_fields_axioms)
    thus ?thesis
      using SA_is_UP_cring assms SA_poly_to_SA_fun_is_SA[of "f UP (SA n)g"] SA_car_in_Qp_funs_car[of "Suc n"]
      by blast
  qed
  show "SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g  carrier (function_ring (carrier (QpSuc n)) Qp)"
    using SA_is_UP_cring SA_poly_to_SA_fun_is_SA assms
    by (meson SA_car_memE(2) SA_mult_closed_left less_Suc_eq_0_disj padic_fields.SA_car_memE(1) padic_fields_axioms)
  show "a. a  carrier (QpSuc n)  SA_poly_to_SA_fun n (f UP (SA n)g) a = (SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) a"
  proof- fix a assume  A: "a  carrier (QpSuc n)"
    then obtain t b where tb_def: "a = t#b"
      using cartesian_power_car_memE[of a Qp "Suc n"] by (meson length_Suc_conv)
    have t_closed: "t  carrier Qp"
      using tb_def A cartesian_power_head[of a Qp n] by (metis list.sel(1))
    have b_closed: "b  carrier (Qpn)"
      using tb_def A cartesian_power_tail[of a Qp n] by (metis list.sel(3))
    have 0: "f UP (SA n)g  carrier (UP (SA n))"
      using assms by (meson UP_SA_n_is_cring cring.cring_simprules(5))
    have 1: "SA_poly_to_SA_fun n (f UP (SA n)g) a  = (SA_poly_to_Qp_poly n b (f UP (SA n)g))t"
      using SA_poly_to_SA_fun_formula[of "f UP (SA n)g" n b t] t_closed b_closed tb_def  0  assms(1)
      by blast
    have 2: "(SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) a =
            SA_poly_to_SA_fun n f a  SA_poly_to_SA_fun n g a"
      using SA_poly_to_SA_fun_is_fun assms A SA_mult by blast
    hence 3: "(SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) a =
            ((SA_poly_to_Qp_poly n b f)  t)  ((SA_poly_to_Qp_poly n b g)  t)"
      using SA_poly_to_SA_fun_formula assms
      by (metis b_closed t_closed tb_def)
    have 4: "SA_poly_to_SA_fun n (f UP (SA n)g) a  = ((SA_poly_to_Qp_poly n b f)  t)  ((SA_poly_to_Qp_poly n b g)  t)"
      using 1 assms SA_poly_to_Qp_poly_closed[of b] SA_poly_to_Qp_poly_mult UPQ.to_fun_mult b_closed t_closed
      by presburger
    show " SA_poly_to_SA_fun n (f UP (SA n)g) a = (SA_poly_to_SA_fun n f SA (Suc n)SA_poly_to_SA_fun n g) a"
      using "3" "4" by blast
  qed
qed

lemma SA_poly_to_SA_fun_one:
  shows "SA_poly_to_SA_fun n (𝟭UP (SA n)) = 𝟭SA (Suc n)⇙"
proof(rule function_ring_car_eqI[of _ "Suc n"])
  have "𝟭UP (SA n) carrier (UP (SA n))"
    using UP_SA_n_is_cring cring.cring_simprules(6) by blast
  thus " SA_poly_to_SA_fun n 𝟭UP (SA n) carrier (function_ring (carrier (QpSuc n)) Qp)"
    using SA_poly_to_SA_fun_is_SA[of "𝟭UP (SA n)⇙"] SA_car_in_Qp_funs_car[of "Suc n"]  SA_is_UP_cring[of n]
    by blast
  show "𝟭SA (Suc n) carrier (function_ring (carrier (QpSuc n)) Qp)"
    unfolding SA_one
    using function_one_closed by blast
  show "a. a  carrier (QpSuc n)  SA_poly_to_SA_fun n 𝟭UP (SA n)a = 𝟭SA (Suc n)a"
  proof- fix a assume A: "  a  carrier (QpSuc n)"
    then obtain t b where tb_def: "a = t#b"
      using cartesian_power_car_memE[of a Qp "Suc n"] by (meson length_Suc_conv)
    have t_closed: "t  carrier Qp"
      using tb_def A cartesian_power_head[of a Qp n] by (metis list.sel(1))
    have b_closed: "b  carrier (Qpn)"
      using tb_def A cartesian_power_tail[of a Qp n] by (metis list.sel(3))
    have 0: "SA_poly_to_SA_fun n 𝟭UP (SA n)a  = (SA_poly_to_Qp_poly n b 𝟭UP (SA n))t"
      using SA_poly_to_SA_fun_formula 𝟭UP (SA n) carrier (UP (SA n))  b_closed t_closed tb_def
      by blast
    have 1: "SA_poly_to_Qp_poly n b 𝟭UP (SA n)= 𝟭UP Qp⇙"
      using  SA_poly_to_Qp_poly_is_hom[of b] ring_hom_one[of _ "UP (SA n)" "UP Qp"] b_closed
      by blast
    thus "SA_poly_to_SA_fun n 𝟭UP (SA n)a = 𝟭SA (Suc n)a"
      using "0" A function_one_eval SA_one UPQ.to_fun_one t_closed by presburger
  qed
qed

lemma SA_poly_to_SA_fun_ring_hom:
  shows "SA_poly_to_SA_fun n  ring_hom (UP (SA n)) (SA (Suc n))"
  apply(rule ring_hom_memI)
  using SA_poly_to_SA_fun_is_SA  apply blast
  apply (meson SA_poly_to_SA_fun_mult)
  apply (meson SA_poly_to_SA_fun_add)
  by (meson SA_poly_to_SA_fun_one)

lemma SA_poly_to_SA_fun_taylor_term:
  assumes "F  carrier (UP (SA n))"
  assumes "c  carrier (SA n)"
  assumes "x  carrier (Qpn)"
  assumes "t  carrier Qp"
  assumes "f = SA_poly_to_Qp_poly n x F"
  shows "SA_poly_to_SA_fun n (UP_cring.taylor_term (SA n) c F k) (t#x) = (taylor_expansion Qp (c x) f k) (t  c x)[^]Qpk"
proof-
  have 0: "hd (t#x) = t"
    by simp
  have 1: "tl (t#x) = x"
    by auto
  have 2: "(t#x)  carrier (QpSuc n)"
    by (metis Suc_eq_plus1 assms cartesian_power_cons)
  show ?thesis
    using assms 0 1 2 Pi_iff  SA_car_memE(3)
          SA_poly_to_Qp_poly_closed SA_poly_to_Qp_poly_comm_taylor_term[of F n c x k] restrict_apply'
    unfolding SA_poly_to_SA_fun_def
    by (metis (no_types, lifting) UPQ.taylor_def UPQ.to_fun_taylor_term)
qed

lemma SA_finsum_eval:
  assumes "finite I"
  assumes "F  I  carrier (SA m)"
  assumes "x  carrier (Qpm)"
  shows "(SA miI. F i) x = (iI. F i x)"
proof-
  have "F  I  carrier (SA m)  (SA miI. F i) x = (iI. F i x)"
    apply(rule finite.induct[of I])
    apply (simp add: assms(1))
    using  abelian_monoid.finsum_empty[of "SA m" F] assms unfolding Qp.finsum_empty
    using SA_is_abelian_monoid SA_zeroE apply presburger
  proof fix A a assume IH: "finite A" " F  A  carrier (SA m)  finsum (SA m) F A x = (iA. F i x)"
                            "F  insert a A  carrier (SA m)"
    then have 0: "F  A  carrier (SA m)"
      by blast
    then have 1: "finsum (SA m) F A x = (iA. F i x)"
      using IH by blast
    show "finsum (SA m) F (insert a A) x = (iinsert a A. F i x)"
    proof(cases "a  A")
      case True
      then show ?thesis using insert_absorb[of a A] IH
        by presburger
    next
      case False
      have F0: "(λi. F i x)  A  carrier Qp" proof fix i assume "i  A" thus "F i x  carrier Qp"
        using 0 assms(3) SA_car_memE(3)[of "F i" m] by blast qed
      have F1: "F a x  carrier Qp"
        using IH assms(3) SA_car_memE(3)[of "F a" m] by blast
      have F2: "finsum (SA m) F (insert a A) x = F a x  finsum (SA m) F A x"
      proof-
        have " finsum (SA m) F (insert a A) = F a SA mfinsum (SA m) F A"
          using False IH abelian_monoid.finsum_insert[of "SA m" A a F ] 0
          by (meson Pi_split_insert_domain SA_is_abelian_monoid)
        thus ?thesis using    abelian_monoid.finsum_closed[of "SA m" F A] 1 F0  F1
            SA_add[of x m "F a" "finsum (SA m) F A"]
          using assms(3) by presburger
      qed
      show ?thesis using False 0 IH  Qp.finsum_insert[of A a "λi. F i x"] unfolding F2 1
        using F0 F1 by blast
    qed
  qed
  thus ?thesis using assms by blast
qed

lemma(in ring) finsum_ring_hom:
  assumes "ring S"
  assumes "h  ring_hom R S"
  assumes "F  I  carrier R"
  assumes "finite I"
  shows "h (iI. F i) = (SiI. h (F i))"
proof-
  have "F  I  carrier R  h (iI. F i) = (SiI. h (F i))"
    apply(rule finite.induct[of I])
      apply (simp add: assms(4))
    unfolding finsum_empty using assms abelian_monoid.finsum_empty[of S]
    unfolding ring_def abelian_group_def
    apply (simp add: f. abelian_monoid S  finsum S f {} = 𝟬S⇙› assms(1) local.ring_axioms ring_hom_zero)
  proof fix A a assume A: "finite A" "  F  A  carrier R  h (finsum R F A) = (SiA. h (F i))"
       "   F  insert a A  carrier R"
    then have 0: "F  A  carrier R "
      by blast
    have 1: "h (finsum R F A) = (SiA. h (F i))"
      using "0" A(2) by linarith
    have 2: "(finsum R F A)  carrier R"
      using finsum_closed[of F A] 0 by blast
    have 3: "(SiA. h (F i))  carrier S"
    using assms 1 2 ring_hom_closed
    by metis
    have 4: "F a  carrier R" using A by blast
    have 5: "h (F a)  carrier S"
      using assms 4
      by (meson ring_hom_closed)
    show "h (finsum R F (insert a A)) = (Siinsert a A. h (F i))"
      apply(cases "a  A")
      using insert_absorb 1
       apply metis
    proof- assume C: "a A"
      have 6: "finsum R F (insert a A) = F a  finsum R F A"
      using A finsum_insert[of A a F]  C by blast
    have 7: "(Siinsert a A. h (F i)) = h (F a) S(SiA. h (F i))"
    apply(rule abelian_monoid.finsum_insert[of S A a "λi. h (F i)"])
          apply (simp add: abelian_group.axioms(1) assms(1) ring.is_abelian_group)
         apply (simp add: A(1))
        apply (simp add: C)
       apply(intro Pi_I ring_hom_closed[of h R S] assms)
      using 0 A 5 assms by auto
    thus ?thesis
    using    0 1 2 3 4 5 6  7 assms ring_hom_add[of h R S "F a" "finsum R F A" ]

      unfolding 1 ring_def abelian_group_def
      by presburger
  qed
  qed
  thus ?thesis using assms by auto
qed

lemma SA_poly_to_SA_fun_finsum:
  assumes "finite I"
  assumes "F  I  carrier (UP (SA m))"
  assumes "f = (UP (SA m)iI. F i)"
  assumes "x  carrier (QpSuc m)"
  shows "SA_poly_to_SA_fun m f x = (iI. SA_poly_to_SA_fun m (F i) x)"
proof-
  have "SA_poly_to_SA_fun m  ring_hom (UP (SA m)) (SA (Suc m))"
    using SA_poly_to_SA_fun_ring_hom by blast
  have f_closed: "f  carrier (UP (SA m))"
    unfolding assms apply(rule finsum_closed)  using assms by blast
  have 0: "SA_poly_to_SA_fun m f = (SA (Suc m)iI. SA_poly_to_SA_fun m (F i))"
    unfolding assms
    apply(rule finsum_ring_hom)
       apply (simp add: SA_is_ring)
    using SA_poly_to_SA_fun m  ring_hom (UP (SA m)) (SA (Suc m)) apply blast
    using assms(2) apply blast
    by (simp add: assms(1))
  show ?thesis unfolding 0 apply(rule SA_finsum_eval)
    using assms apply blast using assms
    apply (meson Pi_I Pi_mem padic_fields.SA_poly_to_SA_fun_is_SA padic_fields_axioms)
    using assms by blast
qed

lemma SA_poly_to_SA_fun_taylor_expansion:
  assumes "f  carrier (UP (SA m))"
  assumes "c  carrier (SA m)"
  assumes "x  carrier (QpSuc m)"
  shows "SA_poly_to_SA_fun m f x = (i{..deg (SA m) f}. taylor_expansion (SA m) c f i (tl x)  (hd x  c (tl x)) [^] i)"
proof-
  have 0: "f = (UP (SA m)i{..deg (SA m) f}. UP_cring.taylor_term (SA m) c f i)"
    using taylor_sum[of f m "deg (SA m) f" c] assms unfolding UPSA.taylor_term_def UPSA.taylor_def
    by blast
  have 1: "SA_poly_to_SA_fun m f x = (i{..deg (SA m) f}. SA_poly_to_SA_fun m (UP_cring.taylor_term (SA m) c f i) x)"
    apply(rule SA_poly_to_SA_fun_finsum)
       apply simp
      apply (meson Pi_I UPSA.taylor_term_closed assms(1) assms(2))
    using 0 apply blast
    using assms by blast
  have hd_closed: "hd x  carrier Qp"
    using assms cartesian_power_head by blast
  have tl_closed: "tl x  carrier (Qpm)"
    using assms cartesian_power_tail by blast
  obtain t a where ta_def: " x = t#a"
    using assms cartesian_power_car_memE[of x Qp "Suc m" ]
    by (metis Suc_length_conv)
  have 2: "t = hd x"
    by (simp add: ta_def)
  have 3: "a = tl x "
    by (simp add: ta_def)
  have 4: "i. SA_poly_to_SA_fun m (UPSA.taylor_term m c f i) x =
    taylor_expansion Qp (c (tl x)) (SA_poly_to_Qp_poly m (tl x) f) i  (lead_coeff x  c (tl x)) [^] i"
    using tl_closed hd_closed assms SA_poly_to_SA_fun_taylor_term[of f m c a t  "SA_poly_to_Qp_poly m a f"  ]
    unfolding 2 3  by (metis "2" "3" ta_def)
  have 5: "SA_poly_to_SA_fun m f x = (i{..deg (SA m) f}. (taylor_expansion Qp (c (tl x)) (SA_poly_to_Qp_poly m (tl x) f) i) ((hd x)  c (tl x))[^]Qpi)"
    using 1 2 unfolding 4 by blast
  have 6: "taylor_expansion Qp (c (tl x)) (SA_poly_to_Qp_poly m (tl x) f) = SA_poly_to_Qp_poly m (tl x) (taylor_expansion (SA m) c f)"

    using SA_poly_to_Qp_poly_taylor_poly[of f m c "tl x"] assms(1) assms(2) tl_closed by presburger
  have 7: "i. taylor_expansion Qp (c (tl x)) (SA_poly_to_Qp_poly m (tl x) f) i =
                taylor_expansion (SA m) c f i (tl x)"
    unfolding 6 using SA_poly_to_Qp_poly_coeff[of "tl x" m "taylor_expansion (SA m) c f"]
    by (metis UPSA.taylor_closed UPSA.taylor_def assms(1) assms(2) tl_closed)
  show ?thesis using 5 unfolding 7 by blast
qed

lemma SA_deg_one_eval:
  assumes "g  carrier (UP (SA m))"
  assumes "deg (SA m) g = 1"
  assumes "ξ  carrier (Fun⇘mQp)"
  assumes "UP_ring.lcf (SA m) g  Units (SA m)"
  assumes "x  carrier (Qpm). (SA_poly_to_SA_fun m g) (ξ x#x) = 𝟬"
  shows "ξ = SA m(g 0)SA m(invSA m(g 1))"
proof(rule ext)
  fix x show " ξ x = (SA mg 0 SA minvSA mg 1) x"
  proof(cases "x  carrier (Qpm)")
    case True
    then have "(SA_poly_to_SA_fun m g) (ξ x#x) = 𝟬"
      using assms by blast
    then have T0: "SA_poly_to_Qp_poly m x g  ξ x = 𝟬"
      using SA_poly_to_SA_fun_formula[of g m x "ξ x"] assms
            Qp.function_ring_car_mem_closed True by metis
    have "deg Qp (SA_poly_to_Qp_poly m x g) = 1"
    proof(rule UPQ.deg_eqI)
      show "SA_poly_to_Qp_poly m x g  carrier (UP Qp)"
        using assms True SA_poly_to_Qp_poly_closed by blast
      show "deg Qp (SA_poly_to_Qp_poly m x g)  1"
        using SA_poly_to_Qp_poly_deg_bound by (metis True assms(1) assms(2) )
      show " SA_poly_to_Qp_poly m x g 1  𝟬"
        using SA_poly_to_Qp_poly_coeff[of x m g 1] assms SA_Units_memE' True by presburger
    qed
    hence T1: "SA_poly_to_Qp_poly m x g  ξ x = SA_poly_to_Qp_poly m x g 0  SA_poly_to_Qp_poly m x g 1  (ξ x)"
      using UP_cring.deg_one_eval[of Qp _ "ξ x"]
      by (meson Qp.function_ring_car_mem_closed SA_poly_to_Qp_poly_closed True UPQ.deg_one_eval assms)
    hence T2: "g 0 x  g 1 x  (ξ x) = 𝟬"
      using True T0 assms SA_poly_to_Qp_poly_coeff[of x m g]
      by metis
    have T3: "g 0 x  carrier Qp"
      using True assms UP_ring.cfs_closed
      by (metis SA_poly_to_Qp_poly_closed SA_poly_to_Qp_poly_coeff UPQ.is_UP_ring)
    have T4: "ξ x  carrier Qp"
      using True assms Qp.function_ring_car_memE by blast
    have T5: "g 1 x  nonzero Qp"
      using True assms
      by (metis Qp.function_ring_car_memE SA_Units_closed SA_Units_memE' SA_car_memE(2) not_nonzero_Qp)
    have T6: " (g 0 x) = g 1 x  (ξ x)"
      using T2 T3 T4 T5 by (metis Qp.m_closed Qp.minus_equality Qp.minus_minus Qp.nonzero_closed)
    hence T7: " (g 0 x)  inv (g 1 x)= ξ x"
      using T5 by (metis Qp.inv_cancelR(2) Qp.m_closed Qp.nonzero_closed T4 Units_eq_nonzero)
    have T8: "(invSA mg 1) x = inv (g 1 x)"
      using assms True Qp_funs_m_inv SA_Units_Qp_funs_Units SA_Units_Qp_funs_inv by presburger
    have T9: "(SA mg 0) x =  (g 0 x)"
      using SA_a_inv_eval[of "g 0" m x] UP_ring.cfs_closed[of "SA m" g 0] assms True SA_is_ring
      unfolding UP_ring_def by blast
    have T11: "((SA mg 0) SA minvSA mg 1) x =  (g 0 x)  inv (g 1 x)"
      using assms UP_ring.cfs_closed T8 T9 T7 True SA_mult by presburger
    thus "ξ x = (SA mg 0 SA minvSA mg 1) x"
      using T7 by blast
  next
    case False
    then show ?thesis
      using SA_mult' SA_times assms function_ring_not_car by auto
  qed
qed

lemma SA_deg_one_eval':
  assumes "g  carrier (UP (SA m))"
  assumes "deg (SA m) g = 1"
  assumes "ξ  carrier (Fun⇘mQp)"
  assumes "UP_ring.lcf (SA m) g  Units (SA m)"
  assumes "x  carrier (Qpm). (SA_poly_to_SA_fun m g) (ξ x#x) = 𝟬"
  shows "ξ  carrier (SA m)"
proof-
  have 0: "ξ = SA m(g 0)SA m(invSA m(g 1))"
    using assms SA_deg_one_eval by blast
  have 1: "(invSA m(g 1))  carrier (SA m)"
    using assms SA_Units_inv_closed by presburger
  have "(g 0)  carrier (SA m)"
    using assms(1) assms(2) UP_ring.cfs_closed[of "SA m" g 0] SA_is_ring[of m ] unfolding UP_ring_def
    by blast
  hence 2: "SA m(g 0)  carrier (SA m)"
    by (meson SA_is_cring assms(1) cring.cring_simprules(3))
  show ?thesis
    using 0 1 2  SA_imp_semialg SA_mult_closed_left assms(1) by blast
qed

lemma Qp_pow_ConsI:
  assumes "t  carrier Qp"
  assumes "x  carrier (Qpm)"
  shows "t#x  carrier (QpSuc m)"
  using assms cartesian_power_cons[of x Qp m t] Suc_eq_plus1 by presburger

lemma Qp_pow_ConsE:
  assumes "x  carrier (QpSuc m)"
  shows "tl x  carrier (Qpm)"
        "hd x  carrier Qp"
  using assms cartesian_power_tail apply blast
  using assms cartesian_power_head by blast

lemma(in ring) add_monoid_one:
"𝟭add_monoid R= 𝟬"
  by simp

lemma(in ring) add_monoid_carrier:
"carrier (add_monoid R) = carrier R"
  unfolding Congruence.partial_object.simps(1)
  by simp

lemma(in ring) finsum_mono_neutral_cong:
  assumes "F  I  carrier R"
  assumes "finite I"
  assumes "i. i  J  F i = 𝟬"
  assumes "J  I"
  shows "finsum R F I = finsum R F J"
  unfolding finsum_def apply(rule comm_monoid.finprod_mono_neutral_cong)
  using local.add.comm_monoid_axioms apply blast
  using assms(2) assms(4) rev_finite_subset apply blast
  apply (simp add: assms(2))
  using assms(4) apply blast
  unfolding add_monoid_one using assms apply blast
  apply blast
  using add_monoid_carrier assms(1) apply blast
  using add_monoid_carrier assms  by blast

text‹
  This lemma helps to formalize statements like "by passing to a partition, we can assume the
  Taylor coefficients are either always zero or never zero"›

lemma SA_poly_to_SA_fun_taylor_on_refined_set:
  assumes "f  carrier (UP (SA m))"
  assumes "c  carrier (SA m)"
  assumes "is_semialgebraic m A"
  assumes "i. A  SA_zero_set m (taylor_expansion (SA m) c f i)  A  SA_nonzero_set m (taylor_expansion (SA m) c f i)"
  assumes "a = to_fun_unit m  taylor_expansion (SA m) c f"
  assumes "inds = {i. i  deg (SA m) f  A  SA_nonzero_set m (taylor_expansion (SA m) c f i)}"
  assumes "x  A"
  assumes "t  carrier Qp"
  shows "SA_poly_to_SA_fun m f (t#x) = (iinds. (a i x)(t  c x)[^]i)"
proof-
  have x_closed: "x  carrier (Qpm)"
    using assms(3) assms(7)
  by (metis (no_types, lifting) Diff_eq_empty_iff Diff_iff empty_iff is_semialgebraic_closed)
  have tx_closed: "t#x  carrier (QpSuc m)"
    using x_closed assms(8) cartesian_power_cons[of x Qp m t] Qp_pow_ConsI by blast
  have "SA_poly_to_SA_fun m f (t # x) =
    (i{..deg (SA m) f}. taylor_expansion (SA m) c f i (tl (t # x))  (hd (t # x)  c (tl (t # x))) [^] i)"
    using tx_closed assms SA_poly_to_SA_fun_taylor_expansion[of f m c "t#x"]
    by linarith
  then have 0: "SA_poly_to_SA_fun m f (t#x) = (i{..deg (SA m) f}. taylor_expansion (SA m) c f i x  (t  c x) [^] i)"
    unfolding list_tl list_hd by blast
  have 1: "i. i  inds  taylor_expansion (SA m) c f i x = 𝟬"
  proof- fix i assume A: "i  inds"
    show "taylor_expansion (SA m) c f i x = 𝟬"
    proof(cases "i  deg (SA m) f")
    case True
    then have "A  {x  carrier (Qpm). taylor_expansion (SA m) c f i x = 𝟬}"
      using A assms(8) assms(4)[of i] unfolding assms mem_Collect_eq SA_zero_set_def
      by linarith
    thus ?thesis using x_closed assms by blast
   next
    case False
    then have "taylor_expansion (SA m) c f i = 𝟬SA m⇙"
      using assms taylor_deg[of c m f] unfolding UPSA.taylor_def
      by (metis (no_types, lifting) UPSA.taylor_closed UPSA.taylor_def UPSA.deg_eqI nat_le_linear)
    then show ?thesis
      using x_closed SA_car_memE SA_zeroE by presburger
    qed
  qed
  hence 2: "i. i  inds  taylor_expansion (SA m) c f i x  (t  c x) [^] i = 𝟬"
    using assms x_closed  SA_car_memE
    by (metis (no_types, lifting) Qp.cring_simprules(26) Qp.function_ring_car_memE Qp.minus_closed Qp.nat_pow_closed)
  have 3: "(λi. taylor_expansion (SA m) c f i x  (t  c x) [^] i)  {..deg (SA m) f}  carrier Qp"
  proof fix i assume A: "i  {..deg (SA m) f}"
    have 30: "taylor_expansion (SA m) c f i  carrier (SA m)"
      using  assms UPSA.taylor_closed[of f m c] unfolding UPSA.taylor_def
      using UPSA.UP_car_memE(1) by blast
    hence 31: "taylor_expansion (SA m) c f i x  carrier Qp"
      using x_closed function_ring_car_closed SA_car_memE(2) by blast
    have 32: "c x  carrier Qp"
      using assms x_closed SA_car_memE(3) by blast
    hence 33: "(t  c x)[^] i  carrier Qp"
      using assms by blast
    show "taylor_expansion (SA m) c f i x  (t  c x) [^] i  carrier Qp"
      using 30 31 32 33 by blast
  qed
  have 4: "SA_poly_to_SA_fun m f (t#x) = (iinds. taylor_expansion (SA m) c f i x  (t  c x) [^] i)"
    unfolding 0 apply(rule Qp.finsum_mono_neutral_cong)
    using assms UPSA.taylor_closed[of f m c] unfolding UPSA.taylor_def
    using "3" apply blast
      apply blast
    using 2 apply blast
    unfolding assms by blast
  have A: "i. i  inds  a i x = taylor_expansion (SA m) c f i x"
    unfolding assms mem_Collect_eq SA_nonzero_set_def comp_apply
    apply(rule to_fun_unit_eq[of _ m x])
    using UPSA.taylor_closed[of f m c] assms unfolding UPSA.taylor_def
    using UPSA.UP_car_memE(1) apply blast
    using x_closed apply blast
    using assms(7) by blast
  have a_closed: "a  UNIV  carrier (SA m)"
    apply(rule Pi_I) unfolding assms comp_apply apply(rule to_fun_unit_closed[of _ m])
    apply(rule cfs_closed) using assms UPSA.taylor_closed[of f m c] unfolding UPSA.taylor_def by blast
  have 5: "(λi. a i x  (t  c x) [^] i)  inds  carrier Qp"
  proof fix i assume A: "i  inds"
    show "a i x  (t  c x) [^] i  carrier Qp"
    using assms(8) x_closed a_closed SA_car_memE(3)[of c m] SA_car_memE(3)[of "a i" m]
         assms(2) by blast
  qed
  have 6: "i. i  inds  taylor_expansion (SA m) c f i x  (t  c x) [^] i = a i x  (t  c x) [^] i"
    unfolding A by blast
  show ?thesis
    unfolding 4 apply(rule Qp.finsum_cong') apply blast
    using 5 apply blast
    using 6 by blast
qed

lemma SA_poly_to_Qp_poly_taylor_cfs:
  assumes "f  carrier (UP (SA m))"
  assumes "x  carrier (Qpm)"
  assumes "c  carrier (SA m)"
  shows "taylor_expansion (SA m) c f i x =
    taylor_expansion Qp (c x) (SA_poly_to_Qp_poly m x f) i"
proof-
  have 0: "SA_poly_to_Qp_poly m x (taylor_expansion (SA m) c f) =
         taylor_expansion Qp (c x) (SA_poly_to_Qp_poly m x f)"
    using SA_poly_to_Qp_poly_taylor_poly[of f m c x] assms by blast
  hence 1: "SA_poly_to_Qp_poly m x (taylor_expansion (SA m) c f) i =
  taylor_expansion Qp (c x) (SA_poly_to_Qp_poly m x f) i"
    by presburger
  thus ?thesis
    using assms SA_poly_to_Qp_poly_coeff[of x m "taylor_expansion (SA m) c f" i]
      UPSA.taylor_closed[of f m c] unfolding UPSA.taylor_def assms by blast
qed

(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Common Morphisms on Polynomial Rings›
(**************************************************************************************************)
(**************************************************************************************************)


text‹Evaluation homomorphism from multivariable polynomials to semialgebraic functions›

definition Qp_ev_hom where
"Qp_ev_hom n P = restrict (Qp_ev P) (carrier (Qpn))"

lemma Qp_ev_hom_ev:
  assumes "a  carrier (Qpn)"
  shows "Qp_ev_hom n P a = Qp_ev P a"
  using assms unfolding Qp_ev_hom_def
  by (meson restrict_apply')

lemma Qp_ev_hom_closed:
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  shows "Qp_ev_hom n f  carrier (Qpn)  carrier Qp"
  using Qp_ev_hom_ev assms by (metis Pi_I eval_at_point_closed)

lemma Qp_ev_hom_is_semialg_function:
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  shows "is_semialg_function n (Qp_ev_hom n f)"
  unfolding Qp_ev_hom_def
  using assms poly_is_semialg[of  f] restrict_is_semialg by blast

lemma Qp_ev_hom_closed':
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  shows "Qp_ev_hom n f  carrier (Fun⇘nQp)"
  apply(rule Qp.function_ring_car_memI)
  using  Qp_ev_hom_closed[of f n] assms apply blast
  unfolding Qp_ev_hom_def using assms by (meson restrict_apply)

lemma Qp_ev_hom_in_SA:
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  shows "Qp_ev_hom n f  carrier (SA n)"
  apply(rule SA_car_memI)
  using Qp_ev_hom_closed' assms(1)  apply blast
  using Qp_ev_hom_is_semialg_function assms(1) by blast

lemma Qp_ev_hom_add:
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  assumes "g  carrier (Qp[𝒳⇘n⇙])"
  shows "Qp_ev_hom n (f Qp[𝒳⇘n⇙]⇙ g) = (Qp_ev_hom n f) SA n(Qp_ev_hom n g)"
  apply(rule function_ring_car_eqI[of _ n])
  using assms MP.add.m_closed Qp_ev_hom_closed' apply blast
  using assms Qp_ev_hom_closed' fun_add_closed SA_plus apply presburger
proof- fix a assume A: "a  carrier (Qpn)"
  have " Qp_ev_hom n (f Qp [𝒳⇘n⇙]⇙ g) a =  Qp_ev_hom n f a  Qp_ev_hom n g a"
    using A Qp_ev_hom_ev assms eval_at_point_add by presburger
  then show "Qp_ev_hom n (f Qp [𝒳⇘n⇙]⇙ g) a = (Qp_ev_hom n f SA nQp_ev_hom n g) a"
    using A SA_add by blast
qed

lemma Qp_ev_hom_mult:
  assumes "f  carrier (Qp[𝒳⇘n⇙])"
  assumes "g  carrier (Qp[𝒳⇘n⇙])"
  shows "Qp_ev_hom n (f Qp[𝒳⇘n⇙]⇙ g) = (Qp_ev_hom n f) SA n(Qp_ev_hom n g)"
  apply(rule function_ring_car_eqI[of _ n])
  using assms MP.m_closed Qp_ev_hom_closed' apply blast
  using assms Qp_ev_hom_closed' fun_mult_closed SA_mult
  apply (meson Qp_ev_hom_in_SA SA_car_memE(2) SA_imp_semialg SA_mult_closed)
proof- fix a assume A: "a  carrier (Qpn)"
  have " Qp_ev_hom n (f Qp [𝒳⇘n⇙]⇙ g) a =  Qp_ev_hom n f a  Qp_ev_hom n g a"
    using A Qp_ev_hom_ev assms eval_at_point_mult by presburger
  then show "Qp_ev_hom n (f Qp [𝒳⇘n⇙]⇙ g) a = (Qp_ev_hom n f SA nQp_ev_hom n g) a"
    using A SA_mult by blast
qed

lemma Qp_ev_hom_one:
  shows "Qp_ev_hom n 𝟭Qp[𝒳⇘n⇙]⇙ = 𝟭SA n⇙"
  apply(rule function_ring_car_eqI[of _ n])
  using Qp_ev_hom_closed' apply blast
  using  function_one_closed SA_one apply presburger
  unfolding Qp_ev_hom_def
  using function_one_eval Qp_ev_hom_def Qp_ev_hom_ev Qp_ev_one SA_one by presburger

lemma Qp_ev_hom_is_hom:
  shows "Qp_ev_hom n  ring_hom (Qp[𝒳⇘n⇙]) (SA n)"
  apply(rule ring_hom_memI)
  using   Qp_ev_hom_in_SA apply blast
    using  Qp_ev_hom_mult apply blast
      using  Qp_ev_hom_add apply blast
        using  Qp_ev_hom_one by blast

lemma Qp_ev_hom_constant:
  assumes "c  carrier Qp"
  shows "Qp_ev_hom n (Qp.indexed_const c) = 𝔠⇘nc"
  apply(rule function_ring_car_eqI[of _ n])
  using Qp_ev_hom_closed' Qp_to_IP_car assms(1)  apply blast
  using constant_function_closed assms  apply blast
  by (metis Qp_constE Qp_ev_hom_ev assms eval_at_point_const)

notation  Qp.variable ("𝔳⇘_, _")

lemma Qp_ev_hom_pvar:
  assumes "i < n"
  shows "Qp_ev_hom n (pvar Qp i) = 𝔳⇘n, i⇙"
  apply(rule function_ring_car_eqI[of _ n])
  using assms Qp_ev_hom_closed' local.pvar_closed apply blast
  using Qp.var_in_car assms apply blast
  unfolding Qp_ev_hom_def var_def using assms eval_pvar
  by (metis (no_types, lifting) restrict_apply)

definition ext_hd where
"ext_hd m = (λ x  carrier (Qpm). hd x)"

lemma hd_zeroth:
"length x > 0  x!0 = hd x"
  apply(induction x)
  apply simp
  by simp

lemma ext_hd_pvar:
  assumes "m > 0"
  shows "ext_hd m = (λx  carrier (Qpm). eval_at_point Qp x (pvar Qp 0) )"
  unfolding ext_hd_def restrict_def using assms eval_pvar[of 0 m]
  using hd_zeroth
  by (metis (no_types, opaque_lifting) cartesian_power_car_memE)

lemma ext_hd_closed:
  assumes "m > 0"
  shows "ext_hd m  carrier (SA m)"
  using ext_hd_pvar[of m] assms pvar_closed[of 0 m] Qp_ev_hom_def Qp_ev_hom_in_SA by presburger

lemma UP_Qp_poly_to_UP_SA_is_hom:
  shows "poly_lift_hom (Qp[𝒳⇘n⇙]) (SA n) (Qp_ev_hom n)  ring_hom (UP (Qp[𝒳⇘n⇙])) (UP (SA n))"
  using UP_cring.poly_lift_hom_is_hom[of "Qp[𝒳⇘n⇙]" "SA n" "Qp_ev_hom n"]
  unfolding UP_cring_def
  using Qp_ev_hom_is_hom SA_is_cring  coord_cring_cring by blast

definition coord_ring_to_UP_SA where
"coord_ring_to_UP_SA n = poly_lift_hom (Qp[𝒳⇘n⇙]) (SA n) (Qp_ev_hom n)  to_univ_poly (Suc n) 0"

lemma coord_ring_to_UP_SA_is_hom:
  shows "coord_ring_to_UP_SA n  ring_hom (Qp[𝒳⇘Suc n⇙]) (UP (SA n))"
  unfolding coord_ring_to_UP_SA_def
  using UP_Qp_poly_to_UP_SA_is_hom[of n] to_univ_poly_is_hom[of 0 "n"] ring_hom_trans
  by blast

lemma coord_ring_to_UP_SA_add:
  assumes "f  carrier (Qp[𝒳⇘Suc n⇙])"
  assumes "g  carrier (Qp[𝒳⇘Suc n⇙])"
  shows "coord_ring_to_UP_SA n (f Qp[𝒳⇘Suc n⇙]⇙g) = coord_ring_to_UP_SA n f  UP (SA n)coord_ring_to_UP_SA n g"
  using assms coord_ring_to_UP_SA_is_hom ring_hom_add
  by (metis (mono_tags, opaque_lifting))

lemma coord_ring_to_UP_SA_mult:
  assumes "f  carrier (Qp[𝒳⇘Suc n⇙])"
  assumes "g  carrier (Qp[𝒳⇘Suc n⇙])"
  shows "coord_ring_to_UP_SA n (f Qp[𝒳⇘Suc n⇙]⇙g) = coord_ring_to_UP_SA n f  UP (SA n)coord_ring_to_UP_SA n g"
  using assms coord_ring_to_UP_SA_is_hom ring_hom_mult
  by (metis (no_types, opaque_lifting))

lemma coord_ring_to_UP_SA_one:
  shows "coord_ring_to_UP_SA n 𝟭Qp[𝒳⇘Suc n⇙]⇙ = 𝟭UP (SA n)⇙"
  using coord_ring_to_UP_SA_is_hom ring_hom_one
  by blast

lemma coord_ring_to_UP_SA_closed:
  assumes "f  carrier (Qp[𝒳⇘Suc n⇙])"
  shows "coord_ring_to_UP_SA n f  carrier (UP (SA n))"
  using assms coord_ring_to_UP_SA_is_hom ring_hom_closed
  by (metis (no_types, opaque_lifting))

lemma coord_ring_to_UP_SA_constant:
  assumes "c  carrier Qp"
  shows "coord_ring_to_UP_SA n (Qp.indexed_const c) = to_polynomial (SA n) (𝔠⇘nc)"
proof-
  have 0: "pre_to_univ_poly (Suc n) 0 (Qp.indexed_const c) = ring.indexed_const (Qp[𝒳⇘n⇙]) (Qp.indexed_const c)"
    unfolding to_univ_poly_def
    using pre_to_univ_poly_is_hom(5)[of 0 "Suc n" "pre_to_univ_poly (Suc n) 0" c] assms unfolding coord_ring_def
    using diff_Suc_1 zero_less_Suc by presburger
  hence 1: "to_univ_poly (Suc n) 0 (Qp.indexed_const c) = to_polynomial (Qp[𝒳⇘n⇙]) (Qp.indexed_const c) "
    unfolding to_univ_poly_def
    using UP_cring.IP_to_UP_indexed_const[of "Qp[𝒳⇘n⇙]" "Qp.indexed_const c" "0::nat"] assms
          comp_apply[of "IP_to_UP (0::nat)" "pre_to_univ_poly (Suc n) (0::nat)" "Qp.indexed_const c"]
    unfolding UP_cring_def using Qp_to_IP_car coord_cring_cring by presburger
  have 2: "Qp_ev_hom n (Qp.indexed_const c) = 𝔠⇘nc"
    using Qp_ev_hom_constant[of c n] assms by blast
  have 3: "poly_lift_hom (Qp [𝒳⇘n⇙]) (SA n) (Qp_ev_hom n) (to_polynomial (Qp [𝒳⇘n⇙]) (Qp.indexed_const c)) = to_polynomial (SA n) (Qp_ev_hom n (Qp.indexed_const c))"
    using UP_cring.poly_lift_hom_extends_hom[of "Qp[𝒳⇘n⇙]" "SA n" "Qp_ev_hom n" "Qp.indexed_const c"]
    unfolding UP_cring_def coord_ring_def coord_ring_to_UP_SA_def
    by (metis Qp_ev_hom_is_hom Qp_to_IP_car SA_is_cring assms(1) coord_cring_cring coord_ring_def)
  hence 4: "poly_lift_hom (Qp [𝒳⇘n⇙]) (SA n) (Qp_ev_hom n) (to_univ_poly (Suc n) 0 (Qp.indexed_const c) ) = to_polynomial (SA n) (𝔠⇘nc)"
    using "1" "2" by presburger
  show ?thesis
    using assms 4 Qp.indexed_const_closed[of c "{..<n}"]
          comp_apply[of "poly_lift_hom (Pring Qp {..<n}) (SA n) (Qp_ev_hom n)" "to_univ_poly (Suc n) 0" "Qp.indexed_const c"]
    unfolding coord_ring_to_UP_SA_def
    by (metis coord_ring_def)
qed

lemma coord_ring_to_UP_SA_pvar_0:
  shows "coord_ring_to_UP_SA n (pvar Qp 0) = up_ring.monom (UP (SA n)) 𝟭SA n1"
proof-
  have 0: "pre_to_univ_poly (Suc n) 0 (pvar Qp 0) = pvar (Qp[𝒳⇘n⇙]) 0"
    using pre_to_univ_poly_is_hom(3)[of 0 "Suc n" "pre_to_univ_poly (Suc n) 0"] diff_Suc_1 zero_less_Suc
    by presburger
  have 1: "IP_to_UP (0::nat) (pvar (Qp[𝒳⇘n⇙]) 0) = up_ring.monom (UP (Qp[𝒳⇘n⇙])) 𝟭Qp[𝒳⇘n⇙]⇙ 1"
    using cring.IP_to_UP_var[of "Qp[𝒳⇘n⇙]" "0::nat"] unfolding X_poly_def var_to_IP_def
    using coord_cring_cring by blast
  have 2: "to_univ_poly (Suc n) 0 (pvar Qp 0) = up_ring.monom (UP (Qp[𝒳⇘n⇙])) 𝟭Qp[𝒳⇘n⇙]⇙ 1"
    unfolding to_univ_poly_def using 0 1 comp_apply
    by metis
  have 3: "poly_lift_hom (Qp[𝒳⇘n⇙]) (SA n) (Qp_ev_hom n) (up_ring.monom (UP (Qp[𝒳⇘n⇙])) 𝟭Qp[𝒳⇘n⇙]⇙ 1)
             = up_ring.monom (UP (SA n)) 𝟭SA n1"
    using UP_cring.poly_lift_hom_monom[of "Qp[𝒳⇘n⇙]" "SA n" "Qp_ev_hom n"]
    unfolding UP_cring_def
    using MP.one_closed Qp_ev_hom_is_hom Qp_ev_hom_one SA_is_cring coord_cring_cring by presburger
  thus ?thesis unfolding coord_ring_to_UP_SA_def
    using 2 3  comp_apply
    by metis
qed

lemma coord_ring_to_UP_SA_pvar_Suc:
  assumes "i > 0"
  assumes "i < Suc n"
  shows "coord_ring_to_UP_SA n (pvar Qp i) = to_polynomial (SA n) (𝔳⇘n, i-1)"
proof-
  have 0: "pre_to_univ_poly (Suc n) 0 (pvar Qp i) = ring.indexed_const (Qp[𝒳⇘n⇙]) (pvar Qp (i-1))"
    using pre_to_univ_poly_is_hom(4)[of 0 "Suc n" "pre_to_univ_poly (Suc n) 0" i] diff_Suc_1 zero_less_Suc
          assms
    unfolding coord_ring_def by presburger
  have 1: "IP_to_UP (0::nat) (ring.indexed_const (Qp[𝒳⇘n⇙]) (pvar Qp (i-1))) = to_polynomial (Qp[𝒳⇘n⇙]) (pvar Qp (i-1))"
    using UP_cring.IP_to_UP_indexed_const[of "Qp[𝒳⇘n⇙]" "pvar Qp (i-1)" "0::nat"] coord_cring_cring
    unfolding UP_cring_def
    by (metis assms diff_less less_Suc_eq less_imp_diff_less less_numeral_extra(1) local.pvar_closed)
  have 2: "to_univ_poly (Suc n) 0 (pvar Qp i) = to_polynomial (Qp[𝒳⇘n⇙]) (pvar Qp (i-1))"
    unfolding to_univ_poly_def using 0 1 comp_apply
    by metis
  have 3: "poly_lift_hom (Qp[𝒳⇘n⇙]) (SA n) (Qp_ev_hom n) (to_polynomial (Qp[𝒳⇘n⇙]) (pvar Qp (i-1)))
             = to_polynomial (SA n) (𝔳⇘n, (i-1))"
    using UP_cring.poly_lift_hom_extends_hom[of "Qp[𝒳⇘n⇙]" "SA n" "Qp_ev_hom n" "pvar Qp (i-1)"]
    unfolding UP_cring_def
    by (metis (no_types, lifting) Qp_ev_hom_is_hom Qp_ev_hom_pvar SA_is_cring Suc_diff_1
        Suc_less_eq assms  coord_cring_cring local.pvar_closed)
  thus ?thesis unfolding coord_ring_to_UP_SA_def
    using 2 3 assms comp_apply
    by metis
qed

lemma coord_ring_to_UP_SA_eval:
  assumes "f  carrier (Qp[𝒳⇘Suc n⇙])"
  assumes "a  carrier (Qpn)"
  assumes "t  carrier Qp"
  shows "Qp_ev f (t#a) = ((SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n f)))  t"
proof(rule coord_ring_car_induct[of f "Suc n"])
  have ta_closed: "t # a  carrier (QpSuc n)"
    using assms cartesian_power_cons  by (metis Suc_eq_plus1)
  show "f  carrier (Qp [𝒳⇘Suc n⇙])"
    by (simp add: assms)
  show "c. c  carrier Qp  eval_at_point Qp (t # a) (Qp.indexed_const c) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (Qp.indexed_const c))  t"
  proof- fix c assume A: "c  carrier Qp"
    have 0: "eval_at_point Qp (t # a) (Qp.indexed_const c) = c"
      using A eval_at_point_const ta_closed by blast
    have 1: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (Qp.indexed_const c)) = to_polynomial Qp c"
      using coord_ring_to_UP_SA_constant[of c n]  A Qp_constE SA_car
            SA_poly_to_Qp_poly_extends_apply assms  constant_function_in_semialg_functions
      by presburger
    show "eval_at_point Qp (t # a) (Qp.indexed_const c) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (Qp.indexed_const c))  t"
      using 0 1 A UPQ.to_fun_to_poly assms by presburger
  qed
  show " p q. p  carrier (Qp [𝒳⇘Suc n⇙]) 
           q  carrier (Qp [𝒳⇘Suc n⇙]) 
           eval_at_point Qp (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p)  t 
           eval_at_point Qp (t # a) q = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n q)  t 
           eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ q) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ q))  t"
  proof- fix p q assume A: "p  carrier (Qp [𝒳⇘Suc n⇙])" "q  carrier (Qp [𝒳⇘Suc n⇙])"
           "eval_at_point Qp (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p)  t"
           "eval_at_point Qp (t # a) q = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n q)  t"
    have 0: "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ q) = eval_at_point Qp (t # a) p  eval_at_point Qp (t # a) q"
      using A(1) A(2) eval_at_point_add ta_closed by blast
    have 1: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ q))=
    SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p)  UP QpSA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n q)"
      using coord_ring_to_UP_SA_add[of ] assms coord_ring_to_UP_SA_closed A
            SA_poly_to_Qp_poly_add[of  a n "coord_ring_to_UP_SA n p" "coord_ring_to_UP_SA n q"]
      by presburger
    show "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ q) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ q))  t"
      using 0 1  assms SA_poly_to_Qp_poly_closed[of a n] SA_poly_to_Qp_poly_closed A(1) A(2) A(3) A(4) UPQ.to_fun_plus coord_ring_to_UP_SA_closed
      by presburger
  qed
  show "p i. p  carrier (Qp [𝒳⇘Suc n⇙]) 
           i < Suc n 
           eval_at_point Qp (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p)  t 
           eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i))  t"
  proof- fix p i assume A: "p  carrier (Qp [𝒳⇘Suc n⇙])" "i < Suc n"
           "eval_at_point Qp (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p)  t"
    show " eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i))  t"
    proof-
      have 0: "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = eval_at_point Qp (t # a) p   eval_at_point Qp (t # a) (pvar Qp i)"
        using A(1) A(2) eval_at_point_mult local.pvar_closed ta_closed by blast
      have 1: "coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = coord_ring_to_UP_SA n p UP (SA n)coord_ring_to_UP_SA n (pvar Qp i)"
        using A(1) A(2) assms(1) coord_ring_to_UP_SA_mult local.pvar_closed by blast
      hence 2: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i)) =
              SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) UP QpSA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (pvar Qp i))"
        using SA_poly_to_Qp_poly_mult coord_ring_to_UP_SA_closed A(1) A(2) assms local.pvar_closed
        by presburger
      show "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i))  t"
      proof(cases  "i = 0")
        case True
        have T0: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i)) =
              SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) UP Qpup_ring.monom (UP Qp) 𝟭 1"
          using True coord_ring_to_UP_SA_pvar_0 SA_poly_to_Qp_poly_monom
          by (metis "2" function_one_eval Qp.one_closed SA_car SA_one assms constant_function_in_semialg_functions function_one_as_constant)
        have T1: "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = eval_at_point Qp (t # a) p   t"
          using 0 True ta_closed eval_pvar[of 0 "Suc n" "t#a"]
          by (metis A(2) nth_Cons_0)
        then show ?thesis
          using T0 A SA_poly_to_Qp_poly_closed[of a n  "coord_ring_to_UP_SA n p"] UPQ.to_fun_X[of t] to_fun_mult
                coord_ring_to_UP_SA_closed[of  ] UPQ.X_closed[of ] unfolding X_poly_def
          using assms UPQ.to_fun_mult by presburger
      next
        case False
        have "𝔳⇘n, i - 1a = a!(i-1)"
          by (metis A(2) False Qp.varE Suc_diff_1 Suc_less_eq assms  less_Suc_eq_0_disj)
        hence F0: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i)) =
              SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) UP Qpto_polynomial Qp (a!(i-1))"
          using False coord_ring_to_UP_SA_pvar_Suc[of i n] SA_poly_to_Qp_poly_extends_apply[of a n "𝔳⇘n, i - 1⇙"]
          by (metis (no_types, lifting) "2" A(2) Qp_ev_hom_in_SA Qp_ev_hom_pvar Suc_diff_1 Suc_less_eq assms  local.pvar_closed neq0_conv)
        have F1: "eval_at_point Qp (t # a) (p Qp [𝒳⇘Suc n⇙]⇙ pvar Qp i) = eval_at_point Qp (t # a) p  (a!(i-1))"
          using 0 False ta_closed eval_pvar[of i "Suc n" "t#a"]
          by (metis A(2) nth_Cons')
        then show ?thesis
          using F0 A SA_poly_to_Qp_poly_closed[of a n "coord_ring_to_UP_SA n p"] to_fun_mult
                coord_ring_to_UP_SA_closed[of p n]   False UPQ.to_fun_to_poly UPQ.to_poly_closed assms
              eval_at_point_closed eval_pvar local.pvar_closed neq0_conv nth_Cons_pos ta_closed
          by (metis (no_types, lifting) UPQ.to_fun_mult)
      qed
    qed
  qed
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Gluing Semialgebraic Polynomials›
(**************************************************************************************************)
(**************************************************************************************************)


definition SA_poly_glue where
"SA_poly_glue m S f g = (λ n. fun_glue m S (f n) (g n))"

lemma SA_poly_glue_closed:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  shows "SA_poly_glue m S f g  carrier (UP (SA m))"
proof(rule UP_car_memI[of "max (deg (SA m) f) (deg (SA m) g)"])
  fix n assume A: "max (deg (SA m) f) (deg (SA m) g) < n" show "SA_poly_glue m S f g n = 𝟬SA m⇙"
    unfolding SA_poly_glue_def
  proof-
    have 0: "n > (deg (SA m) f)"
      using A by simp
    have 1: "n > (deg (SA m) g)"
      using A by simp
    have 2: "f n = 𝟬SA m⇙"
      using 0 assms UPSA.deg_leE by blast
    have 3: "g n = 𝟬SA m⇙"
      using 1 assms UPSA.deg_leE by blast
    show "fun_glue m S (f n) (g n) = 𝟬SA m⇙"
      unfolding SA_zero function_ring_def ring_record_simps function_zero_def 2 3
    proof fix x
      show " fun_glue m S (λxcarrier (Qpm). 𝟬) (λxcarrier (Qpm). 𝟬) x = (λxcarrier (Qpm). 𝟬) x"
        apply(cases "x  carrier (Qpm)")
        unfolding fun_glue_def restrict_def apply presburger
        by auto
    qed
  qed
  next
  show " n. SA_poly_glue m S f g n  carrier (SA m)"
    unfolding SA_poly_glue_def apply(rule fun_glue_closed)
    by(rule cfs_closed, rule assms, rule cfs_closed, rule assms, rule assms)
qed

lemma SA_poly_glue_deg:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  assumes "deg (SA m) f  d"
  assumes "deg (SA m) g  d"
  shows "deg (SA m) (SA_poly_glue m S f g)  d"
  apply(rule deg_leqI,  rule SA_poly_glue_closed, rule assms, rule assms, rule assms)
proof- fix n assume A: "d < n"
  show "SA_poly_glue m S f g n = 𝟬SA m⇙"
    unfolding SA_poly_glue_def
  proof-
    have 0: "n > (deg (SA m) f)"
      using A assms by linarith
    have 1: "n > (deg (SA m) g)"
      using A assms by linarith
    have 2: "f n = 𝟬SA m⇙"
      using 0 assms UPSA.deg_leE by blast
    have 3: "g n = 𝟬SA m⇙"
      using 1 assms UPSA.deg_leE by blast
    show "fun_glue m S (f n) (g n) = 𝟬SA m⇙"
      unfolding SA_zero function_ring_def ring_record_simps function_zero_def 2 3
    proof fix x
      show " fun_glue m S (λxcarrier (Qpm). 𝟬) (λxcarrier (Qpm). 𝟬) x = (λxcarrier (Qpm). 𝟬) x"
        apply(cases "x  carrier (Qpm)")
        unfolding fun_glue_def restrict_def apply presburger
        by auto
    qed
  qed
qed

lemma UP_SA_cfs_closed:
  assumes "g  carrier (UP (SA m))"
  shows "g k  carrier (SA m)"
  using assms UP_ring.cfs_closed[of "SA m" g k] SA_is_ring[of m] unfolding UP_ring_def
  by blast


lemma SA_poly_glue_cfs1:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  assumes "x  S"
  shows "(SA_poly_glue m S f g) n x = f n x"
  unfolding SA_poly_glue_def fun_glue_def restrict_def
  using assms
  by (metis SA_car local.function_ring_not_car padic_fields.UP_SA_cfs_closed padic_fields_axioms semialg_functions_memE(2))

lemma SA_poly_glue_cfs2:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  assumes "x  S"
  assumes "x  carrier (Qpm)"
  shows "(SA_poly_glue m S f g) n x = g n x"
  unfolding SA_poly_glue_def fun_glue_def restrict_def
  using assms by meson

lemma SA_poly_glue_to_Qp_poly1:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  assumes "x  S"
  shows "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) = SA_poly_to_Qp_poly m x f"
proof fix n
  have x_closed: "x  carrier (Qpm)"
    using assms is_semialgebraic_closed by blast
  have 0: "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) n = (SA_poly_glue m S f g) n x"
    by(rule SA_poly_to_Qp_poly_coeff[of x m "SA_poly_glue m S f g"], rule x_closed, rule SA_poly_glue_closed
          , rule assms, rule assms, rule assms)
  have 1: "(SA_poly_glue m S f g) n x = f n x"
    by(rule SA_poly_glue_cfs1 , rule assms, rule assms, rule assms, rule assms)
  show "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) n = SA_poly_to_Qp_poly m x f n"
    unfolding 0 1 using SA_poly_to_Qp_poly_coeff[of x m f n] assms (1) x_closed by blast
qed

lemma SA_poly_glue_to_Qp_poly2:
  assumes "f  carrier (UP (SA m))"
  assumes "g  carrier (UP (SA m))"
  assumes "is_semialgebraic m S"
  assumes "x  S"
  assumes "x  carrier (Qpm)"
  shows "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) = SA_poly_to_Qp_poly m x g"
proof fix n
  have x_closed: "x  carrier (Qpm)"
    using assms is_semialgebraic_closed by blast
  have 0: "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) n = (SA_poly_glue m S f g) n x"
    by(rule SA_poly_to_Qp_poly_coeff[of x m "SA_poly_glue m S f g"], rule x_closed, rule SA_poly_glue_closed
          , rule assms, rule assms, rule assms)
  have 1: "(SA_poly_glue m S f g) n x = g n x"
    by(rule SA_poly_glue_cfs2 , rule assms, rule assms, rule assms, rule assms, rule x_closed)
  show "SA_poly_to_Qp_poly m x (SA_poly_glue m S f g) n = SA_poly_to_Qp_poly m x g n"
    unfolding 0 1 using SA_poly_to_Qp_poly_coeff[of x m g n] assms x_closed by blast
qed


(**************************************************************************************************)
(**************************************************************************************************)
subsubsection‹Polynomials over the Valuation Ring›
(**************************************************************************************************)
(**************************************************************************************************)


definition integral_on where
"integral_on m B = {f  carrier (UP (SA m)). (x  B. i. SA_poly_to_Qp_poly m x f i  𝒪p)}"

lemma integral_on_memI:
  assumes "f  carrier (UP (SA m))"
  assumes "x i. x  B  SA_poly_to_Qp_poly m x f i  𝒪p"
  shows "f  integral_on m B"
  unfolding integral_on_def mem_Collect_eq using assms by blast

lemma integral_on_memE:
  assumes "f  integral_on m B"
  shows  "f  carrier (UP (SA m))"
          "x. x  B  SA_poly_to_Qp_poly m x f i  𝒪p"
 using assms unfolding integral_on_def mem_Collect_eq apply blast
 using assms unfolding integral_on_def mem_Collect_eq by blast

lemma one_integral_on:
  assumes "B  carrier (Qpm)"
  shows "𝟭UP (SA m) integral_on m B"
  apply(rule integral_on_memI)
  apply blast
proof- fix x i assume A: "x  B"
  have 0: "SA_poly_to_Qp_poly m x 𝟭UP (SA m)= 𝟭UP Qp⇙"
    apply(rule ring_hom_one[of _ "UP (SA m)"])
    using SA_poly_to_Qp_poly_is_hom[of x m] A assms by blast
  show "SA_poly_to_Qp_poly m x 𝟭UP (SA m)i  𝒪p"
    unfolding 0
    apply(rule val_ring_memI)
     apply(rule UPQ.cfs_closed)
     apply blast
    apply(cases "i = 0")
     apply (metis Qp.add.nat_pow_eone Qp.one_closed UPQ.cfs_one val_of_nat_inc val_one)
    by (metis Qp.int_inc_zero UPQ.cfs_one val_of_int_inc)
qed

lemma integral_on_plus:
  assumes "B  carrier (Qpm)"
  assumes "f  integral_on m B"
  assumes "g  integral_on m B"
  shows "f UP (SA m)g  integral_on m B"
proof(rule integral_on_memI)
  show "f UP (SA m)g  carrier (UP (SA m))"
    using assms integral_on_memE by blast
  show "x i. x  B  SA_poly_to_Qp_poly m x (f UP (SA m)g) i  𝒪p"
  proof- fix x i assume A: "x  B"
    have 0: "SA_poly_to_Qp_poly m x (f UP (SA m)g) = SA_poly_to_Qp_poly m x f
                                                          UP QpSA_poly_to_Qp_poly m x g"
      apply(rule SA_poly_to_Qp_poly_add)
      using A assms apply blast using assms integral_on_memE apply blast
      using assms integral_on_memE by blast
    have 1: "SA_poly_to_Qp_poly m x (f UP (SA m)g) i = SA_poly_to_Qp_poly m x f i
                                                            SA_poly_to_Qp_poly m x g i"
      unfolding 0 apply(rule UPQ.cfs_add)
       apply(rule SA_poly_to_Qp_poly_closed)
      using  A assms apply blast
      using assms integral_on_memE  apply blast
       apply(rule SA_poly_to_Qp_poly_closed)
      using  A assms apply blast
      using assms integral_on_memE  by blast
    show "SA_poly_to_Qp_poly m x (f UP (SA m)g) i  𝒪p"
      unfolding 1 using assms integral_on_memE
      using A val_ring_add_closed by presburger
  qed
qed

lemma integral_on_times:
  assumes "B  carrier (Qpm)"
  assumes "f  integral_on m B"
  assumes "g  integral_on m B"
  shows "f UP (SA m)g  integral_on m B"
proof(rule integral_on_memI)
  show "f UP (SA m)g  carrier (UP (SA m))"
    using assms integral_on_memE by blast
  show "x i. x  B  SA_poly_to_Qp_poly m x (f UP (SA m)g) i  𝒪p"
  proof- fix x i assume A: "x  B"
    have 0: "SA_poly_to_Qp_poly m x (f UP (SA m)g) = SA_poly_to_Qp_poly m x f
                                                          UP QpSA_poly_to_Qp_poly m x g"
      apply(rule SA_poly_to_Qp_poly_mult)
      using A assms apply blast using assms integral_on_memE apply blast
      using assms integral_on_memE by blast
    obtain S where S_def: "S = UP (Qp  carrier := 𝒪p )"
      by blast
    have 1: "cring S"
      unfolding S_def apply(rule UPQ.UP_ring_subring_is_ring)
      using val_ring_subring by blast
    have 2: " carrier S = {h  carrier (UP Qp). n. h n  𝒪p}"
      unfolding S_def  using UPQ.UP_ring_subring_car[of 𝒪p] val_ring_subring by blast
    have 3: "SA_poly_to_Qp_poly m x f  carrier S"
      unfolding 2 mem_Collect_eq using assms integral_on_memE SA_poly_to_Qp_poly_closed A
      by blast
    have 4: "SA_poly_to_Qp_poly m x g  carrier S"
      unfolding 2 mem_Collect_eq using assms integral_on_memE SA_poly_to_Qp_poly_closed A
      by blast
    have 5: "SA_poly_to_Qp_poly m x (f UP (SA m)g)  carrier S"
      unfolding 0
      using  cring.cring_simprules(5)[of S]3 4 1 UPQ.UP_ring_subring_mult[of 𝒪p "SA_poly_to_Qp_poly m x f" "SA_poly_to_Qp_poly m x g"]
      using S_def val_ring_subring by metis
    show "SA_poly_to_Qp_poly m x (f UP (SA m)g) i  𝒪p"
      using 5 unfolding 2 mem_Collect_eq by blast
  qed
qed

lemma integral_on_a_minus:
  assumes "B  carrier (Qpm)"
  assumes "f  integral_on m B"
  shows "UP (SA m)f  integral_on m B"
  apply(rule integral_on_memI)
  using assms integral_on_memE(1)[of f m B]
  apply blast
proof- fix x i assume A: "x  B"
  have 0: "SA_poly_to_Qp_poly m x (UP (SA m)f) = UP QpSA_poly_to_Qp_poly m x f"
    apply(rule  UP_cring.ring_hom_uminus[of "UP Qp" "UP (SA m)" "SA_poly_to_Qp_poly m x" f ] )
    unfolding UP_cring_def
       apply (simp add: UPQ.M_cring)
      apply (simp add: UPSA.P.ring_axioms)
     apply(rule SA_poly_to_Qp_poly_is_hom)
    using A assms apply blast
    using assms integral_on_memE by blast
  have 2: "f. f  carrier (UP Qp)  (UP Qpf) i =  (f i)"
    using UPQ.cfs_a_inv by blast
  have 1: "SA_poly_to_Qp_poly m x (UP (SA m)f) i =  (SA_poly_to_Qp_poly m x f) i "
    unfolding 0 apply(rule 2)
    using integral_on_memE  assms A SA_poly_to_Qp_poly_closed[of x m f] by blast
  show "SA_poly_to_Qp_poly m x (UP (SA m)f) i  𝒪p"
    unfolding 1 using A assms integral_on_memE(2)[of f m B x i]
    using val_ring_ainv_closed by blast
qed

lemma integral_on_subring:
  assumes "B  carrier (Qpm)"
  shows "subring (integral_on m B) (UP (SA m))"
proof(rule subringI)
  show "integral_on m B  carrier (UP (SA m))"
    unfolding integral_on_def by blast
  show "𝟭UP (SA m) integral_on m B"
    using one_integral_on assms by blast
  show " h. h  integral_on m B  UP (SA m)h  integral_on m B"
    using integral_on_a_minus assms by blast
  show "h1 h2. h1  integral_on m B  h2  integral_on m B  h1 UP (SA m)h2  integral_on m B"
    using integral_on_times assms by blast
  show "h1 h2. h1  integral_on m B  h2  integral_on m B  h1 UP (SA m)h2  integral_on m B"
    using integral_on_plus assms by blast
qed

lemma val_ring_add_pow:
  assumes "a  carrier Qp"
  assumes "val a  0"
  shows "val ([(n::nat)]a)  0"
proof-
  have 0: "[(n::nat)]a = ([n]𝟭)a"
    using assms Qp.add_pow_ldistr Qp.cring_simprules(12) Qp.one_closed by presburger
  show ?thesis unfolding 0 using assms
    by (meson Qp.nat_inc_closed val_ring_memE val_of_nat_inc val_ringI val_ring_times_closed)
qed

lemma val_ring_poly_eval:
  assumes "f  carrier (UP Qp)"
  assumes " i. f i  𝒪p"
  shows "x. x  𝒪p  f  x  𝒪p"
proof- fix x assume A: "x  𝒪p"
  obtain S where S_def: "S = (Qp  carrier := 𝒪p )"
    by blast
  have 0: "UP_cring S"
    unfolding S_def apply(rule UPQ.UP_ring_subring(1))
    using val_ring_subring by blast
  have 1: "to_function Qp f x = to_function S f x"
    unfolding S_def  apply(rule UPQ.UP_subring_eval)
    using val_ring_subring apply blast
    apply(rule UPQ.poly_cfs_subring)    using val_ring_subring apply blast
    using assms apply blast
    using assms apply blast using A by blast
  have 2: "f  carrier (UP S)"
    unfolding S_def
    using UPQ.UP_ring_subring_car[of 𝒪p] assms val_ring_subring by blast
  have 3: "to_function S f x  𝒪p"
      using UPQ.UP_subring_eval_closed[of 𝒪p f x]
    using 1 0 UP_cring.to_fun_closed[of S f x]
    unfolding S_def
    by (metis "2" A S_def UPQ.to_fun_def val_ring_subring)
  thus "f  x  𝒪p"
    using 1 UPQ.to_fun_def by presburger
qed

lemma SA_poly_constant_res_class_semialg':
  assumes "f  carrier (UP (SA m))"
  assumes "i x. x  B  f i x  𝒪p"
  assumes "deg (SA m) f  d"
  assumes "C  poly_res_classes n d"
  assumes "is_semialgebraic m B"
  shows "is_semialgebraic m {x  B. SA_poly_to_Qp_poly m x f  C}"
proof-
  obtain g where g_def: "g = SA_poly_glue m B f (𝟭UP (SA m))"
    by blast
  have g_closed: "g  carrier (UP (SA m))"
    unfolding g_def by(rule SA_poly_glue_closed, rule assms, blast, rule assms)
  have g_deg: "deg (SA m) g  d"
    unfolding g_def apply(rule SA_poly_glue_deg, rule assms, blast, rule assms, rule assms)
    unfolding deg_one by blast
  have 0: "{x  B. SA_poly_to_Qp_poly m x f  C} = B  {x  carrier (Qpm). SA_poly_to_Qp_poly m x g  C}"
    apply(rule equalityI', rule IntI, blast)
    unfolding mem_Collect_eq apply(rule conjI)
    using assms is_semialgebraic_closed apply blast
    unfolding g_def using SA_poly_glue_cfs1[of f m "𝟭UP (SA m)⇙" B] assms
    using SA_poly_glue_to_Qp_poly1 UPSA.P.cring_simprules(6) apply presburger
    unfolding g_def using SA_poly_glue_cfs1[of f m "𝟭UP (SA m)⇙" B] assms
    using SA_poly_glue_to_Qp_poly1 UPSA.P.cring_simprules(6)
    by (metis (no_types, lifting) Int_iff mem_Collect_eq)
  have 1: "i x. x  B  g i x = f i x"
    unfolding g_def by(rule SA_poly_glue_cfs1, rule assms, blast, rule assms, blast)
  have 2: "i x. x  carrier (Qpm)  g i x  𝒪p"
  proof- fix i x assume A: "x  carrier (Qpm)"
    show "g i x  𝒪p"
      unfolding g_def apply(cases "x  B")
      using SA_poly_glue_cfs1[of f m "𝟭UP (SA m)⇙" B x i]
            assms(1) assms(2)[of x i] 1
      using UPSA.P.cring_simprules(6) assms(5) apply presburger
      using SA_poly_glue_cfs2[of g m "𝟭UP (SA m)⇙" B x i] g_closed assms A
      by (metis (mono_tags, opaque_lifting) SA_poly_glue_cfs2 SA_poly_to_Qp_poly_coeff
          UPSA.P.cring_simprules(6) carrier_is_semialgebraic g_def integral_on_memE(2) is_semialgebraic_closed one_integral_on )
  qed
  have 3: "i x. x  carrier (Qpm)  x  B  g i x = 𝟭UP (SA m)i x"
    unfolding g_def apply(rule SA_poly_glue_cfs2[of f m "𝟭UP (SA m)⇙" B ])
    by(rule assms, blast, rule assms, blast, blast)
  have 4: "i x. x  carrier (Qpm)  x  B  g i x  𝒪p "
    using 3 cfs_one[of m] 2 by blast
  have 5: "i x. x  carrier (Qpm)  g i x  𝒪p"
    using 1 4 assms 2  by blast
  have 6: "is_semialgebraic m {x  carrier (Qpm). SA_poly_to_Qp_poly m x g  C}"
    by(rule SA_poly_constant_res_class_semialg[of _ _ d _ n], rule g_closed, rule 5, blast, rule g_deg, rule assms)
  show ?thesis unfolding 0
    by(rule intersection_is_semialg, rule assms, rule 6)
qed

lemma SA_poly_constant_res_class_decomp:
  assumes "f  carrier (UP (SA m))"
  assumes "i x. x  B  f i x  𝒪p"
  assumes "deg (SA m) f  d"
  assumes "is_semialgebraic m B"
  shows "B = ( C  poly_res_classes n d. {x  B. SA_poly_to_Qp_poly m x f  C})"
proof(rule equalityI')fix x assume A: "x  B"
  obtain g where g_def: "g = SA_poly_glue m B f (𝟭UP (SA m))"
    by blast
  have g_closed: "g  carrier (UP (SA m))"
    unfolding g_def by(rule SA_poly_glue_closed, rule assms, blast, rule assms)
  have g_deg: "deg (SA m) g  d"
    unfolding g_def apply(rule SA_poly_glue_deg, rule assms, blast, rule assms, rule assms)
    unfolding deg_one by blast
  have 1: "i x. x  B  g i x = f i x"
    unfolding g_def by(rule SA_poly_glue_cfs1, rule assms, blast, rule assms, blast)
  have 2: "i x. x  carrier (Qpm)  g i x  𝒪p"
  proof- fix i x assume A: "x  carrier (Qpm)"
    show "g i x  𝒪p"
      unfolding g_def apply(cases "x  B")
      using SA_poly_glue_cfs1[of f m "𝟭UP (SA m)⇙" B x i]
            assms(1) assms(2)[of x i] 1
      using UPSA.P.cring_simprules(6) assms(4) apply presburger
      using SA_poly_glue_cfs2[of g m "𝟭UP (SA m)⇙" B x i] g_closed assms A
      by (metis (mono_tags, opaque_lifting) SA_poly_glue_cfs2 SA_poly_to_Qp_poly_coeff
          UPSA.P.cring_simprules(6) carrier_is_semialgebraic g_def integral_on_memE(2) is_semialgebraic_closed one_integral_on )
  qed
  have 3: "i x. x  carrier (Qpm)  x  B  g i x = 𝟭UP (SA m)i x"
    unfolding g_def apply(rule SA_poly_glue_cfs2[of f m "𝟭UP (SA m)⇙" B ])
    by(rule assms, blast, rule assms, blast, blast)
  have 4: "i x. x  carrier (Qpm)  x  B  g i x  𝒪p "
    using 3 cfs_one[of m] 2 by blast
  have 5: "i x. x  carrier (Qpm)  g i x  𝒪p"
    using 1 4 assms 2  by blast
  have 6: " i x. x  carrier (Qpm)  SA_poly_to_Qp_poly m x g i = g i x"
    by(rule SA_poly_to_Qp_poly_coeff, blast, rule g_closed)
  have 7: " x. x  carrier (Qpm)  SA_poly_to_Qp_poly m x g  val_ring_polys_grad d"
    apply(rule val_ring_polys_grad_memI, rule SA_poly_to_Qp_poly_closed, blast, rule g_closed)
    unfolding  6 apply(rule 5, blast)
    using g_closed  SA_poly_to_Qp_poly_deg_bound[of g m] g_deg le_trans by blast
  have 8: "x. x  carrier (Qpm)  SA_poly_to_Qp_poly m x g  poly_res_class n d (SA_poly_to_Qp_poly m x g)"
    by(rule poly_res_class_refl, rule 7)
  have 9: "x. x  carrier (Qpm)  poly_res_class n d (SA_poly_to_Qp_poly m x g)  poly_res_classes n d"
    unfolding poly_res_classes_def using 7 by blast
  have 10: "x. x  B  SA_poly_to_Qp_poly m x g = SA_poly_to_Qp_poly m x f"
    unfolding g_def by(rule SA_poly_glue_to_Qp_poly1, rule assms, blast, rule assms, blast)
  have 11: "x  carrier (Qpm)"
    using A is_semialgebraic_closed assms by blast
  have 12: "SA_poly_to_Qp_poly m x g = SA_poly_to_Qp_poly m x f"
    by(rule 10, rule A)
  have 13: "x  {x  B. SA_poly_to_Qp_poly m x f  poly_res_class n d (SA_poly_to_Qp_poly m x g)}"
    using 11 10[of x] A 9[of x] 8[of x] unfolding 12 mem_Collect_eq  by blast
  show  "x  (Cpoly_res_classes n d. {x  B. SA_poly_to_Qp_poly m x f  C})"
    using 13 9[of x] 11  unfolding 12 mem_simps(8) mem_Collect_eq by auto
next
  show "x. x  (Cpoly_res_classes n d. {x  B. SA_poly_to_Qp_poly m x f  C})  x  B"
    by blast
qed

end

context UP_cring
begin

lemma pderiv_deg_bound:
  assumes "p  carrier P"
  assumes "deg R p  (Suc d)"
  shows "deg R (pderiv p)  d"
proof-
  have "deg R p  (Suc d)  deg R (pderiv p)  d"
    apply(rule poly_induct[of p])
      apply (simp add: assms(1))
    using deg_zero pderiv_deg_0 apply presburger
  proof fix p assume A: "(q. q  carrier P  deg R q < deg R p  deg R q  Suc d  deg R (pderiv q)  d)"
                        "p  carrier P" "0 < deg R p" "deg R p  Suc d"
    obtain q where q_def: "q  carrier P  deg R q < deg R p  p = q Pltrm p"
      using A  ltrm_decomp by metis
    have 0: "deg R (pderiv (ltrm p))  d"
    proof-
      have 1: "pderiv (ltrm p) = up_ring.monom P ([deg R p]  p (deg R p)) (deg R p - 1)"
        using pderiv_monom[of "lcf p" "deg R p"] A  P_def UP_car_memE(1) by auto
      show ?thesis unfolding 1
        by (metis (no_types, lifting) A(2) A(3) A(4) R.add_pow_closed Suc_diff_1 cfs_closed deg_monom_le le_trans not_less_eq_eq)
    qed
    have "deg R q  Suc d" using A q_def  by linarith
    then have 1: "deg R (pderiv q)  d"
      using A q_def by blast
    hence "max (deg R (pderiv q)) (deg R (pderiv (up_ring.monom P (p (deg R p)) (deg R p))))   d"
      using 0 max.bounded_iff by blast
    thus "deg R (pderiv p)  d "
      using q_def pderiv_add pderiv_monom[of "lcf p" "deg R p"] A deg_add[of "pderiv q" "pderiv (ltrm p)"]
      by (metis "0" "1" ltrm_closed bound_deg_sum pderiv_closed)
  qed
  thus ?thesis
    using assms(2) by blast
qed

lemma(in cring) minus_zero:
"a  carrier R  a  𝟬 = a"
  unfolding a_minus_def
  by (metis add.l_cancel_one' cring_simprules(2) cring_simprules(22))

lemma (in UP_cring) taylor_expansion_at_zero:
  assumes "g  carrier (UP R)"
  shows "taylor_expansion R 𝟬 g = g"
proof-
  have 0: "X_plus 𝟬 = X_poly R"
    unfolding X_poly_plus_def
    by (metis ctrm_degree lcf_eq P.r_zero P_def R.zero_closed UP_cring.ctrm_is_poly UP_cring.to_poly_inverse UP_zero_closed X_closed deg_nzero_nzero is_UP_cring to_fun_ctrm to_fun_zero)
  show ?thesis
    unfolding taylor_expansion_def 0
    using assms UP_cring.X_sub is_UP_cring by blast
qed
end
(**************************************************************************************************)
(**************************************************************************************************)
subsection‹Partitioning Semialgebraic Sets By Zero Sets of Function›
(**************************************************************************************************)
(**************************************************************************************************)

context padic_fields
begin

definition SA_funs_to_SA_decomp where
"SA_funs_to_SA_decomp n Fs S = atoms_of ((∩) S ` ((SA_zero_set n ` Fs)  (SA_nonzero_set n ` Fs))) "

lemma SA_funs_to_SA_decomp_closed_0:
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "(∩) S ` ((SA_zero_set n ` Fs)  (SA_nonzero_set n ` Fs))  semialg_sets n"
proof(rule subsetI)
  fix x assume A: "x  (∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
  show " x  semialg_sets n"
  proof(cases "x  (∩) S ` SA_zero_set n ` Fs")
    case True
    then obtain f where f_def: "f  Fs  x = S  SA_zero_set n f"
      by blast
    then show "x  semialg_sets n" using assms SA_zero_set_is_semialgebraic
      by (meson is_semialgebraicE semialg_intersect subset_iff)
  next
    case False
    then have "x  (∩) S ` SA_nonzero_set n ` Fs"
      using A by blast
    then obtain f where f_def: "f  Fs  x = S  SA_nonzero_set n f"
      using A by blast
    then show "x  semialg_sets n" using assms SA_nonzero_set_is_semialgebraic
      by (meson padic_fields.is_semialgebraicE padic_fields.semialg_intersect padic_fields_axioms subset_iff)
  qed
qed

lemma SA_funs_to_SA_decomp_closed:
  assumes "finite Fs"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "SA_funs_to_SA_decomp n Fs S  semialg_sets n"
  unfolding SA_funs_to_SA_decomp_def semialg_sets_def
  apply(rule atoms_of_gen_boolean_algebra)
  using SA_funs_to_SA_decomp_closed_0[of Fs n S]  assms  unfolding semialg_sets_def
   apply blast
  using assms
  by blast

lemma SA_funs_to_SA_decomp_finite:
  assumes "finite Fs"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "finite (SA_funs_to_SA_decomp n Fs S)"
  unfolding SA_funs_to_SA_decomp_def
  apply(rule finite_set_imp_finite_atoms)
  using assms by blast

lemma SA_funs_to_SA_decomp_disjoint:
  assumes "finite Fs"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "disjoint (SA_funs_to_SA_decomp n Fs S)"
  apply(rule disjointI) unfolding SA_funs_to_SA_decomp_def
  apply(rule atoms_of_disjoint[of _ " ((∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs))"])
    apply blast apply blast by blast

lemma pre_SA_funs_to_SA_decomp_in_algebra:
  shows " ((∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs))  gen_boolean_algebra S (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
proof(rule subsetI) fix x assume A: " x   (∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
  then obtain A where A_def: "A  (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)  x = S  A"
    by blast
  then show "x  gen_boolean_algebra S (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
    using gen_boolean_algebra.generator[of A "SA_zero_set n ` Fs  SA_nonzero_set n ` Fs" S]
    by (metis inf.commute)
qed

lemma SA_funs_to_SA_decomp_in_algebra:
  assumes "finite Fs"
  shows "SA_funs_to_SA_decomp n Fs S  gen_boolean_algebra S (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
  unfolding SA_funs_to_SA_decomp_def  apply(rule atoms_of_gen_boolean_algebra)
  using pre_SA_funs_to_SA_decomp_in_algebra[of S n Fs] apply blast
  using assms by blast

lemma SA_funs_to_SA_decomp_subset:
  assumes "finite Fs"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  assumes "A  SA_funs_to_SA_decomp n Fs S"
  shows "A  S"
proof-
  have "A  gen_boolean_algebra S (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
  using assms SA_funs_to_SA_decomp_in_algebra[of Fs n S]
        atoms_of_gen_boolean_algebra[of _  S "(SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"]
  unfolding SA_funs_to_SA_decomp_def by blast
  then show ?thesis using gen_boolean_algebra_subset by blast
qed

lemma SA_funs_to_SA_decomp_memE:
  assumes "finite Fs"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  assumes "A  (SA_funs_to_SA_decomp n Fs S)"
  assumes "f  Fs"
  shows "A  SA_zero_set n f  A  SA_nonzero_set n f"
proof(cases "A  S  SA_zero_set n f")
case True
  then show ?thesis
    by blast
next
  case False
  have 0: "S  SA_zero_set n f  (∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs)"
    using assms
    by blast
  then have 1: "A  (S  SA_zero_set n f) = {}"
    using False assms atoms_are_minimal[of A "((∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs))" "S  SA_zero_set n f"]
    unfolding SA_funs_to_SA_decomp_def
    by blast
  have 2: "A  S"
    using assms SA_funs_to_SA_decomp_subset by blast
  then show ?thesis
    using 0 1 False zero_set_nonzero_set_covers_semialg_set[of n S] assms(3) by auto
qed

lemma SA_funs_to_SA_decomp_covers:
  assumes "finite Fs"
  assumes "Fs  {}"
  assumes "Fs  carrier (SA n)"
  assumes "is_semialgebraic n S"
  shows "S =  (SA_funs_to_SA_decomp n Fs S)"
proof-
  have 0: "S =  ((∩) S ` ((SA_zero_set n ` Fs)  (SA_nonzero_set n ` Fs)))"
  proof
    obtain f where f_def: "f  Fs"
      using assms by blast
    have 0: "S  SA_nonzero_set n f  ((∩) S ` ((SA_zero_set n ` Fs)  (SA_nonzero_set n ` Fs)))"
      using f_def by blast
    have 1: "S  SA_zero_set n f  ((∩) S ` ((SA_zero_set n ` Fs)  (SA_nonzero_set n ` Fs)))"
      using f_def by blast
    have 2: "S = S  SA_zero_set n f  S  SA_nonzero_set n f"
      by (simp add: assms(4) zero_set_nonzero_set_covers_semialg_set)
    then show "S   ((∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs))"
      using 0 1 Sup_upper2 Un_subset_iff subset_refl by blast
    show " ((∩) S ` (SA_zero_set n ` Fs  SA_nonzero_set n ` Fs))  S"
      by blast
  qed
  show ?thesis
    unfolding SA_funs_to_SA_decomp_def atoms_of_covers' using 0 by blast
qed

end
end