diff --git a/thys/Padic_Field/Cring_Multivariable_Poly.thy b/thys/Padic_Field/Cring_Multivariable_Poly.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Cring_Multivariable_Poly.thy @@ -0,0 +1,8548 @@ +theory Cring_Multivariable_Poly +imports "HOL-Algebra.Indexed_Polynomials" Padic_Ints.Cring_Poly +begin + +(**************************************************************************************************) +(**************************************************************************************************) +section\Multivariable Polynomials Over a Commutative Ring\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + This theory extends the content of \texttt{HOL-Algebra.Indexed\_Polynomials}, mainly in the + context of a commutative base ring. In particular, the ring of polynomials over an arbitrary + variable set is explicitly witnessed as a ring itself, which is commutative if the base is + commutative, and a domain if the base ring is a domain. A universal property for polynomial + evaluation is proved, which allows us to embed polynomial rings in a ring of functions over the + base ring, as well as construe multivariable polynomials as univariate polynomials over a small + base polynomial ring. +\ + +type_synonym 'a monomial = "'a multiset" +type_synonym ('b,'a) mvar_poly = "'a multiset \ 'b" +type_synonym ('a,'b) ring_hom = "'a \ 'b" +type_synonym 'a u_poly = "nat \ 'a" +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Lemmas about multisets\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + Since multisets function as monomials in this formalization, we'll need some simple lemmas + about multisets in order to define degree functions and certain lemmas about factorizations. + Those are provided in this section. +\ + +lemma count_size: + assumes "size m \ K" + shows "count m i \ K" + using assms + by (metis count_le_replicate_mset_subset_eq dual_order.trans order_refl size_mset_mono size_replicate_mset) + +text\The following defines the set of monomials with nonzero coefficients for a given polynomial:\ + +definition monomials_of :: "('a,'c) ring_scheme \ ('a, 'b) mvar_poly \ ('b monomial) set" where +"monomials_of R P = {m. P m \ \\<^bsub>R\<^esub>}" + +context ring +begin + +lemma monomials_of_index_free: + assumes "index_free P i" + assumes "m \ monomials_of R P" + shows "count m i = 0" + using assms + unfolding monomials_of_def index_free_def + by (meson count_inI mem_Collect_eq) + +lemma index_freeI: + assumes "\m. m \ monomials_of R P \ count m i = 0" + shows "index_free P i" + unfolding index_free_def +proof safe + fix m + assume "i \# m" + then have "count m i \ 0" + by simp + then have "m \ monomials_of R P" + using assms + by meson + then show "P m = \" + unfolding monomials_of_def + by blast +qed + +text\\texttt{monomials\_of} R is subadditive\ + +lemma monomials_of_add: +"monomials_of R (P \ Q) \ (monomials_of R P) \ (monomials_of R Q)" +proof + fix x + assume "x \ monomials_of R (P \ Q) " + then have "P x \ Q x \\" + by (simp add: indexed_padd_def monomials_of_def) + then have "P x \\ \ Q x \ \" + by auto + then show "x \ (monomials_of R P) \ (monomials_of R Q)" + by (simp add: monomials_of_def) +qed + +lemma monomials_of_add_finite: + assumes "finite (monomials_of R P)" + assumes "finite (monomials_of R Q)" + shows "finite (monomials_of R (P \ Q))" + by (meson assms(1) assms(2) finite_Un finite_subset monomials_of_add) + +lemma monomials_ofE: + assumes "m \ monomials_of R p" + shows "p m \ \" + using assms + unfolding monomials_of_def + by simp + +lemma complement_of_monomials_of: + assumes "m \ monomials_of R P" + shows "P m = \" + using assms + unfolding monomials_of_def + by blast + +text\Multiplication by an indexed variable corresponds to adding that index to each monomial:\ + +lemma monomials_of_p_mult: +"monomials_of R (P \ i) = {m. \ n \ (monomials_of R P). m = add_mset i n}" +proof + show "monomials_of R (P \ i) \ {m. \n\monomials_of R P. m = add_mset i n}" + proof + fix m + assume A: "m \ monomials_of R (P \ i)" + show "m \ {m. \n\monomials_of R P. m = add_mset i n}" + proof- + have "(P \ i) m \ \" + by (simp add: A monomials_ofE) + then have "P (m - {# i #}) \\" + unfolding indexed_pmult_def + by presburger + then have 0: "(m - {# i #}) \ monomials_of R P" + by (meson complement_of_monomials_of) + have 1: " m = add_mset i (m - {# i #})" + by (metis \(P \ i) m \ \\ add_mset_remove_trivial_If indexed_pmult_def) + then show ?thesis using 0 1 + by blast + qed + qed + show "{m. \n\monomials_of R P. m = add_mset i n} \ monomials_of R (P \ i)" + unfolding monomials_of_def indexed_pmult_def + by auto +qed + +lemma monomials_of_p_mult': +"monomials_of R (p \ i) = add_mset i ` (monomials_of R p)" + using monomials_of_p_mult + by (simp add: monomials_of_p_mult image_def) + +lemma monomials_of_p_mult_finite: + assumes "finite (monomials_of R P)" + shows "finite (monomials_of R (P \ i))" + using assms monomials_of_p_mult'[of P i] + by simp + +text\Monomials of a constant either consist of the empty multiset, or nothing:\ + +lemma monomials_of_const: +"(monomials_of R (indexed_const k)) = (if (k = \) then {} else {{#}})" + unfolding monomials_of_def indexed_const_def + by simp + +lemma monomials_of_const_finite: +"finite (monomials_of R (indexed_const k))" + by (simp add: monomials_of_const) + +text\A polynomial always has finitely many monomials:\ +lemma monomials_finite: + assumes "P \ indexed_pset K I" + shows "finite (monomials_of R P)" + using assms + apply(induction P) + using monomials_of_const_finite apply blast + using monomials_of_add_finite apply blast + by (simp add: monomials_of_p_mult_finite) +end + +(**************************************************************************************************) +(**************************************************************************************************) + subsection\Turning monomials into polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\Constructor for turning a monomial into a polynomial\ + +definition mset_to_IP :: "('a, 'b) ring_scheme \ 'c monomial \ ('a,'c) mvar_poly" where +"mset_to_IP R m n= (if (n = m) then \\<^bsub>R\<^esub> else \\<^bsub>R\<^esub>)" + +definition var_to_IP :: "('a, 'b) ring_scheme \ 'c \ ('a,'c) mvar_poly" ("pvar") where +"var_to_IP R i = mset_to_IP R {#i#}" + +context ring +begin + +lemma mset_to_IP_simp[simp]: +"mset_to_IP R m m = \" + by (simp add: mset_to_IP_def) + +lemma mset_to_IP_simp'[simp]: + assumes "n \m" + shows "mset_to_IP R m n = \" + by (simp add: assms mset_to_IP_def) + +lemma(in cring) monomials_of_mset_to_IP: + assumes "\ \\" + shows "monomials_of R (mset_to_IP R m) = {m}" + unfolding monomials_of_def mset_to_IP_def +proof + show "{ma. (if ma = m then \ else \) \ \} \ {m}" + by auto + show "{m} \ {ma. (if ma = m then \ else \) \ \}" + using assms by auto +qed + +end + +text\The set of monomials of a fixed \P\ which include a given variable:\ + +definition monomials_with :: "('a, 'b) ring_scheme \ 'c \ ('a,'c) mvar_poly \ ('c monomial) set" where +"monomials_with R i P = {m. m \ monomials_of R P \ i \# m}" + +context ring +begin + +lemma monomials_withE: + assumes "m \ monomials_with R i P" + shows "i \# m" + "m \ monomials_of R P" + using assms unfolding monomials_with_def + apply blast + using assms unfolding monomials_with_def + by blast + +lemma monomials_withI: + assumes "i \# m" + assumes "m \ monomials_of R P" + shows "m \ monomials_with R i P" + using assms + unfolding monomials_with_def + by blast + +end + +text\Restricting a polynomial to be zero outside of a given monomial set:\ + +definition restrict_poly_to_monom_set :: + "('a, 'b) ring_scheme \ ('a,'c) mvar_poly \ ('c monomial) set \('a,'c) mvar_poly" where +"restrict_poly_to_monom_set R P m_set m = (if m \ m_set then P m else \\<^bsub>R\<^esub>)" + +context ring +begin + +lemma restrict_poly_to_monom_set_coeff: + assumes "carrier_coeff P" + shows "carrier_coeff (restrict_poly_to_monom_set R P Ms)" + by (metis assms carrier_coeff_def restrict_poly_to_monom_set_def zero_closed) + +lemma restrict_poly_to_monom_set_coeff': + assumes "P \ indexed_pset K I" + assumes "I \ {}" + assumes "\m. P m \ S" + assumes "\ \ S" + shows "\m. (restrict_poly_to_monom_set R P m_set m) \ S" + using assms + unfolding restrict_poly_to_monom_set_def + by simp + +lemma restrict_poly_to_monom_set_monoms: + assumes "P \ indexed_pset I K" + assumes "m_set \ monomials_of R P" + shows "monomials_of R (restrict_poly_to_monom_set R P m_set) = m_set" + using assms + unfolding monomials_of_def restrict_poly_to_monom_set_def + by (simp add: subset_iff) +end + +(**************************************************************************************************) +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Degree Functions\ +(**************************************************************************************************) +(**************************************************************************************************) +(**************************************************************************************************) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Total Degree Function\ + (**************************************************************************************************) + (**************************************************************************************************) + + +lemma multi_set_size_count: + fixes m :: "'c monomial" + shows "size m \ count m i" + by (metis count_le_replicate_mset_subset_eq order_refl size_mset_mono size_replicate_mset) + +text\Total degree function\ + +definition total_degree :: "('a, 'b) ring_scheme \ ('a, 'c) mvar_poly \ nat" where +"total_degree R P = (if (monomials_of R P = {}) then 0 else Max (size ` (monomials_of R P)))" + +context ring +begin + +lemma total_degree_ineq: + assumes "m \ monomials_of R P" + assumes "finite (monomials_of R P)" + shows "total_degree R P \ size m" + unfolding total_degree_def using assms + by force + +lemma total_degree_in_monomials_of: + assumes "monomials_of R P \ {}" + assumes "finite (monomials_of R P)" + obtains m where "m \ monomials_of R P \ size m = total_degree R P" + using assms Max_in[of "(size ` monomials_of R P)"] unfolding total_degree_def + by (metis (mono_tags, lifting) empty_is_image finite_imageI image_iff) + +lemma total_degreeI: + assumes "finite (monomials_of R P)" + assumes "\m. m \ monomials_of R P \ size m = n" + assumes "\m. m \ monomials_of R P \ size m \ n" + shows "n = total_degree R P" +proof(cases "monomials_of R P = {}") + case True + then show ?thesis + using assms by blast +next + case False + obtain m where m_def: "m \ monomials_of R P \ size m = n" + using assms by blast + have 0: "n \ (size ` (monomials_of R P))" + using m_def + by blast + have "\k. k \ (size ` (monomials_of R P)) \ k \ n" + using assms + by blast + then have 1: "\m. m \ (monomials_of R P) \ size m\ n" + by blast + obtain m' where m'_def: "m' \ monomials_of R P \ size m' = total_degree R P" + using assms total_degree_in_monomials_of + by blast + then have 2: "size m' \ n" + using 1 + by blast + have 3: "n \size m'" + using assms m'_def total_degree_ineq by auto + show ?thesis using 2 3 + using dual_order.antisym m'_def by blast +qed +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Degree in One Variable\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition degree_in_var :: + "('a, 'b) ring_scheme \ ('a, 'c) mvar_poly \ 'c \ nat" where +"degree_in_var R P i = (if (monomials_of R P = {}) then 0 else Max ((\m. count m i) ` (monomials_of R P)))" + +context ring +begin + +lemma degree_in_var_ineq: + assumes "m \ monomials_of R P" + assumes "finite (monomials_of R P)" + shows "degree_in_var R P i \ count m i" + unfolding degree_in_var_def using assms + by force + +lemma degree_in_var_in_monomials_of: + assumes "monomials_of R P \ {}" + assumes "finite (monomials_of R P)" + obtains m where "m \ monomials_of R P \ count m i = degree_in_var R P i" + using assms Max_in[of "((\m. count m i) ` (monomials_of R P))"] unfolding degree_in_var_def + by (metis (no_types, lifting) empty_is_image finite_imageI image_iff) + +lemma degree_in_varI: + assumes "finite (monomials_of R P)" + assumes "\m. m \ monomials_of R P \ count m i = n" + assumes "\c. c \ monomials_of R P \ count c i \ n" + shows "n = degree_in_var R P i" +proof- + obtain l where l_def: "l \ monomials_of R P \ count l i = degree_in_var R P i" + by (metis assms(1) assms(2) degree_in_var_in_monomials_of equals0D) + have "degree_in_var R P i \ n" + using assms(3) l_def + by force + then show ?thesis + using assms(1) assms(2) dual_order.antisym degree_in_var_ineq + by fastforce +qed + +text\Total degree bounds degree in a single variable:\ + +lemma total_degree_degree_in_var: + assumes "finite (monomials_of R P)" + shows "total_degree R P \ degree_in_var R P i" +proof(cases " (monomials_of R P) = {}") + case True + then show ?thesis + unfolding total_degree_def degree_in_var_def + by simp +next + case False + then obtain m1 where m1_def: "m1 \ monomials_of R P \ count m1 i = degree_in_var R P i" + by (meson assms degree_in_var_in_monomials_of) + have "size m1 \count m1 i" + by (simp add: multi_set_size_count) + then show ?thesis + by (simp add: assms local.ring_axioms m1_def order.trans ring.total_degree_ineq) +qed +end + +text\The set of monomials of maximal degree in variable \i\:\ + +definition max_degree_monoms_in_var :: + "('a, 'b) ring_scheme \ ('a, 'c) mvar_poly \ 'c \ ('c monomial) set" where +"max_degree_monoms_in_var R P i = {m. m \ monomials_of R P \ count m i = degree_in_var R P i}" + +context ring +begin + +lemma max_degree_monoms_in_var_memI: + assumes "m \ monomials_of R P" + assumes "count m i = degree_in_var R P i" + shows "m \ max_degree_monoms_in_var R P i" + using assms unfolding max_degree_monoms_in_var_def + by blast + +lemma max_degree_monoms_in_var_memE: + assumes "m \ max_degree_monoms_in_var R P i" + shows "m \ monomials_of R P" + "count m i = degree_in_var R P i" + using assms unfolding max_degree_monoms_in_var_def + apply blast + using assms unfolding max_degree_monoms_in_var_def + by blast +end + +text\The set of monomials of \P\ of fixed degree in variable \i\:\ + +definition fixed_degree_in_var :: + "('a, 'b) ring_scheme \ ('a, 'c) mvar_poly \ 'c \ nat \ ('c monomial) set" where +"fixed_degree_in_var R P i n = {m. m \ monomials_of R P \ count m i = n}" + +context ring +begin + +lemma fixed_degree_in_var_subset: +"fixed_degree_in_var R P i n \ monomials_of R P" + unfolding fixed_degree_in_var_def + by blast + +lemma fixed_degree_in_var_max_degree_monoms_in_var: +"max_degree_monoms_in_var R P i = fixed_degree_in_var R P i (degree_in_var R P i)" + unfolding max_degree_monoms_in_var_def fixed_degree_in_var_def + by auto + +lemma fixed_degree_in_varI: + assumes "m \ monomials_of R P" + assumes "count m i = n" + shows "m \ fixed_degree_in_var R P i n" + unfolding fixed_degree_in_var_def + using assms + by blast + +lemma fixed_degree_in_varE: + assumes "m \ fixed_degree_in_var R P i n" + shows "m \ monomials_of R P" + "count m i = n" + apply (meson assms fixed_degree_in_var_subset in_mono) + using assms fixed_degree_in_var_def by force + +definition restrict_to_var_deg :: + "('a,'c) mvar_poly \ 'c \ nat \ 'c monomial \ 'a" where +"restrict_to_var_deg P i n = restrict_poly_to_monom_set R P (fixed_degree_in_var R P i n)" + +lemma restrict_to_var_deg_var_deg: + assumes "finite (monomials_of R P)" + assumes "Q = restrict_to_var_deg P i n" + assumes "monomials_of R Q \ {}" + shows "n = degree_in_var R Q i" + apply(rule degree_in_varI) + apply (metis assms(1) assms(2) fixed_degree_in_varE(1) monomials_ofE restrict_poly_to_monom_set_def + restrict_to_var_deg_def rev_finite_subset subsetI) + apply (metis (full_types) assms(2) assms(3) equals0I fixed_degree_in_varE(2) + monomials_ofE restrict_poly_to_monom_set_def restrict_to_var_deg_def) + by (metis assms(2) eq_iff fixed_degree_in_varE(2) monomials_ofE restrict_poly_to_monom_set_def restrict_to_var_deg_def) + +lemma index_free_degree_in_var[simp]: + assumes "index_free P i" + shows "degree_in_var R P i = 0" +proof(cases "monomials_of R P = {}") + case True + then show ?thesis + using assms + unfolding degree_in_var_def + by simp +next + case False + then have 0: "degree_in_var R P i = Max ((\m. count m i) ` (monomials_of R P))" + unfolding degree_in_var_def + by simp + have"((\m. count m i) ` (monomials_of R P)) = {0}" + proof + show "(\m. count m i) ` monomials_of R P \ {0}" + using False assms monomials_of_index_free[of P i] + by auto + show "{0} \ (\m. count m i) ` monomials_of R P" + using False \(\m. count m i) ` monomials_of R P \ {0}\ + by auto + qed + then show ?thesis using 0 + by simp +qed + +lemma degree_in_var_index_free: + assumes "degree_in_var R P i = 0" + assumes "finite (monomials_of R P)" + shows "index_free P i" + apply(rule index_freeI) + by (metis assms(1) assms(2) degree_in_var_ineq le_zero_eq) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Constructing the Multiplication Operation on the Ring of Indexed Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + + (**********************************************************************) + (**********************************************************************) + subsubsection\The Set of Factors of a Fixed Monomial\ + (**********************************************************************) + (**********************************************************************) + + text\The following function maps a monomial to the set of monomials which divide it:\ + +definition mset_factors :: "'c monomial \ ('c monomial) set" where +"mset_factors m = {n. n \# m}" + +context ring +begin + +lemma monom_divides_factors: +"n \ (mset_factors m)\ n \# m" + unfolding mset_factors_def by auto + +lemma mset_factors_mono: + assumes "n \# m" + shows "mset_factors n \ mset_factors m" + unfolding mset_factors_def + by (simp add: Collect_mono_iff assms subset_mset.order_trans) + +lemma mset_factors_size_bound: + assumes "n \ mset_factors m" + shows "size n \ size m" + using assms + unfolding mset_factors_def + by (simp add: size_mset_mono) + +lemma sets_to_inds_finite: + assumes "finite I" + shows "finite S \ finite (Pi\<^sub>E S (\_. I))" + using assms + by (simp add: finite_PiE) + +end + + (**********************************************************************************************) + (**********************************************************************************************) + subsubsection\Finiteness of the Factor Set of a Monomial\ + (**********************************************************************************************) + (**********************************************************************************************) + +text\ + This section shows that any monomial m has only finitely many factors. This is done by mapping + the set of factors injectively into a finite extensional function set. Explicitly, a monomial is + just mapped to its count function, restricted to the set of indices where the count is nonzero. +\ + +definition mset_factors_to_fun :: + "('a,'b) ring_scheme \ 'c monomial \ 'c monomial \ ('c \ nat)" where +"mset_factors_to_fun R m n = (if (n \ mset_factors m) then + (restrict (count n) (set_mset m)) else undefined)" +context ring +begin + +lemma mset_factors_to_fun_prop: + assumes "size m = n" + shows "mset_factors_to_fun R m \ (mset_factors m) \ (Pi\<^sub>E (set_mset m) (\_. {0.. n}))" +proof + fix x + assume A: "x \ (mset_factors m)" + show "mset_factors_to_fun R m x \ (set_mset m) \\<^sub>E {0..n} " + proof + show "\xa. xa \# m \ mset_factors_to_fun R m x xa \ {0..n}" + proof- + fix y + assume Ay: "y \# m" + show P0: "mset_factors_to_fun R m x y \ {0..n}" + proof- + have "mset_factors_to_fun R m x = restrict (count x) (set_mset m)" + using A unfolding mset_factors_to_fun_def + by simp + then have "mset_factors_to_fun R m x y = count x y" + using Ay + by simp + then show ?thesis + using A INTEG.R.mset_factors_size_bound assms count_size by fastforce + qed + qed + fix z + assume Ay: "z \#m" + show "mset_factors_to_fun R m x z = undefined" + using A Ay unfolding mset_factors_to_fun_def + by simp + qed +qed + +lemma mset_factors_to_fun_inj: + shows "inj_on (mset_factors_to_fun R m) (mset_factors m) " +proof + fix x y + assume A: "x \ mset_factors m" "y \ mset_factors m" + show "mset_factors_to_fun R m x = mset_factors_to_fun R m y \ x = y" + proof- + assume A0: "mset_factors_to_fun R m x = mset_factors_to_fun R m y" + show " x = y" + proof- + have "\i. count x i = count y i" + proof- fix i + show "count x i = count y i" + proof(cases "i \# m") + case True + then show ?thesis using A0 A unfolding mset_factors_to_fun_def + by (metis restrict_def) + next + case False + then show ?thesis + using A0 A unfolding mset_factors_to_fun_def + by (metis monom_divides_factors count_inI mset_subset_eqD) + qed + qed + then show ?thesis + using multiset_eqI by blast + qed + qed +qed + +lemma finite_target: +"finite (Pi\<^sub>E (set_mset m) (\_. {0..(n::nat)}))" +proof- + have 0: "finite (set_mset m)" + by simp + have 1: "finite ({0..n})" + by simp + then show ?thesis using 0 + by (simp add: finite_PiE) +qed + +text\A multiset has only finitely many factors:\ + +lemma mset_factors_finite[simp]: +"finite (mset_factors m)" +proof- + have 0: "inj_on (mset_factors_to_fun R m) (mset_factors m) " + by (simp add: mset_factors_to_fun_inj) + have 1: "(mset_factors_to_fun R m) \ (mset_factors m) \ (Pi\<^sub>E (set_mset m) (\_. {0 .. (size m)}))" + by (metis mset_factors_to_fun_prop) + have 2: "finite (Pi\<^sub>E (set_mset m) (\_. {0 .. (size m)}))" + by (simp add: finite_target) + have 3: "((mset_factors_to_fun R m) ` (mset_factors m)) \ (Pi\<^sub>E (set_mset m) (\_. {0 .. (size m)}))" + using 1 2 + by blast + then have "finite ((mset_factors_to_fun R m) ` (mset_factors m))" + using 2 finite_subset by auto + then show ?thesis using 0 1 2 + finite_imageD[of "mset_factors_to_fun R m" "mset_factors m" ] + by blast +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Definition of Indexed Polynomial Multiplication.\ +(**************************************************************************************************) +(**************************************************************************************************) + +context ring +begin + +text\Monomial division:\ + +lemma monom_divide: + assumes "n \ mset_factors m" + shows "(THE k. n + k = m ) = m - n " + apply(rule the_equality) + using assms unfolding mset_factors_def + apply simp + using assms unfolding mset_factors_def + by auto + +text\A monomial is a factor of itself:\ + +lemma m_factor[simp]: +"m \ mset_factors m" + using local.ring_axioms ring.monom_divides_factors by blast + +end + +text\The definition of polynomial multiplication:\ + +definition P_ring_mult :: "('a, 'b) ring_scheme \ ('a,'c) mvar_poly \ ('a,'c) mvar_poly \ 'c monomial \ 'a" + where +"P_ring_mult R P Q m = (finsum R (\x. (P x) \\<^bsub>R\<^esub> (Q (m - x))) (mset_factors m))" + +abbreviation(in ring) P_ring_mult_in_ring (infixl "\\<^sub>p" 65)where +"P_ring_mult_in_ring \ P_ring_mult R" + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Distributivity Laws for Polynomial Multiplication\ +(**************************************************************************************************) +(**************************************************************************************************) + +context ring +begin + +lemma P_ring_rdistr: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + shows "a \\<^sub>p (b \ c) = (a \\<^sub>p b)\ (a \\<^sub>p c)" +proof + fix m + show "(a \\<^sub>p (b \ c)) m = (a \\<^sub>p b \ (a \\<^sub>p c)) m" + proof- + have RHS: "(a \\<^sub>p b \ (a \\<^sub>p c)) m = + (\x\mset_factors m. a x \ b (m - x)) \ (\x\mset_factors m. a x \ c (m - x))" + unfolding indexed_padd_def P_ring_mult_def by auto + have LHS: "(a \\<^sub>p (b \ c)) m= + ( \x\mset_factors m. a x \ b (m - x) \ a x \ c (m - x))" + unfolding indexed_padd_def P_ring_mult_def + by (meson assms(1) assms(2) assms(3) local.ring_axioms r_distr ring.carrier_coeff_def) + have RHS': "(a \\<^sub>p b \ (a \\<^sub>p c)) m = + (\x\mset_factors m. a x \ b (m - x) \ a x \ c (m - x))" + proof- + have 0: "(\x. a x \ b (m - x)) \ mset_factors m \ carrier R" + proof + fix x assume "x \ mset_factors m" + then show "a x \ b (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + have 1: "(\x. a x \ c (m - x)) \ mset_factors m \ carrier R" + proof + fix x assume "x \ mset_factors m" + then show "a x \ c (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + then show ?thesis + using 0 1 RHS assms finsum_addf[of "\x. a x \ b (m - x)" "mset_factors m" + "\x. a x \ c (m - x)"] + by metis + qed + show ?thesis using LHS RHS' by auto + qed +qed + +lemma P_ring_ldistr: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + shows " (b \ c) \\<^sub>p a = (b \\<^sub>p a)\ (c \\<^sub>p a)" +proof + fix m + show "((b \ c) \\<^sub>p a) m = ((b \\<^sub>p a)\ (c \\<^sub>p a)) m" + proof- + have RHS: "((b \\<^sub>p a)\ (c \\<^sub>p a)) m = + (\x\mset_factors m. b x \ a (m - x)) \ (\x\mset_factors m. c x \ a (m - x))" + unfolding indexed_padd_def P_ring_mult_def by auto + have LHS: "((b \ c) \\<^sub>p a) m= + ( \x\mset_factors m. b x \ a (m - x) \ c x \ a (m - x))" + unfolding indexed_padd_def P_ring_mult_def + by (meson assms(1) assms(2) assms(3) l_distr local.ring_axioms ring.carrier_coeff_def) + have RHS': "((b \\<^sub>p a)\ (c \\<^sub>p a)) m = + (\x\mset_factors m. b x \ a (m - x) \ c x \ a (m - x))" + proof- + have 0: "(\x. b x \ a (m - x)) \ mset_factors m \ carrier R" + proof + fix x assume "x \ mset_factors m" + then show "b x \ a (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + have 1: "(\x. c x \ a (m - x)) \ mset_factors m \ carrier R" + proof + fix x assume "x \ mset_factors m" + then show "c x \ a (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + then show ?thesis + using 0 1 RHS assms finsum_addf[of "\x. b x \ a (m - x)" "mset_factors m" + "\x. c x \ a (m - x)"] + by metis + qed + show ?thesis using LHS RHS' by auto + qed +qed +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Multiplication Commutes with \texttt{indexed\_pmult}\ +(**************************************************************************************************) +(**************************************************************************************************) + +context ring +begin + +text\This lemma shows how we can write the factors of a monomial $m$ times a variable $i$ in terms + of the factors of m:\ + +lemma mset_factors_add_mset: +"mset_factors (add_mset i m) = mset_factors m \ add_mset i ` (mset_factors m)" +proof + show "mset_factors (add_mset i m) \ mset_factors m \ add_mset i ` mset_factors m" + proof fix x + assume A: "x \ mset_factors (add_mset i m)" + show "x \ mset_factors m \ add_mset i ` mset_factors m" + proof(cases "i \# x") + case True + then have "x - {#i#} \# m" + using A INTEG.R.monom_divides_factors subset_eq_diff_conv by force + then show ?thesis + by (metis INTEG.R.monom_divides_factors True UnI2 add_mset_remove_trivial image_iff multi_member_split) + next + case False + have 0: "x \# add_mset i m" + using A INTEG.R.monom_divides_factors by blast + hence "x \# m" + using mset_subset_eqI[of x m] 0 False + by (metis count_add_mset count_greater_zero_iff count_inI less_Suc_eq_le + subseteq_mset_def union_single_eq_member) + thus ?thesis + using INTEG.R.monom_divides_factors by blast + qed + qed + show "mset_factors m \ add_mset i ` mset_factors m \ mset_factors (add_mset i m)" + proof fix x assume A: "x \ mset_factors m \ add_mset i ` mset_factors m" + have "x \# add_mset i m" + proof(cases "x \ mset_factors m") + case True + then have "x \# m" + by (simp add: INTEG.R.monom_divides_factors) + hence "(\a. count x a \ count (add_mset i m) a)" + using mset_subset_eq_count[of x m] count_add_mset[of i m] + nat_le_linear not_less_eq_eq by fastforce + thus ?thesis + using mset_subset_eqI by blast + next + case False + then obtain n where n_def: "n \ mset_factors m \ x = add_mset i n" + using A by blast + then have "x \# add_mset i m" + by (simp add: INTEG.R.monom_divides_factors) + then show ?thesis + by simp + qed + thus "x \ mset_factors (add_mset i m)" + by (simp add: INTEG.R.monom_divides_factors) + qed +qed + +end +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Associativity of Polynomial Multiplication.\ +(**************************************************************************************************) +(**************************************************************************************************) +context ring +begin + +lemma finsum_eq: + assumes "f \ S \ carrier R" + assumes "g \ S \ carrier R" + assumes "(\ x \ S. f x) = (\ x \ S. g x)" + shows " finsum R f S = finsum R g S" + by (metis assms(1) assms(3) finsum_cong' restrict_apply') + +lemma finsum_eq_induct: + assumes "f \ S \ carrier R" + assumes "g \ T \ carrier R" + assumes "finite S" + assumes "finite T" + assumes "bij_betw h S T" + assumes "\s. s \ S \ f s = g (h s)" + shows "finite U \ U \ S \ finsum R f U = finsum R g (h ` U)" + apply(induct rule: finite_induct) + apply (simp; fail) +proof- + fix x + fix F :: "'c set" + assume F_fin: "finite F" + assume x_not_in: "x \ F" + assume IH: "(F \ S \ finsum R f F = finsum R g (h ` F)) " + show "insert x F \ S \ finsum R f (insert x F) = finsum R g (h ` insert x F)" + proof- + assume A: "insert x F \ S" + then have F_sub: "F \S" + by simp + have x_in: "x \ S" + using A by blast + have fx_in: "f x \ carrier R" + using x_in assms(1) + by auto + have ghx_fx_eq: + "f x = g (h x)" + using x_in assms + by blast + have I: "finsum R f F = finsum R g (h ` F)" + using F_sub IH by blast + have f_fact: "f \ F \ carrier R " + using assms(1) F_sub + by blast + have "finsum R f (insert x F) = (f x) \\<^bsub>add_monoid R\<^esub> (finsum R f F) " + using F_fin x_not_in comm_monoid.finprod_insert[of "(add_monoid R)" F x f ] + unfolding finsum_def + using f_fact fx_in local.add.comm_monoid_axioms + by auto + have "finsum R g (h ` insert x F) = finsum R g (insert (h x) (h ` F))" + by simp + then have "finsum R g (h ` insert x F) = (g (h x)) \\<^bsub>add_monoid R\<^esub> finsum R g (h ` F)" + proof- + have 0: "finite (h ` F)" + by (simp add: F_fin) + have 1: "h x \ h ` F" + proof- + have 10: "bij_betw h F (h` F)" + using assms(5) F_sub bij_betw_subset + by blast + show ?thesis + proof + assume "h x \ h ` F " + then obtain s where s_def: "s \ F \ h s = h x" + using 10 + by auto + have "s \ S" + using F_sub s_def by blast + then have "s = x" + using x_in assms(5) + s_def + unfolding bij_betw_def inj_on_def + by blast + then show False using x_not_in + using s_def by blast + qed + qed + have 2: "g \ h ` F \ carrier (add_monoid R)" + proof + fix y + assume Ay: "y \ h ` F" + then obtain s where s_def: "y = h s \ s \ F" + by blast + then have s_S: "s \ S" + using F_sub by blast + have "h ` F \ T" + using F_sub assms(5) + unfolding bij_betw_def + by blast + then have "g y \ carrier R" + using Ay assms(2) + by blast + then show "g y \ carrier (add_monoid R)" + by simp + qed + have "g (h x) \ carrier (add_monoid R)" + using fx_in ghx_fx_eq + by auto + then show ?thesis + using 0 1 2 comm_monoid.finprod_insert[of "(add_monoid R)" "(h ` F)" "h x" g ] + unfolding finsum_def + by (simp add: local.add.comm_monoid_axioms) + qed + then have "finsum R g (h ` insert x F) = f x \\<^bsub>add_monoid R\<^esub> (finsum R f F)" + using I ghx_fx_eq by auto + then show ?thesis + by (simp add: \finsum R f (insert x F) = f x \\<^bsub>add_monoid R\<^esub> finsum R f F\) + qed +qed + +lemma finsum_bij_eq: + assumes "f \ S \ carrier R" + assumes "g \ T \ carrier R" + assumes "finite S" + assumes "bij_betw h S T" + assumes "\s. s \ S \ f s = g (h s)" + shows "finsum R f S = finsum R g T" +proof- + have 0: "finite T" + using assms bij_betw_finite + by blast + have 1: "(\s. s \ S \ f s = g (h s)) " + using assms + by blast + have "(\s. s \ S \ f s = g (h s)) \ finite S \ S \ S \ finsum R f S = finsum R g (h ` S)" + using 0 assms finsum_eq_induct[of f S g T h S ] + by blast + then have "finsum R f S = finsum R g (h ` S)" + using assms(5) assms(3) + by blast + then show ?thesis + using assms(5) assms(4) bij_betw_imp_surj_on + by blast +qed + +lemma monom_comp: + assumes "x \# m" + assumes "y \# m - x" + shows "x \# m - y" +using assms + by (metis add_diff_cancel_left' subset_eq_diff_conv + subset_mset.le_diff_conv2 subset_mset.order_refl subset_mset.order_trans) + +lemma monom_comp': + assumes "x \# m" + assumes "y = m - x" + shows "x = m - y" + using assms + by (metis add_diff_cancel_right' subset_mset.add_diff_inverse) + +text \ + This lemma turns iterated sums into sums over a product set. The first lemma is only a technical + phrasing of \texttt{double\_finsum'} to facilitate an inductive proof, and likely can and should + be simplified. +\ + +lemma double_finsum_induct: + assumes "finite S" + assumes "\s. s \ S \ finite (F s)" + assumes "P = (\S. {(s, t). s \ S \ t \ (F s)})" + assumes "\s y. s \ S \ y \ (F s) \ g s y \ carrier R" + shows "finite T \ T \ S \ finsum R (\s. finsum R (g s) (F s)) T = + finsum R (\ c. g (fst c) (snd c)) (P T)" + apply(induct rule: finite_induct) + apply (simp add: assms(3); fail) +proof- + fix x + fix T :: "'c set" + assume AT: "finite T" + assume x_notin: "x \ T" + assume IH: "(T \ S \ (\s\T. finsum R (g s) (F s)) = (\c\P T. g (fst c) (snd c)))" + assume A: "insert x T \S" + show " (\s\insert x T. finsum R (g s) (F s)) = (\c\P (insert x T). g (fst c) (snd c))" + proof- + have 0: "(\s. finsum R (g s) (F s)) \ T \ carrier R" + proof + fix v + assume A0: "v \ T" + then have A1: "v \ S" + using A by blast + then show "finsum R (g v) (F v) \ carrier R" + proof- + have 00: "finite (F v)" + using assms A + using \v \ T\ by blast + have "g v \ F v \ carrier R" + proof + fix l assume "l \ F v" + then show "g v l \ carrier R" + using A1 assms(4)[of v l] + by blast + qed + then show ?thesis + using finsum_closed by blast + qed + qed + have 1: "finsum R (g x) (F x) \ carrier R" + proof- + have "x \ S" + using A by blast + then show ?thesis using assms(4) finsum_closed[of "(g x)" "F x"] + by blast + qed + have 2: " (\s\insert x T. finsum R (g s) (F s)) = finsum R (g x) (F x) \ (\s\ T. finsum R (g s) (F s))" + using 0 1 finsum_insert[of T x "(\s. finsum R (g s) (F s))"] AT x_notin + by blast + have 3: "P (insert x T) = P {x} \ P T" + proof + show "P (insert x T) \ P {x} \ P T" + proof + fix y + assume "y \ P (insert x T)" + then show "y \ P {x} \ P T" + using assms + by blast + qed + show "P {x} \ P T \ P (insert x T)" + proof + fix y + assume Ay: "y \ P {x} \ P T" + show "y \ P (insert x T)" + proof(cases "y \ P {x}") + case True + then have "y \ ({(s, t). s \ {x} \ t \ F s})" + using assms(3) by auto + then have "fst y = x \ snd y \ F x" + using Product_Type.Collect_case_prodD by auto + then show ?thesis using assms(3) + using fst_def by auto + next + case False + then have "y \ P T" + using Ay + by blast + then have "y \ ({(s, t). s \ T \ t \ F s})" + using assms(3) by blast + then obtain s t where st_def: "y = (s, t) \s \ T \ t \ F s" + by blast + then have "y = (s, t) \s \ (insert x T) \ t \ F s" + by blast + then show ?thesis using assms(3) + by simp + qed + qed + qed + have 4: "P {x} \ P T = {}" + proof + show "P {x} \ P T \ {}" + proof + fix y + assume B: "y \ P {x} \ P T" + then have "fst y = x" + proof- + have "y \ {(s, t). s \ {x} \ t \ F s}" + using assms(3) B by auto + then obtain s t where "y = (s, t) \ s \ {x} \ t \ F s" + by blast + then show ?thesis + by auto + qed + have "fst y \ T" + proof- + have "y \ {(s, t). s \ T \ t \ F s}" + using assms(3) B by auto + then obtain s t where "y = (s, t) \ s \ T \ t \ F s" + by blast + then show ?thesis + by simp + qed + then have False + using x_notin + by (simp add: \fst y = x\) + then show "y \ {}" + by simp + qed + show "{} \ P {x} \ P T" + by simp + qed + have 5: "(\c. g (fst c) (snd c)) \ P {x} \ carrier R" + proof + fix y + assume X0: "y \ P {x}" + obtain s t where st_def: "y = (s, t) \ (s, t) \ P {x}" + by (metis X0 old.prod.exhaust) + then have st: "s = x \ t \ F x" + using assms(3) by blast + then have "g (fst y) (snd y) = g x t \t \ F x" + by (simp add: st_def) + then show "g (fst y) (snd y) \ carrier R" + using assms(4)[of x t] A + by simp + qed + have 6: "(\c. g (fst c) (snd c)) \ P T \ carrier R" + proof + fix y + assume X0: "y \ P T" + obtain s t where st_def: "y = (s, t) \ (s, t) \ P T" + by (metis X0 old.prod.exhaust) + then have st: "s \ T \ t \ F s" + using assms(3) by blast + then have "g (fst y) (snd y) = g s t \t \ F s" + by (simp add: st_def) + then show "g (fst y) (snd y) \ carrier R" + using assms(4)[of s t] A st + by auto + qed + have 07: "\x. x \ S \ finite (P {x})" + proof- + fix x + assume "x \ S" + have "bij_betw snd (P {x}) (F x)" + unfolding bij_betw_def + proof + show "inj_on snd (P {x})" + proof + fix a b + assume Aa: "a \ P {x}" + assume Ab: "b \ P {x}" + have 0: "fst a = x" + proof- + have" a \ {(s, t). s \ {x} \ t \ F s} " + using Aa assms(3) + by blast + then obtain s t where st_def: "a = (s, t) \ s \ {x} \ t \ F s" + by blast + then show ?thesis + by auto + qed + have 1: "fst b = x" + proof- + have" b \ {(s, t). s \ {x} \ t \ F s} " + using Ab assms(3) + by blast + then obtain s t where st_def: "b = (s, t) \ s \ {x} \ t \ F s" + by blast + then show ?thesis + by auto + qed + show "snd a = snd b \ a = b" + using 0 1 + by (simp add: prod_eqI) + qed + show "snd ` P {x} = F x" + proof + show "snd ` P {x} \ F x" + proof + fix y + assume 0: "y \ snd ` P {x}" + then obtain q where q_def: " q \ P {x} \ y = snd q" + by blast + then obtain s t where st: "q = (s, t) \ s \ {x} \ t \ F s" + using assms(3) by blast + then have 1: "s = x" + by blast + have 2: "snd q = t" + by (simp add: st) + then show "y \ F x" + using q_def st by blast + qed + show "F x \ snd ` P {x}" + proof + fix y + assume "y \ F x" + then have C: "(x, y) \ P {x}" + using assms(3) + by simp + then have "y = snd (x, y)" + by auto + then show "y \ snd ` P {x}" + using C + by blast + qed + qed + qed + then show "finite (P {x})" + using assms(2)[of x] bij_betw_finite + \x \ S\ + by blast + qed + have 7: "finite (P {x})" + using 07 A + by blast + have 8: "finite (P T)" + proof- + have "\D. finite D \ D \ S \ finite (P D)" + proof- + fix D + show "finite D \ D \ S \ finite (P D)" + apply(induct rule: finite_induct) + proof- + have "P {} = {}" using assms(3) + by blast + then show "finite (P {})" + by auto + show "\x F. finite F \ x \ F \ (F \ S \ finite (P F)) \ insert x F \ S \ finite (P (insert x F))" + proof- + fix x + fix Q :: "'c set " + assume fin: "finite Q" + assume notin: "x \ Q" + assume fin_pf: "(Q \ S \ finite (P Q))" + assume I: "insert x Q \ S " + show "finite (P (insert x Q))" + proof- + have x_in: "x \ S" + using I by blast + have 0: "(P (insert x Q)) \ (P Q) \ P {x}" + proof + fix y + assume "y \ P (insert x Q)" + obtain s t where st: "y = (s, t) \ s \ (insert x Q) \ t \ F s" + using assms(3) x_in \y \ P (insert x Q)\ + by blast + show "y \ (P Q) \ P {x}" + proof(cases "s \ Q") + case True + then have "y \ P Q" + using st assms(3) + by simp + then show ?thesis using st assms(3) + by blast + next + case False + then have "s = x" + using st by blast + then have "s \{x} \ t \ F s" + using st by blast + then have "y \ P {x}" + using st assms(3) + by auto + then show ?thesis by auto + qed + qed + have 1: "finite (P Q)" + using I fin_pf by blast + have "finite (P {x})" + using "07" x_in by blast + then show ?thesis using 0 1 + using finite_subset by auto + qed + qed + qed + qed + then show ?thesis + using A AT by blast + qed + have 9: "(\q\P (insert x T). g (fst q) (snd q)) = + (\q\P T. g (fst q) (snd q)) \ (\q\P {x}. g (fst q) (snd q))" + using 8 7 6 5 4 3 finsum_Un_disjoint[of "P {x}" "P T" "\q. g (fst q) (snd q)"] + by (simp add: "7" \\finite (P {x}); finite (P T); P {x} \ P T = {}; (\p. g (fst p) (snd p)) + \ P {x} \ carrier R; (\p. g (fst p) (snd p)) \ P T \ carrier R\ \ + (\p\P {x} \ P T. g (fst p) (snd p)) = (\p\P {x}. g (fst p) (snd p)) \ + (\p\P T. g (fst p) (snd p))\ add.m_comm) + have 10: "(\p\P (insert x T). g (fst p) (snd p)) + = (\s\T. finsum R (g s) (F s))\ (\p\P {x}. g (fst p) (snd p))" + using IH 9 A + by auto + have 11: "(\p\P {x}. g (fst p) (snd p)) = finsum R (g x) (F x)" + proof- + obtain h :: "('c \ 'd) \ 'd" where h_def: "h = snd" + by simp + have 110: "bij_betw h (P {x}) (F x)" + unfolding bij_betw_def + proof + show "inj_on h (P {x})" + unfolding inj_on_def + proof + fix q + assume Ap: "q \ P {x}" + show " \y\P {x}. h q = h y \ q = y" + proof + fix y + assume q_def: "y \ P {x}" + then have C0: "fst y = x" + using assms(3) + by (simp add: case_prod_beta) + have C1: "fst q = x" + using assms(3) Ap + by (simp add: case_prod_beta) + show "h q = h y \ q = y" + using C0 C1 h_def + by (simp add: prod_eq_iff) + qed + qed + show "h ` P {x} = F x" + proof + show "h ` P {x} \ F x" + proof + fix y + assume "y \ h ` P {x}" + then obtain d where d_def: "y = h d \ d \ P {x}" + by blast + then obtain s t where st_def: "d = (s, t)" + by (meson surj_pair) + then have "s = x \ t \ F x" + using assms(3) d_def + by blast + then show "y \ F x" + using assms(3) + by (simp add: d_def h_def st_def) + qed + show "F x \ h ` P {x}" + proof + fix y + assume E0: "y \ F x" + have E1: "y = h (x , y)" + by (simp add: h_def) + have "(x, y) \ P {x}" + using E0 assms(3) by blast + then show "y \ h ` P {x} " + using E1 assms(3) by blast + qed + qed + qed + have 111: "g x \ F x \ carrier R " + proof + fix y + assume "y \ F x" + then show "g x y \ carrier R" + using assms A + by blast + qed + have 112: "(\s. s \ P {x} \ g (fst s) (snd s) = g x (h s))" + proof- + fix s + assume "s \ P{x}" + then have "s \ {(s, t). s = x \ t \ F s}" + using assms(3) by blast + then obtain t where "s = (x, t) \ t \ F x" + using assms(3) + by blast + then show "g (fst s) (snd s) = g x (h s)" + by (simp add: h_def) + qed + + + + show ?thesis using 5 7 110 111 112 finsum_bij_eq[of "\p. g (fst p) (snd p)" "P {x}" "g x" "F x" h ] + by auto + qed + have 12: "(\p\P (insert x T). g (fst p) (snd p)) + = (\s\T. finsum R (g s) (F s))\ finsum R (g x) (F x)" + using 10 11 by auto + then show ?thesis using 2 + by (simp add: "2" "0" "1" add.m_comm) + qed +qed + +lemma double_finsum: + assumes "finite S" + assumes "\s. s \ S \ finite (F s)" + assumes "P = {(s, t). s \ S \ t \ (F s)}" + assumes "\s y. s \ S \ y \ (F s) \ g s y \ carrier R" + shows "finsum R (\s. finsum R (g s) (F s)) S = + finsum R (\ p. g (fst p) (snd p)) P" +proof- + obtain P' where P'_def: "P' = (\S. {(s, t). s \ S \ t \ (F s)})" + by simp + have "finsum R (\s. finsum R (g s) (F s)) S = + finsum R (\ p. g (fst p) (snd p)) (P' S)" + using double_finsum_induct[of S F P' g S] + assms P'_def + by blast + then show ?thesis + using P'_def assms + by blast +qed + +end + +text\ + The product index set for the double sums in the coefficients of the + $((a \otimes_p b) \otimes_p c)$. It is the set of pairs $(x,y)$ of monomials, such that + $x$ is a factor of monomial $m$, and $y$ is a factor of monomial $x$. +\ + +definition right_cuts :: "'c monomial \ ('c monomial \ 'c monomial) set" where +"right_cuts m = {(x, y). x \# m \ y \# x}" + +context ring +begin + +lemma right_cuts_alt_def: +"right_cuts m = {(x, y). x \ mset_factors m \ y \ mset_factors x}" + unfolding mset_factors_def right_cuts_def + by simp + +lemma right_cuts_finite: +"finite (right_cuts m)" +proof- + have "finite (mset_factors m \ mset_factors m)" + using mset_factors_finite + by blast + have "right_cuts m \ (mset_factors m \ mset_factors m)" + proof + fix p + assume p_def: "p \ right_cuts m" + obtain x y where xy: "p = (x , y) \ x \# m \ y \#x" + using p_def unfolding right_cuts_def + by blast + then have "x \ mset_factors m \ y \ mset_factors m" + using monom_divides_factors + by auto + then show "p \ (mset_factors m \ mset_factors m)" + by (simp add: xy) + qed + then show ?thesis + using \finite (mset_factors m \ mset_factors m)\ finite_subset + by blast +qed + +lemma assoc_aux0: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + assumes "g = (\x y. a x \ (b y \ c (m - x - y)))" + shows "\s y. s \ mset_factors m \ y \ mset_factors (m - x) + \ g s y \ carrier R" + using assms carrier_coeffE by blast + +lemma assoc_aux1: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + assumes "g = (\x y. (a y \ b (x - y)) \ c (m - x))" + shows "\s y. s \ mset_factors m \ y \ mset_factors x \ g s y \ carrier R" + using assms carrier_coeffE by blast +end + +text\ + The product index set for the double sums in the coefficients of the + $(a \otimes_p (b \otimes_p c))$. It is the set of pairs $(x,y)$ such that $x$ is a factor + of $m$ and $y$ is a factor of $m/x$. +\ + +definition left_cuts :: "'c monomial \ ('c monomial \ 'c monomial) set" where +"left_cuts m = {(x, y). x \#m \ y \# (m - x)}" + +context ring +begin + +lemma left_cuts_alt_def: +"left_cuts m = {(x, y). x \ mset_factors m \ y \ mset_factors (m - x)}" + unfolding mset_factors_def left_cuts_def + by simp + +text\This lemma witnesses the bijection between left and right cuts for the proof of associativity:\ + +lemma left_right_cuts_bij: +"bij_betw (\ (x,y). (x + y, x)) (left_cuts m) (right_cuts m)" + unfolding bij_betw_def right_cuts_def left_cuts_def +proof + show "inj_on (\(x, y). (x + y, x)) {(x, y). x \# m \ y \# m - x}" + unfolding inj_on_def + by auto + show "(\(x, y). (x + y, x)) ` {(x, y). x \# m \ y \# m - x} = {(x, y). x \# m \ y \# x}" + proof + show "(\(x, y). (x + y, x)) ` {(x, y). x \# m \ y \# m - x} \ {(x, y). x \# m \ y \# x}" + proof + fix p + assume "p \ (\(x, y). (x + y, x)) ` {(x, y). x \# m \ y \# m - x}" + then obtain a b where ab: "a \# m \ b \# m - a \ p = (\(x, y). (x + y, x)) (a, b)" + by blast + then have "p = (a + b, a)" + by simp + then show "p \ {(x, y). x \# m \ y \# x}" + using ab + by (metis (no_types, lifting) case_prodI mem_Collect_eq mset_subset_eq_add_left + subset_mset.le_diff_conv2 union_commute) + qed + show "{(x, y). x \# m \ y \# x} \ (\(x, y). (x + y, x)) ` {(x, y). x \# m \ y \# m - x}" + proof + fix p + assume p_def: "p \ {(x, y). x \# m \ y \# x}" + then obtain a b where ab: "p = (a, b) \ a \# m \ b \# a" + by blast + then have "p = (\(x, y). (x + y, x)) (b, a - b)" + using ab + by simp + then show " p \ (\(x, y). (x + y, x)) ` {(x, y). x \# m \ y \# m - x}" + by (metis (mono_tags, lifting) ab case_prodI image_eqI mem_Collect_eq + subset_mset.diff_add subset_mset.dual_order.trans subset_mset.le_diff_conv2) + qed + qed +qed + +lemma left_cuts_sum: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + shows "(a \\<^sub>p (b \\<^sub>p c)) m = (\p \ left_cuts m. a (fst p) \ (b (snd p) \ c (m - (fst p) - (snd p))))" +proof- + have U: "(a \\<^sub>p (b \\<^sub>p c)) m = (\x\mset_factors m. (\xa\mset_factors (m - x). a x \ (b xa \ c (m - x - xa))))" + unfolding P_ring_mult_def + proof- + obtain f where f_def: + "f = (\x . a x \ (\xa\mset_factors (m - x). b xa \ c (m - x - xa)))" + by simp + obtain g where g_def: + "g = (\x . \xa\mset_factors (m - x). a x \ (b xa \ c (m - x - xa)))" + by simp + have 0: "restrict f (mset_factors m) = restrict g (mset_factors m)" + proof + fix x + show "restrict f (mset_factors m) x = restrict g (mset_factors m) x" + proof(cases "x \ mset_factors m") + case True + have T0: "restrict f (mset_factors m) x = a x \ (\xa\mset_factors (m - x). b xa \ c (m - x - xa))" + using f_def True by auto + have T1: "restrict g (mset_factors m) x = (\xa\mset_factors (m - x). a x \ (b xa \ c (m - x - xa)))" + using True g_def by auto + show "restrict f (mset_factors m) x = restrict g (mset_factors m) x" + using assms finsum_rdistr[of "mset_factors (m - x)" "a x" "\ xa. b xa \ c (m - x - xa)"] + by (metis (mono_tags, lifting) mset_factors_finite Pi_I T0 T1 carrier_coeffE m_closed) + next + case False + then have "restrict f (mset_factors m) x = undefined" using f_def + by (simp add: restrict_def) + have "restrict g (mset_factors m) x = undefined" using g_def False + using restrict_def by auto + then show ?thesis using f_def g_def + using \restrict f (mset_factors m) x = undefined\ + by auto + + qed + qed + have 1: "f \ mset_factors m \ carrier R" + using f_def assms + by (simp add: carrier_coeffE) + have 2: "g \ mset_factors m \ carrier R" + using g_def assms + by (simp add: carrier_coeffE) + show "(\x\mset_factors m. a x \ (\xa\mset_factors (m - x). b xa \ c (m - x - xa))) = + (\x\mset_factors m. \xa\mset_factors (m - x). a x \ (b xa \ c (m - x - xa)))" + using f_def g_def finsum_eq[of "f" + "mset_factors m" "g"] 0 1 2 + by blast + qed + have 0: "(\s. s \ mset_factors m \ finite (mset_factors (m - s)))" + by simp + have 1: "finite (mset_factors m)" + by simp + have 2: "(\s y. s \ mset_factors m \ y \ mset_factors (m - s) \ a s \ (b y \ c (m - s - y)) \ carrier R)" + using assms assoc_aux0[of a b c ] + by blast + have "(\x\mset_factors m. (\xa\mset_factors (m - x). a x \ (b xa \ c (m - x - xa)))) = + (\p \ left_cuts m. a (fst p) \ (b (snd p) \ c (m - (fst p) - (snd p)))) " + using assms left_cuts_alt_def[of m] 0 1 2 + double_finsum[of "mset_factors m" "\x. mset_factors (m - x)" "left_cuts m" "(\x y. a x \ (b y \ c (m - x - y)))"] + by blast + then show ?thesis using U + by auto +qed + +lemma right_cuts_sum: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + shows "(a \\<^sub>p b \\<^sub>p c) m = (\p \ right_cuts m. a (snd p) \ (b ((fst p) -(snd p)) \ c (m - (fst p))))" +proof- + have 0: "finite (mset_factors m)" + by simp + have 1: "(\s. s \ mset_factors m \ finite (mset_factors s))" + by auto + have 2: "right_cuts m = {(s, t). s \ mset_factors m \ t \ mset_factors s}" + unfolding right_cuts_def + by (simp add: monom_divides_factors) + have 3: "(\s y. s \ mset_factors m \ y \ mset_factors s \ a y \ b (s - y) \ c (m - s) \ carrier R)" + using assoc_aux1 assms(1) assms(2) assms(3) + by blast + have 4: "(\s\mset_factors m. (\y\mset_factors s. a y \ b (s - y) \ c (m - s))) = + (\p\right_cuts m. a (snd p) \ b (fst p - snd p) \ c (m - fst p))" + using 0 1 2 3 + double_finsum[of "mset_factors m" _ "right_cuts m" + "(\x y. (a y \ b (x - y)) \ c (m - x))"] + by auto + have 5: "(\x\mset_factors m. (\xa\mset_factors x. a xa \ b (x - xa)) \ c (m - x)) = + (\x\mset_factors m. \xa\mset_factors x. a xa \ b (x - xa) \ c (m - x))" + proof- + obtain f where f_def: "f =( \x. (\xa\mset_factors x. a xa \ b (x - xa)) \ c (m - x))" + by simp + obtain g where g_def: "g = (\x. \xa\mset_factors x. a xa \ b (x - xa) \ c (m - x))" + by simp + have 50: "\s. s \ (mset_factors m) \ f s = g s" + proof- + fix x + assume As: "x \ mset_factors m" + show "f x = g x" + proof- + have f_eq: "f x = (\xa\mset_factors x. a xa \ b (x - xa)) \ c (m - x)" + using f_def + by auto + have g_eq: "g x = (\xa\mset_factors x. a xa \ b (x - xa) \ c (m - x))" + using g_def + by auto + have f_eq': "f x = (\xa\mset_factors x. (a xa \ b (x - xa)) \ c (m - x))" + using f_eq finsum_ldistr[of "mset_factors x" "c (m - x)" "\xa. (a xa \ b (x - xa))" ] assms + by (simp add: \\finite (mset_factors x); c (m - x) \ carrier R; (\xa. a xa \ b (x - xa)) \ mset_factors x \ carrier R\ \ (\i\mset_factors x. a i \ b (x - i)) \ c (m - x) = (\i\mset_factors x. a i \ b (x - i) \ c (m - x))\ Pi_I carrier_coeffE) + then show "f x = g x" + by (simp add: g_eq) + qed + qed + have 51: "f \ mset_factors m \ carrier R" + proof + fix x + assume "x \ mset_factors m" + have "(\xa\mset_factors x. a xa \ b (x - xa)) \ c (m - x )\ carrier R" + using assms carrier_coeffE finsum_closed[of "\xa. a xa \ b (x - xa)" "mset_factors x"] + by blast + then show "f x \ carrier R" using assms f_def by auto + qed + have 52: "g \ mset_factors m \ carrier R" + proof + fix x + assume "x \ mset_factors m" + show "g x \ carrier R" + using assms finsum_closed[of "\xa. a xa \ b (x - xa) \ c (m - x)" "mset_factors x"] g_def + by (metis (no_types, lifting) "50" "51" Pi_iff \x \ mset_factors m\) + qed + + show ?thesis + using 50 51 52 finsum_eq[of f "(mset_factors m) " g] + by (metis (mono_tags, lifting) f_def finsum_cong' g_def) + qed + then have 6: "(\x\mset_factors m. (\xa\mset_factors x. a xa \ b (x - xa)) \ c (m - x)) = + (\p\right_cuts m. a (snd p) \ b (fst p - snd p) \ c (m - fst p))" + by (simp add: "4") + have 7: "(\p\right_cuts m. a (snd p) \ b (fst p - snd p) \ c (m - fst p)) + = (\p\right_cuts m. a (snd p) \ (b (fst p - snd p) \ c (m - fst p)))" + using assms + by (meson local.ring_axioms m_assoc ring.carrier_coeff_def) + then show ?thesis + using assms 5 6 + unfolding P_ring_mult_def + by simp +qed + +text\The Associativity of Polynomial Multiplication:\ + +lemma P_ring_mult_assoc: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + assumes "carrier_coeff c" + shows "a \\<^sub>p (b \\<^sub>p c) = (a \\<^sub>p b) \\<^sub>p c" +proof + fix m + show "(a \\<^sub>p (b \\<^sub>p c)) m = (a \\<^sub>p b \\<^sub>p c) m" + proof- + obtain f where f_def: "f = (\p. a (snd p) \ (b ((fst p) -(snd p)) \ c (m - (fst p))))" + by simp + obtain g where g_def: "g = (\p. a (fst p) \ (b (snd p) \ c (m - (fst p) - (snd p))))" + by simp + have f_dom: "f \ right_cuts m \ carrier R" + using assms f_def unfolding right_cuts_def + by (simp add: carrier_coeffE) + have g_dom: "g \ left_cuts m \ carrier R" + using assms g_def unfolding left_cuts_def + by (simp add: carrier_coeffE) + have 0: "finite (right_cuts m)" + by (simp add: right_cuts_finite) + have 1: "bij_betw (\ (x,y). (x + y, x)) (left_cuts m) (right_cuts m)" + by (simp add: left_right_cuts_bij) + have 2: "(\s. s \ left_cuts m \ g s = f (case s of (x, y) \ (x + y, x)))" + proof- + fix s + assume "s \ left_cuts m" + then obtain x y where xy: "s = (x, y) \ x \# m \ y \# m - x" + using left_cuts_def + by blast + then have g_eq: "g s = a x \ (b y \ c (m - x - y))" + using g_def fst_conv + by auto + have f_eq: "f (case s of (x, y) \ (x + y, x)) = f (x + y, x)" + by (simp add: xy) + then have f_eq': "f (case s of (x, y) \ (x + y, x)) = a x \ (b ((x + y) - x) \ c (m - (x + y)))" + using f_def + by simp + then show "g s = f (case s of (x, y) \ (x + y, x))" + by (simp add: g_eq) + qed + have 3: "finsum R g (left_cuts m) = finsum R f (right_cuts m)" + using 0 1 2 finsum_bij_eq[of g "left_cuts m" f "right_cuts m" "\ (x,y). (x + y, x)" ] + using bij_betw_finite f_dom g_dom by blast + then show ?thesis using assms right_cuts_sum left_cuts_sum + by (metis (mono_tags, lifting) f_def f_dom finsum_cong' g_def g_dom) + qed +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Commutativity of Polynomial Multiplication\ +(**************************************************************************************************) +(**************************************************************************************************) + +context ring +begin + +lemma mset_factors_bij: +"bij_betw (\x. m - x) (mset_factors m) (mset_factors m)" + apply(rule bij_betwI') + apply (metis monom_comp' monom_divides_factors) + apply (simp add: monom_divides_factors) + by (meson monom_comp' monom_divides_factors diff_subset_eq_self) + +lemma(in cring) P_ring_mult_comm: + assumes "carrier_coeff a" + assumes "carrier_coeff b" + shows "a \\<^sub>p b = b \\<^sub>p a" +proof + fix m + show "(a \\<^sub>p b) m = (b \\<^sub>p a) m" + unfolding P_ring_mult_def + apply (rule finsum_bij_eq[of "\ x. a x \ b (m - x)" "mset_factors m" + "\x. b x \ a (m - x)" "mset_factors m" "\x. m - x"]) + proof + show "\x. x \ mset_factors m \ a x \ b (m - x) \ carrier R" + proof- + fix x + assume "x \ mset_factors m" + show "a x \ b (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + show "(\x. b x \ a (m - x)) \ mset_factors m \ carrier R" + proof + fix x + assume "x \ mset_factors m" + show "b x \ a (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + show "finite (mset_factors m)" + by simp + show "bij_betw ((-) m) (mset_factors m) (mset_factors m)" + by (simp add: mset_factors_bij) + show "\s. s \ mset_factors m \ a s \ b (m - s) = b (m - s) \ a (m - (m - s))" + proof- + fix s + assume "s \ mset_factors m" + then show "a s \ b (m - s) = b (m - s) \ a (m - (m - s))" + using assms carrier_coeffE + by (metis local.ring_axioms m_comm ring.monom_comp' ring.monom_divides_factors) + qed + qed +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Closure properties for multiplication\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\Building monomials from polynomials:\ + +lemma indexed_const_P_mult_eq: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "(indexed_const a) \\<^sub>p (indexed_const b) = indexed_const (a \ b)" +proof- + have 0: "monomials_of R (indexed_const a) = (if (a = \) then {} else {{#}})" + unfolding monomials_of_def indexed_const_def + by auto + have 1: "monomials_of R (indexed_const b) = (if (b = \) then {} else {{#}})" + unfolding monomials_of_def indexed_const_def + by auto + show ?thesis + unfolding P_ring_mult_def + proof + fix m:: "'c monomial" + show "(\x\mset_factors m. indexed_const a x \ indexed_const b (m - x)) = indexed_const (a \ b) m " + proof(cases "m = {#}") + case True + then have T0: "mset_factors m = {{#}}" + unfolding mset_factors_def + by simp + have "(\x\mset_factors m. indexed_const a x \ indexed_const b (m - x)) = + indexed_const a {#} \ indexed_const b ({#} - {#}) " + proof- + have 0: "(\x. indexed_const a x \ indexed_const b (m - x)) \ {} \ carrier R" + by blast + have 1: "indexed_const a {#} \ indexed_const b (m - {#}) \ carrier R" + by (simp add: assms(1) assms(2) indexed_const_def) + have 2: "(\x\{{#}}. indexed_const a x \ indexed_const b (m - x)) = + indexed_const a {#} \ indexed_const b (m - {#}) \ (\x\{}. indexed_const a x \ indexed_const b (m - x))" + using True T0 0 1 finsum_insert[of "{}" "{#}" "\x. indexed_const a x \ indexed_const b (m - x)" ] + by (simp add: indexed_const_def) + then show ?thesis + using True T0 0 1 finsum_insert[of "{}" "{#}" "\x. indexed_const a x \ indexed_const b (m - x)" ] + finsum_empty[of "\x. indexed_const a x \ indexed_const b (m - x)"] + by (simp add: \\finite {}; {#} \ {}; (\x. indexed_const a x \ indexed_const b (m - x)) \ {} \ carrier R; indexed_const a {#} \ indexed_const b (m - {#}) \ carrier R\ \ (\x\{{#}}. indexed_const a x \ indexed_const b (m - x)) = indexed_const a {#} \ indexed_const b (m - {#}) \ (\x\{}. indexed_const a x \ indexed_const b (m - x))\ indexed_const_def) + qed + then show ?thesis + by (simp add: True indexed_const_def) + next + case False + then show ?thesis using 0 1 + by (simp add: add.finprod_one_eqI assms(1) assms(2) + indexed_const_def ) + qed + qed +qed + +lemma indexed_const_P_multE: + assumes "P \ indexed_pset I (carrier R)" + assumes "c \ carrier R" + shows "(P \\<^sub>p (indexed_const c)) m = (P m) \ c" + unfolding P_ring_mult_def +proof- + have 3: "m \ mset_factors m - {m}" + by simp + have 4: "finite ((mset_factors m) - {m})" + by simp + have 5: "P m \ indexed_const c (m - m) \ carrier R " + by (metis assms(1) assms(2) cancel_comm_monoid_add_class.diff_cancel + carrier_coeffE indexed_const_def indexed_pset_in_carrier local.ring_axioms + ring.ring_simprules(5) subsetI) + have 0: "(\x. P x \ indexed_const c (m - x)) \ mset_factors m - {m} \ carrier R" + proof + fix x + assume "x \ mset_factors m - {m}" + then show "P x \ indexed_const c (m - x) \ carrier R " + using assms + by (meson carrier_coeffE indexed_const_in_carrier indexed_pset_in_carrier m_closed subsetI) + qed + have 1: "P m \ indexed_const c ( m - m) = (P m \ c)" + by (simp add: indexed_const_def) + have 2: "\x. x \ mset_factors m \ P x \ indexed_const c (m - x) = (if x = m then (P m \ c) else \)" + proof- + fix x + assume "x \ mset_factors m " + then have "indexed_const c (m - x) = (if x = m then c else \)" + unfolding indexed_const_def + by (metis cancel_comm_monoid_add_class.diff_cancel + diff_zero monom_comp' monom_divides_factors) + then show "P x \ indexed_const c (m - x) = (if x = m then (P m \ c) else \)" + by (metis assms(1) carrier_coeffE indexed_pset_in_carrier + local.semiring_axioms semiring.r_null set_eq_subset) + qed + then have "(\x\(mset_factors m) - {m}. P x \ indexed_const c (m - x)) = \" + using assms + by (metis (no_types, lifting) DiffD1 DiffD2 add.finprod_one_eqI singletonI) + then show "(\x\(mset_factors m). P x \ indexed_const c (m - x)) = P m \ c" + using assms 0 1 3 4 5 finsum_insert[of "(mset_factors m) - {m}" + m "\x. P x \ indexed_const c (m - x) "] + by (metis (no_types, lifting) m_factor add.l_cancel_one + insert_Diff_single insert_absorb zero_closed) +qed + +lemma indexed_const_var_mult: + assumes "P \ indexed_pset I (carrier R)" + assumes "c \ carrier R" + assumes "i \ I" + shows "P \ i \\<^sub>p indexed_const c = (P \\<^sub>p (indexed_const c)) \ i " +proof + fix m + show "(P \ i \\<^sub>p indexed_const c) m = (P \\<^sub>p indexed_const c \ i) m" + proof(cases "i \# m") + case True + then have T0: "(P \ i \\<^sub>p indexed_const c) m = (P \ i) m \ c" + using assms indexed_const_P_multE[of "P \ i" I c m] + by (simp add: indexed_pset.indexed_pmult) + then show ?thesis using assms indexed_const_P_multE[of P I c m] + unfolding indexed_pmult_def + using True indexed_const_P_multE by fastforce + next + case False + then have T0: "(P \ i \\<^sub>p indexed_const c) m = (P \ i) m \ c" + using assms indexed_const_P_multE[of "P \ i" I c m] + by (simp add: indexed_pset.indexed_pmult) + then show ?thesis using assms indexed_const_P_multE[of P I c m] + unfolding indexed_pmult_def + using False by auto + qed +qed + +lemma indexed_const_P_mult_closed: + assumes "a \ indexed_pset I (carrier R)" + assumes "c \ carrier R" + shows "a \\<^sub>p (indexed_const c) \ indexed_pset I (carrier R)" + apply(rule indexed_pset.induct[of a I "(carrier R)" ]) + apply (simp add: assms(1); fail) +proof- + show "\k. (k \ carrier R) \ ((indexed_const k) \\<^sub>p (indexed_const c)) \ ((carrier R) [\\<^bsub>I\<^esub>])" + using assms + by (metis indexed_const_P_mult_eq indexed_pset.simps m_closed) + show "\P Q. P \ ((carrier R) [\\<^bsub>I\<^esub>]) \ + P \\<^sub>p indexed_const c \ ((carrier R) [\\<^bsub>I\<^esub>]) \ + Q \ ((carrier R) [\\<^bsub>I\<^esub>]) \ Q \\<^sub>p indexed_const c \ ((carrier R) [\\<^bsub>I\<^esub>]) \ P \ Q \\<^sub>p indexed_const c \ ((carrier R) [\\<^bsub>I\<^esub>])" + using P_ring_ldistr + by (metis assms(2) carrier_coeff_def indexed_const_in_carrier indexed_pset.indexed_padd indexed_pset_in_carrier subsetI) + show "\P i. P \ ((carrier R) [\\<^bsub>I\<^esub>]) \ P \\<^sub>p indexed_const c \ ((carrier R) [\\<^bsub>I\<^esub>]) \ i \ I \ P \ i \\<^sub>p indexed_const c \ ((carrier R) [\\<^bsub>I\<^esub>])" + proof- + fix P i + assume A0: "P \ ((carrier R) [\\<^bsub>I\<^esub>])" + assume A1: " (P \\<^sub>p indexed_const c) \ ((carrier R) [\\<^bsub>I\<^esub>])" + assume A2: "i \ I" + show "P \ i \\<^sub>p indexed_const c \ (carrier R [\\<^bsub>I\<^esub>])" + using assms A0 A1 A2 indexed_const_var_mult local.ring_axioms ring.indexed_pset.simps + by (metis (no_types, opaque_lifting)) + qed +qed + +lemma monom_add_mset: +"mset_to_IP R (add_mset i m) = mset_to_IP R m \ i" + unfolding indexed_pmult_def + by (metis (no_types, opaque_lifting) add_mset_diff_bothsides diff_empty mset_to_IP_def multi_member_split union_single_eq_member) + +text\Monomials are closed under multiplication:\ + +lemma monom_mult: +"mset_to_IP R m \\<^sub>p mset_to_IP R n = mset_to_IP R (m + n)" +proof- + have "(\ma. \x\mset_factors ma. (if x = m then \ else \) \ (if ma - x = n then \ else \)) + = (\ma. (if ma = n + m then \ else \))" + proof + fix ma + show "(\x\mset_factors ma. (if x = m then \ else \) \ (if ma - x = n then \ else \)) = (if ma = n + m then \ else \)" + + proof- + have 0: "\x. x \ mset_factors ma \ + (\ x. (if x = m then \ else \) \ (if ma - x = n then \ else \)) x = + (if (x = m) then (if ma - x = n then \ else \) else \) " + by simp + then have 1: "(\x\((mset_factors ma) - {m}). (if x = m then \ else \) \ (if ma - x = n then \ else \)) = + \" + by (metis (no_types, lifting) add.finprod_one_eqI mem_Collect_eq set_diff_eq singletonI) + have 2: "finite (mset_factors ma - {m}) " + by simp + have 3: "m \ mset_factors ma - {m}" + by simp + have 4: "(\x. (if x = m then \ else \) \ (if ma - x = n then \ else \)) \ mset_factors ma - {m} \ carrier R" + proof + fix x + assume " x \ mset_factors ma - {m}" + show "(if x = m then \ else \) \ (if ma - x = n then \ else \) \ carrier R " + by auto + qed + show ?thesis + proof(cases "m \ (mset_factors ma)") + case True + have T0: " (\x\insert m (mset_factors ma - {m}). (if x = m then \ else \) \ (if ma - x = n then \ else \)) = + (if m = m then \ else \) \ (if ma - m = n then \ else \) \ + (\x\mset_factors ma - {m}. (if x = m then \ else \) \ (if ma - x = n then \ else \))" + using True 1 2 3 4 finsum_insert[of "((mset_factors ma) - {m})" m + "\x. (if x = m then \ else \) \ (if ma - x = n then \ else \)"] + using m_closed one_closed zero_closed by presburger + have T1: "(mset_factors ma) = insert m (mset_factors ma - {m})" + using True + by blast + then have "(\x\(mset_factors ma). (if x = m then \ else \) \ (if ma - x = n then \ else \)) + = (if m = m then \ else \) \ (if ma - m = n then \ else \)" + using T0 T1 1 add.l_cancel_one insert_Diff_single insert_absorb2 l_one mk_disjoint_insert one_closed zero_closed + by presburger + then have "(\x\(mset_factors ma). (if x = m then \ else \) \ (if ma - x = n then \ else \)) + = (if ma - m = n then \ else \)" + by simp + then show ?thesis + by (metis (no_types, lifting) monom_divides_factors True add.commute add_diff_cancel_right' subset_mset.add_diff_inverse) + next + case False + then show ?thesis + by (metis (no_types, lifting) "0" monom_divides_factors add.finprod_one_eqI mset_subset_eq_add_right) + qed + qed + qed + then show ?thesis + unfolding mset_to_IP_def P_ring_mult_def + by (simp add: union_commute) +qed + +lemma poly_index_mult: + assumes "a \ indexed_pset I (carrier R)" + assumes "i \ I" + shows "a \ i = a \\<^sub>p mset_to_IP R {#i#}" +proof + fix m + show "(a \ i) m = (a \\<^sub>p mset_to_IP R {#i#}) m" + proof(cases "i \# m") + case True + have T0: "(a \ i) m = a (m - {#i#})" + by (simp add: True indexed_pmult_def) + have T1: "(a \\<^sub>p mset_to_IP R {#i#}) m = a (m - {#i#})" + proof- + have T10: "(a \\<^sub>p mset_to_IP R {#i#}) m + = (\x\mset_factors m. a x \ mset_to_IP R {#i#} (m - x)) " + unfolding P_ring_mult_def + by auto + have T11: "\x. x \ mset_factors m \ + mset_to_IP R {#i#} (m - x) = (if x = (m - {#i#}) then \ else \)" + unfolding mset_to_IP_def + by (metis Multiset.diff_cancel Multiset.diff_right_commute True diff_single_eq_union + diff_single_trivial diff_zero monom_comp' monom_divides_factors multi_drop_mem_not_eq) + have T12: "\x. x \ mset_factors m \ + a x \ mset_to_IP R {#i#} (m - x) = (if x = (m - {#i#}) then a (m - {#i#}) else \)" + proof- + fix x + assume "x \ mset_factors m" + then show " a x \ mset_to_IP R {#i#} (m - x) = (if x = m - {#i#} then a (m - {#i#}) else \)" + apply(cases "x = m - {#i#}") + apply (metis T0 T11 \x \ mset_factors m\ assms(1) carrier_coeffE empty_subsetI + ideal_is_subalgebra indexed_pset_in_carrier oneideal r_one subalgebra_in_carrier) + by (metis T11 \x \ mset_factors m\ assms(1) carrier_coeffE carrier_is_subalgebra + empty_subsetI indexed_pset_in_carrier r_null subalgebra_in_carrier) + qed + have T13: "(\x\(mset_factors m) - {m - {#i#}}. a x \ mset_to_IP R {#i#} (m - x)) = \" + using T12 + by (metis (no_types, lifting) DiffE add.finprod_one_eqI singletonI) + have T14: "finite (mset_factors m - {m - {#i#}})" + using mset_factors_finite by blast + have T15: "m - {#i#} \ mset_factors m - {m - {#i#}}" + by simp + have T16: " (\x. a x \ mset_to_IP R {#i#} (m - x)) \ mset_factors m - {m - {#i#}} \ carrier R" + proof + fix x + assume "x \ mset_factors m - {m - {#i#}}" + then show " a x \ mset_to_IP R {#i#} (m - x) \ carrier R" + using assms T12 by auto + qed + have "m - (m - {#i#}) = {#i#}" + by (metis monom_comp' True single_subset_iff) + then have T17: " a (m - {#i#}) \ mset_to_IP R {#i#} (m - (m - {#i#})) \ carrier R" + unfolding mset_to_IP_def apply auto using assms + by (meson carrier_coeffE carrier_is_subalgebra empty_subsetI indexed_pset_in_carrier m_closed one_closed subalgebra_in_carrier) + have T18:"(\x\insert (m - {#i#}) (mset_factors m - {m - {#i#}}). a x \ mset_to_IP R {#i#} (m - x)) = + a (m - {#i#}) \ mset_to_IP R {#i#} (m - (m - {#i#})) \ (\x\mset_factors m - {m - {#i#}}. a x \ mset_to_IP R {#i#} (m - x))" + using T12 T13 T14 T15 T16 T17 unfolding P_ring_mult_def + using finsum_insert[of "mset_factors m - {m - {#i#}}" "m - {#i#}" + "\x. a x \ mset_to_IP R {#i#} (m - x)"] + by blast + have T19: "a (m - {#i#}) \ mset_to_IP R {#i#} (m - (m - {#i#})) = a (m - {#i#}) " + proof- + have " (m - (m - {#i#})) = {#i#}" + using True + by (metis monom_comp' single_subset_iff) + then have "mset_to_IP R {#i#} (m - (m - {#i#})) = \" + by (metis mset_to_IP_def) + have "a (m - {#i#}) \ carrier R" + using assms + by (meson carrier_coeffE carrier_is_subalgebra exp_base_closed + indexed_pset_in_carrier one_closed subalgebra_in_carrier) + then show ?thesis using assms + using \mset_to_IP R {#i#} (m - (m - {#i#})) = \\ by auto + qed + have T20: "(\x\insert (m - {#i#}) ((mset_factors m) - {m - {#i#}}). a x \ mset_to_IP R {#i#} (m - x)) = + a (m - {#i#}) \ (\x\mset_factors m - {m - {#i#}}. a x \ mset_to_IP R {#i#} (m - x))" + using T18 T19 by auto + have T21: "insert (m - {#i#}) ((mset_factors m) - {m - {#i#}}) = mset_factors m" + proof- + have "(m - {#i#}) \ mset_factors m" + by (simp add: monom_divides_factors) + then show ?thesis + by blast + qed + then show ?thesis using T13 T20 True unfolding P_ring_mult_def + using T17 T19 by auto + qed + then show ?thesis + by (simp add: T0) + next + case False + have F0: "(a \ i) m = \ " + by (simp add: False indexed_pmult_def) + have F1: " (a \\<^sub>p mset_to_IP R {#i#}) m = \" + unfolding P_ring_mult_def + proof- + have "\x. x \ mset_factors m \ a x \ mset_to_IP R {#i#} (m - x) = \" + proof- + fix x + assume A: "x \ mset_factors m" + have B: "m - x \ {#i#}" using False A + by (metis diff_subset_eq_self single_subset_iff) + then show "a x \ mset_to_IP R {#i#} (m - x) = \" + using assms False unfolding mset_to_IP_def + using carrier_coeffE indexed_pset_in_carrier by fastforce + qed + then show "(\x\mset_factors m. a x \ mset_to_IP R {#i#} (m - x)) = \" + by (simp add: add.finprod_one_eqI) + qed + then show ?thesis + by (simp add: F0) + qed +qed + +lemma mset_to_IP_mult_closed: + assumes "a \ indexed_pset I (carrier R)" + shows "set_mset m \ I \ a \\<^sub>p mset_to_IP R m \ indexed_pset I (carrier R)" +proof(induction m) + case empty + then have "mset_to_IP R {#} = indexed_const \" + unfolding mset_to_IP_def indexed_const_def by auto + then show ?case + by (simp add: \mset_to_IP R {#} = indexed_const \\ assms indexed_const_P_mult_closed) +next + case (add x m) + fix x + fix m :: "'c monomial" + assume A: "set_mset (add_mset x m) \ I" + assume B: "set_mset m \ I \ a \\<^sub>p mset_to_IP R m \ (carrier R [\\<^bsub>I\<^esub>])" + show " a \\<^sub>p mset_to_IP R (add_mset x m) \ (carrier R [\\<^bsub>I\<^esub>])" + proof- + have 0: "x \ I" + using A by auto + have 1: "set_mset m \ I" + using A by auto + have 2: " a \\<^sub>p mset_to_IP R m \ (carrier R [\\<^bsub>I\<^esub>])" + using "1" B by blast + have 3: " a \\<^sub>p mset_to_IP R (add_mset x m) = (a \\<^sub>p mset_to_IP R m) \ x" + proof- + have 30: " a \\<^sub>p mset_to_IP R (add_mset x m) = a \\<^sub>p (mset_to_IP R m \ x)" + using monom_add_mset[of x m] + by auto + have 31: "(a \\<^sub>p mset_to_IP R m) \ x = a \\<^sub>p (mset_to_IP R m \ x)" + proof + fix y + show "(a \\<^sub>p mset_to_IP R m \ x) y = (a \\<^sub>p (mset_to_IP R m \ x)) y " + proof- + have 310: "(a \\<^sub>p mset_to_IP R m \ x) y = (a \\<^sub>p mset_to_IP R m \\<^sub>p mset_to_IP R {#x#}) y" + using poly_index_mult "0" "2" + by fastforce + have 311: " (mset_to_IP R m \ x) = mset_to_IP R m \\<^sub>p mset_to_IP R {#x#}" + using "0" "2" + by (metis add_mset_add_single monom_add_mset monom_mult) + have 312: "carrier_coeff a " + using assms indexed_pset_in_carrier by blast + have 313: "carrier_coeff (mset_to_IP R m)" + by (simp add: carrier_coeff_def mset_to_IP_def) + have 314: "carrier_coeff (mset_to_IP R {#x#})" + by (metis carrier_coeff_def mset_to_IP_def one_closed zero_closed) + show ?thesis + using 310 311 312 313 314 + P_ring_mult_assoc[of a "mset_to_IP R m" "mset_to_IP R {#x#}"] + by simp + qed + qed + then show ?thesis + by (simp add: monom_add_mset) + qed + then show ?thesis + by (simp add: "0" "1" B indexed_pset.indexed_pmult) + qed +qed + +text\ + A predicate which identifies when the variables used in a given polynomial $P$ are only + drawn from a fixed variable set $I$. +\ +abbreviation monoms_in where +"monoms_in I P \ (\m \ monomials_of R P. set_mset m \ I)" + +lemma monoms_of_const: +"monomials_of R (indexed_const k) = (if k = \ then {} else {{#}})" + unfolding indexed_const_def monomials_of_def + by auto + +lemma const_monoms_in: +"monoms_in I (indexed_const k)" + unfolding indexed_const_def monomials_of_def + using count_empty count_eq_zero_iff monomials_ofE subsetI + by simp + +lemma mset_to_IP_indices: + shows "P \ indexed_pset I (carrier R) \ monoms_in I P" + apply(erule indexed_pset.induct[of]) + apply (simp add: const_monoms_in; fail) +proof- + show "\P Q. P \ (carrier R [\\<^bsub>I\<^esub>]) \ + \m\monomials_of R P. set_mset m \ I \ + Q \ (carrier R [\\<^bsub>I\<^esub>]) \ \m\monomials_of R Q. set_mset m \ I \ \m\monomials_of R (P \ Q). set_mset m \ I" + proof + fix P Q + fix m + assume A: "P \ (carrier R [\\<^bsub>I\<^esub>])" "\m\monomials_of R P. set_mset m \ I" + "Q \ (carrier R [\\<^bsub>I\<^esub>])" "\m\monomials_of R Q. set_mset m \ I" + " m \ monomials_of R (P \ Q)" + show "set_mset m \ I" + proof- + have "m \ monomials_of R P \ m \ monomials_of R Q" + using A using monomials_of_add[of P Q] + by blast + then show ?thesis using A + by blast + qed + qed + show "\P i. P \ (carrier R [\\<^bsub>I\<^esub>]) \ \m\monomials_of R P. set_mset m \ I \ i \ I \ \m\monomials_of R (P \ i). set_mset m \ I" + proof + fix P + fix i + fix m + assume A: "P \ (carrier R [\\<^bsub>I\<^esub>])" "\m\monomials_of R P. set_mset m \ I" "i \ I" "m \ monomials_of R (P \ i)" + obtain n where "n \ monomials_of R P \ m = add_mset i n" + using A + by (metis image_iff monomials_of_p_mult') + then show "set_mset m \ I" + using A + by auto + qed +qed + +lemma mset_to_IP_indices': + assumes "P \ indexed_pset I (carrier R)" + assumes "m \ monomials_of R P" + shows "set_mset m \ I" + using assms(1) assms(2) mset_to_IP_indices by blast + +lemma one_mset_to_IP: + "mset_to_IP R {#} = indexed_const \" + unfolding mset_to_IP_def indexed_const_def + by blast + +lemma mset_to_IP_closed: + shows "set_mset m \ I \mset_to_IP R m \ indexed_pset I (carrier R) " + apply(induction m) + apply (simp add: indexed_pset.indexed_const one_mset_to_IP) + by (metis (no_types, lifting) add_mset_commute indexed_pset.simps + monom_add_mset mset_add subset_iff union_single_eq_member) + +lemma mset_to_IP_closed': + assumes "P \ indexed_pset I (carrier R)" + assumes "m \ monomials_of R P" + shows "mset_to_IP R m \ indexed_pset I (carrier R)" + by (meson assms(1) assms(2) mset_to_IP_closed mset_to_IP_indices') + +text\This lemma expresses closure under multiplcation for indexed polynomials.\ + +lemma P_ring_mult_closed: + assumes "carrier_coeff P" + assumes "carrier_coeff Q" + shows "carrier_coeff (P_ring_mult R P Q)" + unfolding carrier_coeff_def +proof + fix m + have "(\x. P x \ Q (m - x)) \ mset_factors m \ carrier R" + proof + fix x + assume "x \ mset_factors m" + then show "P x \ Q (m - x) \ carrier R" + using assms carrier_coeffE + by blast + qed + then show "(P \\<^sub>p Q) m \ carrier R" + using assms finsum_closed[of "(\x. (P x) \ (Q (m - x)))" "mset_factors m"] + unfolding carrier_coeff_def P_ring_mult_def + by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Multivariable Polynomial Induction\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma mpoly_induct: + assumes "\Q. Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) = 0 \ P Q" + assumes "\n. (\Q. Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) \ n \ P Q) + \ (\Q. Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) \ (Suc n) \ P Q)" + assumes "Q \ indexed_pset I (carrier R)" + shows "P Q" +proof- + have "\m. \Q. Q \ indexed_pset I (carrier R) \ card(monomials_of R Q) \ m \ P Q" + proof- + fix m + show "\Q. Q \ indexed_pset I (carrier R) \ card(monomials_of R Q) \ m \ P Q" + apply(induction m) + using assms(1) + apply blast + using assms + by blast + qed + then show ?thesis + using assms(3) by blast +qed + +lemma monomials_of_card_zero: + assumes "Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) = 0" + shows "Q = indexed_const \" +proof- + have 0: "carrier_coeff Q" + using assms indexed_pset_in_carrier + by blast + have "\m. Q m = \" + unfolding monomials_of_def + using assms monomials_finite + by (metis card_0_eq complement_of_monomials_of empty_iff) + then show ?thesis + using 0 assms unfolding indexed_const_def + by auto +qed + +text\Polynomial induction on the number of monomials with nonzero coefficient:\ + +lemma mpoly_induct': + assumes "P (indexed_const \)" + assumes "\n. (\Q. Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) \ n \ P Q) + \ (\Q. Q \ indexed_pset I (carrier R) \ card (monomials_of R Q) = (Suc n) \ P Q)" + assumes "Q \ indexed_pset I (carrier R)" + shows "P Q" + apply(rule mpoly_induct) + using assms(1) monomials_of_card_zero apply blast +proof- + show "\n Q. (\Q. Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ n \ P Q) \ Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ Suc n \ P Q" + proof- + fix n + fix Q + assume A0: "(\Q. Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ n \ P Q)" + assume A1: "Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ Suc n" + show "P Q" + apply(cases "card (monomials_of R Q) = Suc n") + using assms A0 A1 + apply blast + using assms A0 A1 + using le_SucE by blast + qed + show "Q \ (carrier R [\\<^bsub>I\<^esub>])" + using assms by auto +qed + +lemma monomial_poly_split: + assumes "P \ indexed_pset I (carrier R)" + assumes "m \ monomials_of R P" + shows "(restrict_poly_to_monom_set R P ((monomials_of R P) - {m})) \ ((mset_to_IP R m) \\<^sub>p (indexed_const (P m))) = P" +proof fix x + show "(restrict_poly_to_monom_set R P (monomials_of R P - {m}) \ (mset_to_IP R m \\<^sub>p indexed_const (P m))) x = P x" + proof(cases "x = m") + case True + have T0: "(restrict_poly_to_monom_set R P (monomials_of R P - {m})) x = \" + unfolding restrict_poly_to_monom_set_def + by (simp add: True) + have T1: "(mset_to_IP R m \\<^sub>p indexed_const (P m)) x = P x" + using assms True indexed_const_P_multE[of "mset_to_IP R m" I "P m" m] + mset_to_IP_simp + by (metis Idl_subset_ideal' carrier_coeffE genideal_one indexed_pset_in_carrier + l_one mset_to_IP_closed' one_closed) + then show ?thesis + using T0 True + unfolding indexed_padd_def + using assms(1) carrier_coeffE indexed_pset_in_carrier l_zero + by fastforce + next + case False + then have F0: "(restrict_poly_to_monom_set R P (monomials_of R P - {m})) x = P x" + proof(cases "x \ monomials_of R P") + case True + then show ?thesis + by (simp add: False restrict_poly_to_monom_set_def) + next + case False + then show ?thesis + by (simp add: complement_of_monomials_of restrict_poly_to_monom_set_def) + qed + have F1: "mset_to_IP R m \ (carrier R [\\<^bsub>I\<^esub>])" + using assms(1) assms(2) mset_to_IP_closed' by blast + have F2: "P m \ carrier R" + using assms(1) carrier_coeffE indexed_pset_in_carrier by blast + have F3: "(mset_to_IP R m \\<^sub>p indexed_const (P m)) x = \" + using False F1 F2 assms indexed_const_P_multE[of "mset_to_IP R m" I "P m" x] + by (simp add: F1 \m \ monomials_of R P\) + then show ?thesis using F0 unfolding indexed_padd_def + using assms(1) carrier_coeffE indexed_pset_in_carrier + by fastforce + qed +qed + +lemma restrict_not_in_monoms: + assumes "a \ monomials_of R P" + shows "restrict_poly_to_monom_set R P A = restrict_poly_to_monom_set R P (insert a A)" +proof + fix x + show " restrict_poly_to_monom_set R P A x = restrict_poly_to_monom_set R P (insert a A) x " + unfolding restrict_poly_to_monom_set_def using assms unfolding monomials_of_def + by simp +qed + +lemma restriction_closed': + assumes "P \ indexed_pset I (carrier R)" + assumes "finite ms" + shows "(restrict_poly_to_monom_set R P ms) \ indexed_pset I (carrier R)" + apply(rule finite.induct[of ms]) + apply (simp add: assms(2); fail) +proof- + show "restrict_poly_to_monom_set R P {} \ (carrier R [\\<^bsub>I\<^esub>])" + proof- + have "restrict_poly_to_monom_set R P {} = indexed_const \" + unfolding restrict_poly_to_monom_set_def indexed_const_def by auto + then show ?thesis + by (simp add: indexed_pset.indexed_const) + qed + show "\A a. finite A \ restrict_poly_to_monom_set R P A \ (carrier R [\\<^bsub>I\<^esub>]) \ restrict_poly_to_monom_set R P (insert a A) \ (carrier R [\\<^bsub>I\<^esub>])" + proof- + fix A + fix a + assume A: "restrict_poly_to_monom_set R P A \ (carrier R [\\<^bsub>I\<^esub>])" + show "restrict_poly_to_monom_set R P (insert a A) \ (carrier R [\\<^bsub>I\<^esub>])" + proof(cases "a \ monomials_of R P") + case True + then have T0: "mset_to_IP R a \ (carrier R [\\<^bsub>I\<^esub>])" + using assms(1) mset_to_IP_closed' by blast + then have T1: "(mset_to_IP R a \\<^sub>p indexed_const (P a)) \ (carrier R [\\<^bsub>I\<^esub>])" + by (meson assms(1) carrier_coeffE indexed_pset_in_carrier + local.ring_axioms ring.indexed_const_P_mult_closed subset_refl) + + show ?thesis + proof(cases "a \ A") + case True + then show ?thesis + by (simp add: A insert_absorb) + next + case False + have "restrict_poly_to_monom_set R P (insert a A) = restrict_poly_to_monom_set R P A \ (mset_to_IP R a \\<^sub>p indexed_const (P a))" + proof + fix x + show "restrict_poly_to_monom_set R P (insert a A) x = (restrict_poly_to_monom_set R P A \ (mset_to_IP R a \\<^sub>p indexed_const (P a))) x" + proof(cases "x = a") + case True + then have FT0: "(if x \ insert a A then P x else \) = P x" + by simp + have FT1: "(if x \ A then P x else \) = \" + by (simp add: False True) + have FT2: "P a \ carrier R" + using assms(1) carrier_coeffE indexed_pset_in_carrier by blast + have FT3: "(mset_to_IP R a \\<^sub>p indexed_const (P a)) x = P x" + using T0 FT2 True indexed_const_P_multE[of "mset_to_IP R a" I "P a" x] mset_to_IP_simp[of a] + by simp + then show ?thesis + using False + unfolding restrict_poly_to_monom_set_def indexed_padd_def + using FT0 FT1 FT3 assms + by (simp add: FT2 True) + next + case F: False + then have FT0: "(if x \ insert a A then P x else \) = (if x \ A then P x else \)" + by simp + have FT2: "P a \ carrier R" + using assms(1) carrier_coeffE indexed_pset_in_carrier by blast + have FT3: "(mset_to_IP R a \\<^sub>p indexed_const (P a)) x = \" + using T0 FT2 True indexed_const_P_multE[of "mset_to_IP R a" I "P a" x] mset_to_IP_simp[of a] + by (simp add: F mset_to_IP_def) + show ?thesis + unfolding restrict_poly_to_monom_set_def indexed_padd_def + proof- + + show "(if x \ insert a A then P x else \) = (if x \ A then P x else \) \ (mset_to_IP R a \\<^sub>p indexed_const (P a)) x" + apply(cases "x \ A") + using FT0 FT2 FT3 F False assms carrier_coeffE indexed_pset_in_carrier local.ring_axioms + apply fastforce + using FT0 FT2 FT3 F False assms + by simp + qed + qed + qed + then show ?thesis + using A T1 indexed_pset.simps by blast + qed + next + case False + then show ?thesis + using A restrict_not_in_monoms by fastforce + qed + qed +qed + +lemma restriction_restrict: +"restrict_poly_to_monom_set R P ms = restrict_poly_to_monom_set R P (ms \ monomials_of R P)" +proof + fix x + show "restrict_poly_to_monom_set R P ms x = restrict_poly_to_monom_set R P (ms \ monomials_of R P) x" + unfolding restrict_poly_to_monom_set_def monomials_of_def + by simp +qed + +lemma restriction_closed: + assumes "P \ indexed_pset I (carrier R)" + assumes "Q = restrict_poly_to_monom_set R P ms" + shows "Q \ indexed_pset I (carrier R)" +proof- + have "Q = restrict_poly_to_monom_set R P (ms \ monomials_of R P)" + using assms restriction_restrict + by blast + then show ?thesis + using assms restriction_closed'[of P I "(ms \ monomials_of R P)"] + using monomials_finite + by blast +qed + +lemma monomial_split_card: + assumes "P \ indexed_pset I (carrier R)" + assumes "m \ monomials_of R P" + shows "card (monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m})))= + card (monomials_of R P) -1" +proof- + have 0: "(monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m}))) = + (monomials_of R P) - {m}" + using assms(1) + by (meson Diff_subset restrict_poly_to_monom_set_monoms) + then have 1: "card (monomials_of R (restrict_poly_to_monom_set R P ((monomials_of R P) - {m}))) = + card ((monomials_of R P) - {m})" + by auto + have " card ((monomials_of R P) - {m}) = card (monomials_of R P) - 1" + using assms(2) + using assms(1) monomials_finite by fastforce + then show ?thesis + by (simp add: "1") +qed + +lemma P_ring_mult_closed': + assumes "a \ indexed_pset I (carrier R)" + assumes "b \ indexed_pset I (carrier R)" + shows "a \\<^sub>p b \ indexed_pset I (carrier R)" + apply(rule mpoly_induct'[of "\b. a \\<^sub>p b \ indexed_pset I (carrier R)" I b]) + using assms(1) indexed_const_P_mult_closed apply blast +proof- + show "b \ (carrier R [\\<^bsub>I\<^esub>])" + using assms by auto + show "\n Q. (\Q. Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ n \ a \\<^sub>p Q \ (carrier R [\\<^bsub>I\<^esub>])) \ + Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) = Suc n \ a \\<^sub>p Q \ (carrier R [\\<^bsub>I\<^esub>])" + proof- + fix n + fix Q + assume A0: "(\Q. Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) \ n \ a \\<^sub>p Q \ (carrier R [\\<^bsub>I\<^esub>]))" + assume A1: "Q \ (carrier R [\\<^bsub>I\<^esub>]) \ card (monomials_of R Q) = Suc n" + obtain m where m_def: "m \ monomials_of R Q" + using A1 + by fastforce + obtain P where P_def: "P = (restrict_poly_to_monom_set R Q ((monomials_of R Q) - {m}))" + by simp + have "P \ (carrier R [\\<^bsub>I\<^esub>])" + using P_def A1 restriction_closed + by blast + have 0: "a \\<^sub>p P \ (carrier R [\\<^bsub>I\<^esub>])" + using A0 P_def + by (metis A1 One_nat_def \P \ (carrier R [\\<^bsub>I\<^esub>])\ diff_Suc_Suc m_def + minus_nat.diff_0 monomial_split_card nat_le_linear) + have 1: "a \\<^sub>p (mset_to_IP R m) \ (carrier R [\\<^bsub>I\<^esub>])" + using assms mset_to_IP_mult_closed[of a I m] A1 m_def mset_to_IP_indices + by blast + have 2: "a \\<^sub>p (mset_to_IP R m \\<^sub>p indexed_const (Q m)) \ (carrier R [\\<^bsub>I\<^esub>])" + using P_ring_mult_assoc[of a "mset_to_IP R m" " indexed_const (Q m)"] + 1 + by (metis A1 assms(1) carrier_coeffE indexed_pset.indexed_const + indexed_pset_in_carrier local.ring_axioms m_def mset_to_IP_closed' + ring.indexed_const_P_mult_closed set_eq_subset) + have 3: "(P \ (mset_to_IP R m \\<^sub>p indexed_const (Q m))) = Q" + using P_def monomial_poly_split[of Q I m] + using A1 m_def by blast + have 4: "(a \\<^sub>p P ) \ (a \\<^sub>p (mset_to_IP R m \\<^sub>p indexed_const (Q m))) = a \\<^sub>p Q" + using assms 2 3 P_ring_rdistr[of a P "(mset_to_IP R m \\<^sub>p indexed_const (Q m))"] + by (metis A1 Idl_subset_ideal' P_ring_mult_closed \P \ (carrier R [\\<^bsub>I\<^esub>])\ + carrier_coeffE indexed_pset.indexed_const indexed_pset_in_carrier m_def + mset_to_IP_closed' onepideal principalideal.generate) + then show " a \\<^sub>p Q \ (carrier R [\\<^bsub>I\<^esub>])" + by (metis "0" "2" indexed_pset.indexed_padd) + qed +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Subtraction of Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition P_ring_uminus :: "('a,'b) ring_scheme \ ('a,'c) mvar_poly \ ('a,'c) mvar_poly" where +"P_ring_uminus R P = (\m. \\<^bsub>R\<^esub> (P m))" + +context ring + +begin + + +lemma P_ring_uminus_eq: + assumes "a \ indexed_pset I (carrier R)" + shows "P_ring_uminus R a = a \\<^sub>p (indexed_const (\ \))" +proof + fix x + have "(a \\<^sub>p indexed_const (\ \)) x = a x \ \ \" + using indexed_const_P_multE[of a I "\ \" x] assms + by blast + then show "P_ring_uminus R a x = (a \\<^sub>p indexed_const (\ \)) x" + unfolding P_ring_uminus_def + using assms + by (metis Idl_subset_ideal' carrier_coeffE indexed_pset_in_carrier + one_closed onepideal principalideal.generate r_minus r_one) +qed + +lemma P_ring_uminus_closed: + assumes "a \ indexed_pset I (carrier R)" + shows "P_ring_uminus R a \ indexed_pset I (carrier R)" + using assms P_ring_uminus_eq + by (metis add.l_inv_ex indexed_const_P_mult_closed minus_equality one_closed) + +lemma P_ring_uminus_add: + assumes "a \ indexed_pset I (carrier R)" + shows "P_ring_uminus R a \ a = indexed_const \" +proof + fix x + show "(P_ring_uminus R a \ a) x = indexed_const \ x" + using assms + unfolding P_ring_uminus_def indexed_const_def indexed_padd_def + by (meson carrier_coeffE indexed_pset_in_carrier + local.ring_axioms ring.ring_simprules(9) set_eq_subset) +qed + +text\multiplication by 1\ + +lemma one_mult_left: + assumes "a \ indexed_pset I (carrier R)" + shows "(indexed_const \) \\<^sub>p a = a" +proof + fix m + show "(indexed_const \ \\<^sub>p a) m = a m " + unfolding indexed_const_def P_ring_mult_def + proof- + have 0: "(\x\((mset_factors m) - {{#}}). (if x = {#} then \ else \) \ a (m - x)) = \" + proof- + have "\x. x\((mset_factors m) - {{#}}) \ (if x = {#} then \ else \) \ a (m - x) = \ \ a (m - x)" + by simp + then have "\x. x\((mset_factors m) - {{#}}) \ (if x = {#} then \ else \) \ a (m - x) = \" + using assms + by (metis carrier_coeffE indexed_pset_in_carrier local.ring_axioms ring.ring_simprules(24) set_eq_subset) + then show ?thesis + by (meson add.finprod_one_eqI) + qed + have 1: "(\x. (if x = {#} then \ else \) \ a (m - x)) {#} = a m" + by (metis (full_types) assms carrier_coeffE diff_zero + indexed_pset_in_carrier local.ring_axioms ring.ring_simprules(12) set_eq_subset) + have 2: "(\x\insert {#} (mset_factors m - {{#}}). (if x = {#} then \ else \) \ a (m - x)) = + (\x. (if x = {#} then \ else \) \ a (m - x)) {#} \ + (\x\((mset_factors m) - {{#}}). (if x = {#} then \ else \) \ a (m - x))" + proof- + have 20:"finite (mset_factors m - {{#}})" + by simp + have 21: " {#} \ mset_factors m - {{#}}" + by blast + have 22: " (\x. (if x = {#} then \ else \) \ a (m - x)) \ mset_factors m - {{#}} \ carrier R" + proof + fix x + assume A: "x \ mset_factors m - {{#}}" + show "(if x = {#} then \ else \) \ a (m - x) \ carrier R" + apply(cases "x = {#}") + using A + apply blast + using assms carrier_coeffE indexed_pset_in_carrier local.ring_axioms + one_closed ring.ring_simprules(5) set_eq_subset zero_closed + by (metis (mono_tags, opaque_lifting)) + qed + have 23: "(if {#} = {#} then \ else \) \ a (m - {#}) \ carrier R" + by (metis "1" assms carrier_coeffE indexed_pset_in_carrier set_eq_subset) + show ?thesis + using 20 21 22 23 finsum_insert[of "((mset_factors m) - {{#}})" "{#}" + " (\x. (if x = {#} then \ else \) \ a (m - x))"] + by blast + qed + have 3: "insert {#} (mset_factors m - {{#}}) = mset_factors m" + proof- + have "{#} \ mset_factors m" + using monom_divides_factors subset_mset.zero_le + by blast + then show ?thesis + by blast + qed + show "(\x\mset_factors m. (if x = {#} then \ else \) \ a (m - x)) = a m" + using 0 1 2 3 assms + by (metis (no_types, lifting) Idl_subset_ideal' carrier_coeffE + genideal_one indexed_pset_in_carrier one_closed r_zero) + qed +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) + subsection\The Carrier of the Ring of Indexed Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + + +abbreviation(input) Pring_set where +"Pring_set R I \ ring.indexed_pset R I (carrier R) " + +context ring + +begin + +lemma Pring_set_zero: + assumes "f \ Pring_set R I" + assumes "\ set_mset m \ I" + shows " f m = \\<^bsub>R\<^esub>" +proof- + have "\ set_mset m \ I \ f m = \\<^bsub>R\<^esub>" + apply(induction m) + apply simp + by (meson assms complement_of_monomials_of mset_to_IP_indices') + then show ?thesis + using assms(2) by blast +qed + +lemma(in ring) Pring_cfs_closed: + assumes "P \ Pring_set R I" + shows "P m \ carrier R" + using assms carrier_coeffE indexed_pset_in_carrier + by blast + +lemma indexed_pset_mono_aux: + assumes "P \ indexed_pset I S" + shows "S \ T \ P \ indexed_pset I T" + using assms + apply(induction P) + using indexed_pset.indexed_const apply blast + using indexed_pset.indexed_padd apply blast + by (simp add: indexed_pset.indexed_pmult) + +lemma indexed_pset_mono: + assumes "S \ T" + shows "indexed_pset I S \ indexed_pset I T" + using assms indexed_pset_mono_aux + by blast + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Scalar Multiplication\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition poly_scalar_mult :: "('a, 'b) ring_scheme \ 'a \ ('a,'c) mvar_poly \ ('a,'c) mvar_poly" where +"poly_scalar_mult R c P = (\ m. c \\<^bsub>R\<^esub> P m)" + +lemma(in cring) poly_scalar_mult_eq: + assumes "c \ carrier R" + shows "P \ Pring_set R (I :: 'c set) \ poly_scalar_mult R c P = indexed_const c \\<^sub>p P" +proof(erule indexed_pset.induct) + show "\k. k \ carrier R \ poly_scalar_mult R c (indexed_const k) = indexed_const c \\<^sub>p indexed_const k" + proof- + fix k + assume A0: "k \ carrier R" + show "poly_scalar_mult R c (indexed_const k) = (indexed_const c \\<^sub>p indexed_const k) " + unfolding poly_scalar_mult_def + proof + show "\m. c \ indexed_const k m = (indexed_const c \\<^sub>p indexed_const k) m" + using indexed_const_P_mult_eq + by (metis A0 assms indexed_const_P_mult_eq indexed_const_def local.semiring_axioms semiring.r_null) + qed + qed + show "\P Q. P \ Pring_set R I \ + poly_scalar_mult R c P = indexed_const c \\<^sub>p P \ + Q \ Pring_set R I \ poly_scalar_mult R c Q = indexed_const c \\<^sub>p Q \ poly_scalar_mult R c (P \ Q) = indexed_const c \\<^sub>p (P \ Q)" + proof + fix P Q :: "'c monomial \ 'a" + fix x :: "'c monomial" + assume A0: "P \ Pring_set R I" + assume A1: "poly_scalar_mult R c P = indexed_const c \\<^sub>p P" + assume A2: "Q \ Pring_set R I" + assume A3: "poly_scalar_mult R c Q = indexed_const c \\<^sub>p Q" + show "poly_scalar_mult R c (P \ Q) x = (indexed_const c \\<^sub>p (P \ Q)) x" + proof- + have P0: "poly_scalar_mult R c (P \ Q) = (poly_scalar_mult R c P) \ (poly_scalar_mult R c Q)" + unfolding poly_scalar_mult_def + proof + fix m + show "c \ (P \ Q) m = ((\m. c \ P m) \ (\m. c \ Q m)) m" + proof- + have LHS: "c \ (P \ Q) m = c \ ( P m \ Q m) " + by (simp add: indexed_padd_def) + then have LHS1: "c \ (P \ Q) m = (c \ P m) \ (c \Q m) " + proof- + have 0: "carrier_coeff P" + using A0 indexed_pset_in_carrier + by blast + have 1: "P m \ carrier R" + using 0 carrier_coeffE by blast + have 2: "carrier_coeff Q" + using A2 indexed_pset_in_carrier + by blast + have 3: "Q m \ carrier R" using 2 + using carrier_coeff_def + by blast + show ?thesis using 1 3 assms + by (simp add: LHS r_distr) + qed + then show ?thesis + by (simp add: indexed_padd_def) + qed + qed + have P1: "poly_scalar_mult R c (P \ Q) = (indexed_const c \\<^sub>p P) \ (indexed_const c \\<^sub>p Q)" + using P0 A1 A3 + by simp + have P2: "indexed_const c \ Pring_set R I" + using assms indexed_pset.indexed_const by blast + show ?thesis + using P2 A0 A2 P1 + by (metis P_ring_rdistr indexed_pset_in_carrier set_eq_subset) + qed + qed + show "\P i. P \ Pring_set R I \ + poly_scalar_mult R c P = indexed_const c \\<^sub>p P \ i \ I \ poly_scalar_mult R c (P \ i) = indexed_const c \\<^sub>p (P \ i)" + proof- + fix P + fix i + assume A0: "P \ Pring_set R I" + assume A1: "poly_scalar_mult R c P = indexed_const c \\<^sub>p P" + assume A2: "i \ I" + show "poly_scalar_mult R c (P \ i) = indexed_const c \\<^sub>p (P \ i)" + proof + fix x + show "poly_scalar_mult R c (P \ i) x = (indexed_const c \\<^sub>p (P \ i)) x" + proof- + have RHS: "(indexed_const c \\<^sub>p (P \ i)) = (indexed_const c \\<^sub>p P) \ i" + proof- + have B0: "P \ i = P \\<^sub>p (mset_to_IP R {#i#})" + by (meson A0 A2 poly_index_mult) + have B1: " (indexed_const c \\<^sub>p P) \ i = (indexed_const c \\<^sub>p P) \\<^sub>p (mset_to_IP R {#i#})" + by (meson A0 A2 P_ring_mult_closed' assms indexed_pset.simps poly_index_mult) + have B2: "(indexed_const c \\<^sub>p (P \ i)) = indexed_const c \\<^sub>p P \\<^sub>p (mset_to_IP R {#i#})" + by (metis A0 A2 B1 assms cring.P_ring_mult_comm indexed_const_var_mult + indexed_pmult_in_carrier indexed_pset.indexed_const indexed_pset_in_carrier + is_cring set_eq_subset) + show ?thesis using A0 A1 A2 B2 B1 assms + by (simp add: A0) + then have RHS': "(indexed_const c \\<^sub>p (P \ i)) = (poly_scalar_mult R c P) \ i" + using A0 A1 A2 assms + by simp + qed + show ?thesis apply(cases "i \# x") + unfolding poly_scalar_mult_def + apply (metis A1 RHS poly_scalar_mult_def indexed_pmult_def ) + by (metis RHS assms indexed_pmult_def r_null) + qed + qed + qed +qed + +lemma(in cring) poly_scalar_mult_const: + assumes "c \ carrier R" + assumes "k \ carrier R" + shows "poly_scalar_mult R k (indexed_const c) = indexed_const (k \ c)" + using assms poly_scalar_mult_eq + by (simp add: poly_scalar_mult_eq indexed_const_P_mult_eq indexed_pset.indexed_const) + +lemma(in cring) poly_scalar_mult_closed: + assumes "c \ carrier R" + assumes "P \ Pring_set R I" + shows "poly_scalar_mult R c P \ Pring_set R I" + using assms poly_scalar_mult_eq + by (metis P_ring_mult_closed' indexed_pset.indexed_const) + +lemma(in cring) poly_scalar_mult_zero: + assumes "P \ Pring_set R I" + shows "poly_scalar_mult R \ P = indexed_const \" +proof + fix x + show "poly_scalar_mult R \ P x = indexed_const \ x" + unfolding poly_scalar_mult_def + using assms + by (metis Pring_cfs_closed indexed_zero_def l_null) +qed + +lemma(in cring) poly_scalar_mult_one: + assumes "P \ Pring_set R I" + shows "poly_scalar_mult R \ P = P" +proof + fix x show "poly_scalar_mult R \ P x = P x" + using assms + by (metis one_closed one_mult_left poly_scalar_mult_eq) +qed + +lemma(in cring) times_poly_scalar_mult: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + assumes "k \ carrier R" + shows "P \\<^sub>p (poly_scalar_mult R k Q) = poly_scalar_mult R k (P \\<^sub>p Q)" +proof- + have "P \\<^sub>p (poly_scalar_mult R k Q) = P \\<^sub>p (indexed_const k) \\<^sub>p Q" + by (metis assms(1) assms(2) assms(3) indexed_pset_mono_aux + local.ring_axioms poly_scalar_mult_eq ring.P_ring_mult_assoc ring.indexed_pset.intros(1) + ring.indexed_pset_in_carrier subset_refl) + then have "P \\<^sub>p (poly_scalar_mult R k Q) = (indexed_const k) \\<^sub>p P \\<^sub>p Q" + by (metis P_ring_mult_comm assms(1) assms(3) local.ring_axioms ring.indexed_pset.indexed_const ring.indexed_pset_in_carrier subset_refl) + then show ?thesis + by (metis (no_types, opaque_lifting) P_ring_mult_closed' Pring_cfs_closed assms(1) assms(2) assms(3) + carrier_coeff_def indexed_pset.indexed_const local.ring_axioms poly_scalar_mult_eq ring.P_ring_mult_assoc) +qed + +lemma(in cring) poly_scalar_mult_times: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + assumes "k \ carrier R" + shows " poly_scalar_mult R k (Q \\<^sub>p P) = (poly_scalar_mult R k Q) \\<^sub>p P" + using assms times_poly_scalar_mult + by (metis (no_types, opaque_lifting) P_ring_mult_comm cring.P_ring_mult_comm + cring.poly_scalar_mult_closed is_cring local.ring_axioms ring.indexed_pset_in_carrier subset_refl) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Defining the Ring of Indexed Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) +definition Pring :: "('b, 'e) ring_scheme \ 'a set \ ('b, ('b,'a) mvar_poly) module" where + +"Pring R I \ \ carrier = Pring_set R I, + Group.monoid.mult = P_ring_mult R , + one = ring.indexed_const R \\<^bsub>R\<^esub>, + zero = ring.indexed_const R \\<^bsub>R\<^esub>, + add = ring.indexed_padd R, + smult = poly_scalar_mult R\" + +context ring + +begin + +lemma Pring_car: +"carrier (Pring R I) = Pring_set R I" + unfolding Pring_def + by auto + +text\Definitions of the operations and constants:\ + +lemma Pring_mult: +"a \\<^bsub>Pring R I\<^esub> b = a \\<^sub>p b" + unfolding Pring_def + by simp + +lemma Pring_add: +"a \\<^bsub>Pring R I\<^esub> b = a \ b" + unfolding Pring_def + by simp + +lemma Pring_zero: +"\\<^bsub>Pring R I\<^esub> = indexed_const \" + unfolding Pring_def by simp + +lemma Pring_one: +"\\<^bsub>Pring R I\<^esub> = indexed_const \" + unfolding Pring_def by simp + +lemma Pring_smult: +"(\\<^bsub>Pring R I\<^esub>) = (poly_scalar_mult R)" + unfolding Pring_def by simp + +lemma Pring_carrier_coeff: + assumes "a \ carrier (Pring R I)" + shows "carrier_coeff a" + using assms indexed_pset_in_carrier[of "(carrier R)" a I] Pring_car + by blast + +lemma Pring_carrier_coeff'[simp]: + assumes "a \ carrier (Pring R I)" + shows "a m \ carrier R" + using Pring_car[of I] assms carrier_coeffE indexed_pset_in_carrier + by blast + +lemma Pring_add_closed: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> b \ carrier (Pring R I)" + using assms Pring_def[of R I] + by (simp add: Pring_def[of R I] indexed_pset.indexed_padd) + +lemma Pring_mult_closed: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> b \ carrier (Pring R I)" + using assms P_ring_mult_closed'[of a I b] Pring_car[of I] Pring_mult[of I a b] + by (simp add: \a \\<^bsub>Pring R I\<^esub> b = a \\<^sub>p b\ \carrier (Pring R I) = Pring_set R I\) + +lemma Pring_one_closed: +"\\<^bsub>Pring R I\<^esub> \ carrier (Pring R I)" +proof- + have "indexed_const \ \ carrier (Pring R I)" + using Pring_car indexed_pset.simps by blast + then show ?thesis + unfolding Pring_def by auto +qed + +lemma Pring_zero_closed: +"\\<^bsub>Pring R I\<^esub> \ carrier (Pring R I)" +proof- + have "indexed_const \ \ carrier (Pring R I)" + using Pring_car indexed_pset.simps by blast + then show ?thesis + unfolding Pring_def by auto +qed + +lemma Pring_var_closed: + assumes "i \ I" + shows "var_to_IP R i \ carrier (Pring R I)" +unfolding var_to_IP_def + by (metis Pring_car Pring_one_closed assms indexed_pset.indexed_pmult + local.ring_axioms monom_add_mset one_mset_to_IP ring.Pring_one) + +text\Properties of addition:\ + +lemma Pring_add_assoc: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + assumes "c \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> c) = (a \\<^bsub>Pring R I\<^esub> b) \\<^bsub>Pring R I\<^esub> c" +proof- + have "a \ (b \ c) = (a \ b) \ c" + proof + fix x + have "carrier_coeff a" "carrier_coeff b" "carrier_coeff c" + using assms Pring_car[of I] Pring_carrier_coeff apply blast + using assms Pring_car[of I] Pring_carrier_coeff apply blast + using assms Pring_car[of I] Pring_carrier_coeff by blast + then have "a x \ carrier R" "b x \ carrier R" "c x \ carrier R" + using carrier_coeffE apply blast + using \carrier_coeff b\ carrier_coeffE apply blast + using \carrier_coeff c\ carrier_coeffE by blast + show "(a \ (b \ c)) x = (a \ b \ c) x" + unfolding indexed_padd_def + using assms + by (simp add: \a x \ carrier R\ \b x \ carrier R\ \c x \ carrier R\ add.m_assoc) + qed + then show ?thesis + using assms + by (simp add: Pring_add) +qed + +lemma Pring_add_comm: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> b = b \\<^bsub>Pring R I\<^esub> a" +proof- + have "a \ b = b \ a" + proof fix x + show "(a \ b) x = (b \ a) x" + using assms + by (metis abelian_monoid.a_comm abelian_monoid_axioms + indexed_padd_def local.ring_axioms ring.Pring_carrier_coeff') + qed + then show ?thesis + by (simp add: Pring_add) +qed + +lemma Pring_add_zero: + assumes "a \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub> = a" + "\\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub> a = a" + using assms Pring_zero Pring_add + apply (metis Pring_carrier_coeff indexed_padd_zero(1)) + using assms Pring_zero Pring_add + by (metis Pring_carrier_coeff indexed_padd_zero(2)) + +text\Properties of multiplication\ + +lemma Pring_mult_assoc: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + assumes "c \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> c) = (a \\<^bsub>Pring R I\<^esub> b) \\<^bsub>Pring R I\<^esub> c" + using assms P_ring_mult_assoc[of a b c] + by (metis Pring_carrier_coeff Pring_mult) + +lemma Pring_mult_comm: + assumes "cring R" + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> b = b \\<^bsub>Pring R I\<^esub> a" + using assms Pring_carrier_coeff[of a I] Pring_carrier_coeff[of b I] + Pring_mult[of I b a] Pring_mult[of I a b] cring.P_ring_mult_comm[of R a b] + by metis + +lemma Pring_mult_one: + assumes "a \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub> = a" +proof + fix x + show "(a \\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub>) x = a x " + proof- + have 0: "(a \\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub>) x = (a \\<^sub>p indexed_const \) x" + using assms Pring_mult[of I a "\\<^bsub>Pring R I\<^esub>" ] Pring_one[of I] + by metis + have 1: "\ \ carrier R" + by simp + have 2: "(a \\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub>) x = a x \ \" + using 0 1 indexed_const_P_multE[of a I \ x] + assms Pring_car[of I] + by metis + then show ?thesis + using assms by auto + qed +qed + +lemma Pring_mult_one': + assumes "a \ carrier (Pring R I)" + shows "\\<^bsub>Pring R I\<^esub> \\<^bsub>Pring R I\<^esub> a = a" + using one_mult_left[of a I] + assms Pring_one Pring_mult + by (simp add: Pring_mult Pring_one Pring_car) + +text\Distributive laws\ + +lemma Pring_mult_rdistr: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + assumes "c \ carrier (Pring R I)" + shows "a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> c) = (a \\<^bsub>Pring R I\<^esub> b) \\<^bsub>Pring R I\<^esub> (a \\<^bsub>Pring R I\<^esub> c)" +proof- + have "a \\<^sub>p (b \ c) = a \\<^sub>p b \ (a \\<^sub>p c)" + using P_ring_rdistr[of a b c] + assms Pring_carrier_coeff + by metis + then have "a \\<^sub>p (b \\<^bsub>Pring R I\<^esub> c) = a \\<^sub>p b \\<^bsub>Pring R I\<^esub> (a \\<^sub>p c)" + using Pring_add[of I b c] Pring_add[of I "a \\<^sub>p b" "a \\<^sub>p c"] + by auto + then have "a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> c) = (a \\<^sub>p b) \\<^bsub>Pring R I\<^esub> (a \\<^sub>p c)" + using Pring_mult[of I a "(b \\<^bsub>Pring R I\<^esub> c)"] + by auto + then have "a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> c) = (a \\<^bsub>Pring R I\<^esub> b) \\<^bsub>Pring R I\<^esub> (a \\<^sub>p c)" + using Pring_mult[of I a b] by metis + then show ?thesis + using Pring_mult[of I a c] by metis +qed + +lemma Pring_mult_ldistr: + assumes "a \ carrier (Pring R I)" + assumes "b \ carrier (Pring R I)" + assumes "c \ carrier (Pring R I)" + shows "(b \\<^bsub>Pring R I\<^esub> c) \\<^bsub>Pring R I\<^esub> a = (b \\<^bsub>Pring R I\<^esub> a) \\<^bsub>Pring R I\<^esub> (c \\<^bsub>Pring R I\<^esub> a)" +proof- + have "(b \ c) \\<^sub>p a = b \\<^sub>p a \ (c \\<^sub>p a)" + using P_ring_ldistr[of a b c] + assms Pring_carrier_coeff + by metis + then have " (b \\<^bsub>Pring R I\<^esub> c) \\<^sub>p a = b \\<^sub>p a \\<^bsub>Pring R I\<^esub> (c \\<^sub>p a)" + using Pring_add[of I b c] Pring_add[of I "b \\<^sub>p a" "c \\<^sub>p a"] + by auto + then have " (b \\<^bsub>Pring R I\<^esub> c) \\<^bsub>Pring R I\<^esub> a = (b \\<^sub>p a) \\<^bsub>Pring R I\<^esub> (c \\<^sub>p a)" + using Pring_mult[of I "(b \\<^bsub>Pring R I\<^esub> c)" a] + by auto + then have "(b \\<^bsub>Pring R I\<^esub> c) \\<^bsub>Pring R I\<^esub> a = (b \\<^bsub>Pring R I\<^esub> a) \\<^bsub>Pring R I\<^esub> (c \\<^sub>p a)" + using Pring_mult[of I b a] by metis + then show ?thesis + using Pring_mult[of I c a] by metis +qed + +text\Properties of subtraction:\ + +lemma Pring_uminus: + assumes "a \ carrier (Pring R I)" + shows "P_ring_uminus R a \ carrier (Pring R I)" + using P_ring_uminus_closed[of a I] Pring_car[of I] assms + by metis + +lemma Pring_subtract: + assumes "a \ carrier (Pring R I)" + shows "P_ring_uminus R a \\<^bsub>Pring R I\<^esub> a = \\<^bsub>Pring R I\<^esub>" + "a \\<^bsub>Pring R I\<^esub> P_ring_uminus R a = \\<^bsub>Pring R I\<^esub>" + using assms Pring_add[of I "P_ring_uminus R a " a] Pring_zero[of I] + apply (simp add: Pring_car local.ring_axioms ring.P_ring_uminus_add) + using assms Pring_add[of I "P_ring_uminus R a " a] Pring_zero[of I] + by (metis P_ring_uminus_add P_ring_uminus_closed Pring_add_comm Pring_car) + +text\Pring R I is a ring\ + +lemma Pring_is_abelian_group: + shows "abelian_group (Pring R I)" + apply(rule abelian_groupI) + apply (simp add: Pring_add_closed) + apply (simp add: local.ring_axioms ring.Pring_zero_closed) + apply (simp add: Pring_add_assoc) + apply (simp add: Pring_add_comm) + apply (simp add: local.ring_axioms ring.Pring_add_zero(2)) + using Pring_subtract(1) Pring_uminus + by blast + +lemma Pring_is_monoid: +"Group.monoid (Pring R I)" + apply(rule monoidI) + using Pring_mult_closed apply blast + apply (simp add: Pring_one_closed) + apply (metis Pring_mult_assoc) + using Pring_mult_one' + apply blast + using Pring_mult_one by blast + +lemma Pring_is_ring: + shows "ring (Pring R I)" + apply(rule ringI) + apply (simp add: Pring_is_abelian_group) + apply (simp add: Pring_is_monoid) + apply (simp add: Pring_mult_ldistr) + by (simp add: Pring_mult_rdistr) + +lemma Pring_is_cring: + assumes "cring R" + shows "cring (Pring R I)" + apply(rule cringI) + apply (simp add: Pring_is_abelian_group) + apply (simp add: Pring_is_monoid assms local.ring_axioms + monoid.monoid_comm_monoidI ring.Pring_mult_comm) + by (simp add: Pring_mult_ldistr) + +lemma Pring_a_inv: + assumes "P \ carrier (Pring R I)" + shows "\\<^bsub>Pring R I\<^esub> P = P_ring_uminus R P" +proof- + have 0: "P_ring_uminus R P \ carrier (Pring R I)" + using Pring_uminus assms + by blast + have 1: "P_ring_uminus R P \\<^bsub>Pring R I\<^esub> P = \\<^bsub>Pring R I\<^esub>" + using Pring_subtract(1) assms + by blast + show ?thesis using 0 1 assms Pring_is_ring + by (simp add: Pring_car Pring_is_abelian_group abelian_group.minus_equality) +qed + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Defining the R-Algebra of Indexed Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +context cring +begin + +lemma Pring_smult_cfs: + assumes "a \ carrier R" + assumes "P \ carrier (Pring R I)" + shows "(a \\<^bsub>Pring R I\<^esub> P) m = a \ (P m)" + using assms Pring_smult + by (simp add: Pring_smult poly_scalar_mult_def) + +lemma Pring_smult_closed: + "\ a x. [| a \ carrier R; x \ carrier (Pring R I) |] ==> a \\<^bsub>(Pring R I)\<^esub> x \ carrier (Pring R I)" + by (simp add: Pring_car Pring_smult poly_scalar_mult_closed) + +lemma Pring_smult_l_distr: + "!!a b x. [| a \ carrier R; b \ carrier R; x \ carrier (Pring R I) |] ==> + (a \ b) \\<^bsub>(Pring R I)\<^esub> x = (a \\<^bsub>(Pring R I)\<^esub> x) \\<^bsub>(Pring R I)\<^esub> (b \\<^bsub>(Pring R I)\<^esub> x)" +proof- fix a b x assume A: "a \ carrier R" "b \ carrier R" "x \ carrier (Pring R I)" + show "(a \ b) \\<^bsub>Pring R I\<^esub> x = a \\<^bsub>Pring R I\<^esub> x \\<^bsub>Pring R I\<^esub> b \\<^bsub>Pring R I\<^esub> x" + proof fix m + have "(a \ b) \ x m = (a \ (x m)) \ (b \ (x m))" + by (meson A(1) A(2) A(3) Pring_carrier_coeff' local.semiring_axioms semiring.l_distr) + thus "((a \ b) \\<^bsub>Pring R I\<^esub> x) m = (a \\<^bsub>Pring R I\<^esub> x \\<^bsub>Pring R I\<^esub> b \\<^bsub>Pring R I\<^esub> x) m" + using Pring_smult_cfs[of "a \ b" x I m] + Pring_smult_cfs[of "a" x I m] + Pring_smult_cfs[of "b" x I m] + Pring_smult_closed[of a x I] + Pring_smult_closed[of b x I] + Pring_add A + by (simp add: \(a \ b) \ x m = a \ x m \ b \ x m\ Pring_def indexed_padd_def poly_scalar_mult_def) + qed +qed + +lemma Pring_smult_r_distr: + "!!a x y. [| a \ carrier R; x \ carrier (Pring R I); y \ carrier (Pring R I) |] ==> + a \\<^bsub>(Pring R I)\<^esub> (x \\<^bsub>(Pring R I)\<^esub> y) = (a \\<^bsub>(Pring R I)\<^esub> x) \\<^bsub>(Pring R I)\<^esub> (a \\<^bsub>(Pring R I)\<^esub> y)" +proof fix a x y m assume A: "a \ carrier R" "x \ carrier (Pring R I)" "y \ carrier (Pring R I)" + show "(a \\<^bsub>Pring R I\<^esub> (x \\<^bsub>Pring R I\<^esub> y)) m = + (a \\<^bsub>Pring R I\<^esub> x \\<^bsub>Pring R I\<^esub> a \\<^bsub>Pring R I\<^esub> y) m" + using Pring_smult_cfs[of a "x \\<^bsub>Pring R I\<^esub> y" I m] + Pring_smult_cfs[of a x I m] + Pring_smult_cfs[of a y I m] + Pring_smult_closed[of a x I] + Pring_smult_closed[of a y I] + Pring_add A + by (metis (no_types, lifting) Pring_add_closed Pring_carrier_coeff' indexed_padd_def l_distr m_comm) +qed + +lemma Pring_smult_assoc1: + "!!a b x. [| a \ carrier R; b \ carrier R; x \ carrier (Pring R I) |] ==> + (a \ b) \\<^bsub>Pring R I\<^esub> x = a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> x)" +proof fix a b x m assume A: "a \ carrier R" "b \ carrier R" "x \ carrier (Pring R I)" + show " (a \ b \\<^bsub>Pring R I\<^esub> x) m = (a \\<^bsub>Pring R I\<^esub> (b \\<^bsub>Pring R I\<^esub> x)) m" + using Pring_smult_cfs[of "a \ b" x I m] + Pring_smult_cfs[of a "b \\<^bsub>Pring R I\<^esub> x" I m] + Pring_smult_cfs[of "b" x I m] + Pring_smult_closed[of a "b \\<^bsub>Pring R I\<^esub> x" I] + Pring_smult_closed[of b x I] + A(1) A(2) A(3) m_assoc m_closed by auto +qed + +lemma Pring_smult_one: + "!!x. x \ carrier (Pring R I) ==> (one R) \\<^bsub>Pring R I\<^esub> x = x" + by (simp add: Pring_car Pring_smult poly_scalar_mult_one) + + +lemma Pring_smult_assoc2: + "!!a x y. [| a \ carrier R; x \ carrier (Pring R I); y \ carrier (Pring R I) |] ==> + (a \\<^bsub>Pring R I\<^esub> x) \\<^bsub>Pring R I\<^esub> y = a \\<^bsub>Pring R I\<^esub> (x \\<^bsub>Pring R I\<^esub> y)" + by (simp add: Pring_def poly_scalar_mult_times) + +lemma Pring_algebra: +"algebra R (Pring R I)" + apply(rule algebraI) + apply (simp add: is_cring) + apply (simp add: Pring_is_cring is_cring) + apply (simp add: Pring_smult_closed) + apply (simp add: Pring_smult_l_distr) + apply (simp add: Pring_smult_r_distr) + apply (simp add: Pring_smult_assoc1) + apply (simp add: Pring_smult_one) + by (simp add: Pring_smult_assoc2) + + +end + + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Evaluation of Polynomials and Subring Structure\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + In this section the aim is to define the evaluation of a polynomial over its base ring. We define + both total evaluation of a polynomial at all variables, and partial evaluation at only a subset + of variables. The basic input for evaluation is a variable assignment function mapping variables + to ring elements. Once we can evaluate a polynomial $P$ in variables $I$ over a ring $R$ at an + assignment $f: I \to R$, this can be generalized to evaluation of $P$ in some other ring $S$, + given a variable assignment $f: I \to S$ and a ring homomorphism $\phi: R \to S$. We chose to + define this by simply applying $\phi$ to the coefficients of $P$, and then using the first + evaluation function over $S$. This could also have been done the other way around: define + general polynomial evaluation over any ring, given a ring hom $\phi$, and then defining + evaluation over the base ring $R$ as the specialization of this function to the case there + $\phi = \mathit{id}_R$.\ + +definition remove_monom :: +"('a,'b) ring_scheme \ 'c monomial \ ('a, 'c) mvar_poly \ ('a, 'c) mvar_poly" where +"remove_monom R m P = ring.indexed_padd R P (poly_scalar_mult R (\\<^bsub>R\<^esub> P m) (mset_to_IP R m))" + +context cring +begin + +lemma remove_monom_alt_def: + assumes "P \ Pring_set R I" + shows "remove_monom R m P n = (if n = m then \ else P n)" + unfolding remove_monom_def + apply(cases "n = m") + using assms + apply (metis Pring_cfs_closed cring.cring_simprules(3) poly_scalar_mult_def + indexed_padd_def is_cring mset_to_IP_simp r_neg r_one) + using assms + by (metis Pring_cfs_closed add.l_cancel_one cring.cring_simprules(3) + poly_scalar_mult_def indexed_padd_def is_cring mset_to_IP_simp' r_null zero_closed) + +lemma remove_monom_zero: + assumes "m \ monomials_of R P" + assumes "P \ Pring_set R I" + shows "remove_monom R m P = P" +proof + fix x + show "remove_monom R m P x = P x " + apply(cases "x \ monomials_of R P") + using assms unfolding monomials_of_def remove_monom_def + apply (metis cring.remove_monom_alt_def is_cring remove_monom_def) + using assms unfolding monomials_of_def remove_monom_def + by (metis assms(1) cring.remove_monom_alt_def is_cring local.ring_axioms + remove_monom_def ring.complement_of_monomials_of) +qed + +lemma remove_monom_closed: + assumes "P \ Pring_set R I" + shows "remove_monom R m P \ Pring_set R I" + + apply(cases "m \ monomials_of R P") + using assms poly_scalar_mult_closed[of "(\ P m)" "(mset_to_IP R m)" I] mset_to_IP_closed[of m I] + unfolding remove_monom_def + apply (meson Pring_cfs_closed add.inv_closed indexed_pset.indexed_padd mset_to_IP_closed') + by (metis assms remove_monom_def remove_monom_zero) + +lemma remove_monom_monomials: + assumes "P \ Pring_set R I" + shows "monomials_of R (remove_monom R m P) = monomials_of R P - {m}" +proof + show "monomials_of R (remove_monom R m P) \ monomials_of R P - {m}" + using assms remove_monom_alt_def[of P I m] + unfolding monomials_of_def + by auto + show "monomials_of R P - {m} \ monomials_of R (remove_monom R m P)" + using assms remove_monom_alt_def[of P I m] + unfolding monomials_of_def + by auto +qed + +text\The additive decomposition of a polynomial by a monomial\ + +lemma remove_monom_eq: + assumes "P \ Pring_set R I" + shows "P = (remove_monom R a P) \ poly_scalar_mult R (P a) (mset_to_IP R a)" + unfolding remove_monom_def poly_scalar_mult_def +proof + fix x + show "P x = (P \ (\m. \ P a \ mset_to_IP R a m) \ (\m. P a \ mset_to_IP R a m)) x" + apply(cases "x = a") + apply (metis Pring_cfs_closed assms l_minus l_zero local.ring_axioms mset_to_IP_simp one_closed r_neg r_one ring.indexed_padd_def) + proof- + assume A: "x \ a" + show "P x = (P \ (\m. \ P a \ mset_to_IP R a m) \ (\m. P a \ mset_to_IP R a m)) x" + using assms A + unfolding mset_to_IP_def indexed_padd_def + using Pring_cfs_closed by fastforce + qed +qed + +lemma remove_monom_restrict_poly_to_monom_set: + assumes "P \ Pring_set R I" + assumes "monomials_of R P = insert a M" + assumes "a \ M" + shows "(remove_monom R a P) = restrict_poly_to_monom_set R P M" +proof + fix m + show "remove_monom R a P m= restrict_poly_to_monom_set R P M m" + apply(cases "m = a") + using assms apply (metis remove_monom_alt_def restrict_poly_to_monom_set_def) + using assms + by (metis complement_of_monomials_of insert_iff remove_monom_alt_def restrict_poly_to_monom_set_def) +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Nesting of Polynomial Rings According to Nesting of Index Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma(in ring) Pring_carrier_subset: + assumes "J \ I" + shows "(Pring_set R J) \ (Pring_set R I)" +proof + fix P + show "P \ Pring_set R J \P \ Pring_set R I" + apply(erule indexed_pset.induct) + using indexed_pset.indexed_const apply blast + using indexed_pset.indexed_padd apply blast + by (meson assms indexed_pset.indexed_pmult subsetD) +qed + +lemma(in cring) Pring_set_restrict_induct: + shows "finite S \ \P. monomials_of R P = S \ P \ Pring_set R I \ (\ m \ monomials_of R P. set_mset m \ J) \ P \ Pring_set R J" +proof(erule finite.induct) + show "\P. monomials_of R P = {} \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J" + proof + fix P + show "monomials_of R P = {} \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J" + proof + assume A0: "monomials_of R P = {} \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J)" + show "P \ Pring_set R J " + by (metis A0 card_eq_0_iff indexed_pset.indexed_const monomials_of_card_zero zero_closed) + qed + qed + show "\A a. finite A \ + \P. monomials_of R P = A \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J \ + \P. monomials_of R P = insert a A \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J" + proof + fix A :: "('c monomial) set" fix a fix P + assume A1: "finite A" + assume A2: "\P. monomials_of R P = A \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J" + show "monomials_of R P = insert a A \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J) \ P \ Pring_set R J " + proof + assume A3: "monomials_of R P = insert a A \ P \ Pring_set R I \ (\m\monomials_of R P. set_mset m \ J)" + show "P \ Pring_set R J" + apply(cases "a \ A") + apply (metis A2 A3 insert_absorb) + proof- + assume N: "a \ A" + obtain Q where Q_def: "Q = remove_monom R a P" + by simp + have Q0: "monomials_of R Q = A" + proof- + have 0: "monomials_of R P = insert a A" + by (simp add: A3) + have 1: "P \ Pring_set R I" + using A3 by blast + have 2: "monomials_of R (remove_monom R a P) = monomials_of R P - {a}" + using A3 remove_monom_monomials by blast + then show ?thesis + using Q_def 0 + by (simp add: N) + qed + have Q1: "Q \ Pring_set R I" + using A3 Q_def remove_monom_closed by blast + have Q2: "(\m\monomials_of R Q. set_mset m \ J)" + using Q0 A3 + by blast + have "Q \ Pring_set R J" + using A2 Q0 Q1 Q2 by blast + then show "P \ Pring_set R J" + proof- + have "P = Q \ poly_scalar_mult R (P a) (mset_to_IP R a)" + using Q_def remove_monom_eq + using A3 by blast + then show ?thesis + by (metis A3 Pring_cfs_closed \Q \ Pring_set R J\ indexed_pset.indexed_padd insertCI mset_to_IP_closed poly_scalar_mult_closed) + qed + qed + qed + qed +qed + +lemma(in cring) Pring_set_restrict: + assumes "P \ Pring_set R I" + assumes "(\m. m \ monomials_of R P \ set_mset m \ J)" + shows " P \ Pring_set R J" + using assms Pring_set_restrict_induct[of "monomials_of R P"] + by (metis monomials_finite) + +lemma(in ring) Pring_mult_eq: + fixes I:: "'c set" + fixes J:: "'c set" + shows "(\\<^bsub>Pring R I\<^esub>) = (\\<^bsub>Pring R J\<^esub>)" + by (simp add: Pring_def) + +lemma(in ring) Pring_add_eq: + fixes I:: "'c set" + fixes J:: "'c set" + shows "(\\<^bsub>Pring R I\<^esub>) = (\\<^bsub>Pring R J\<^esub>)" + using Pring_def + by (simp add: Pring_def) + +lemma(in ring) Pring_one_eq: + fixes I:: "'c set" + fixes J:: "'c set" + shows "(\\<^bsub>Pring R I\<^esub>) = (\\<^bsub>Pring R J\<^esub>)" + using Pring_def + by (simp add: Pring_def) + +lemma(in ring) Pring_zero_eq: + fixes I:: "'c set" + fixes J:: "'c set" + shows "(\\<^bsub>Pring R I\<^esub>) = (\\<^bsub>Pring R J\<^esub>)" + using Pring_def + by (simp add: Pring_def) + +lemma(in ring) index_subset_Pring_subring: + assumes "J \ I" + shows "subring (carrier (Pring R J)) (Pring R I)" + apply(rule ring.subringI) + apply (simp add: Pring_is_ring; fail) + using assms + apply (simp add: Pring_car Pring_carrier_subset; fail) + using Pring_def + apply (simp add: Pring_def indexed_pset.indexed_const; fail) + proof- + show "\h. h \ carrier (Pring R J) \ \\<^bsub>Pring R I\<^esub> h \ carrier (Pring R J)" + proof- + fix h + assume A: "h \ carrier (Pring R J)" + then have A0: "\\<^bsub>Pring R J\<^esub> h = P_ring_uminus R h" + using Pring_a_inv[of h J] + by auto + have A1: "\\<^bsub>Pring R I\<^esub> h = P_ring_uminus R h" + using assms A Pring_carrier_subset[of J I] Pring_a_inv[of h I] Pring_car + by blast + show "\\<^bsub>Pring R I\<^esub> h \ carrier (Pring R J)" + using A0 A1 A + by (simp add: Pring_uminus) + qed + show " \h1 h2. h1 \ carrier (Pring R J) \ h2 \ carrier (Pring R J) \ h1 \\<^bsub>Pring R I\<^esub> h2 \ carrier (Pring R J)" + using assms Pring_mult_eq + by (metis Pring_mult_closed) + show " \h1 h2. h1 \ carrier (Pring R J) \ h2 \ carrier (Pring R J) \ h1 \\<^bsub>Pring R I\<^esub> h2 \ carrier (Pring R J)" + using assms Pring_add_eq + by (metis Pring_add_closed) + qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Inclusion Maps\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition Pring_inc :: "('a, 'c) mvar_poly \ ('a, 'c) mvar_poly" where +"Pring_inc \ (\P. P)" + +lemma(in ring) Princ_inc_is_ring_hom: + assumes "J \ I" + shows "ring_hom_ring (Pring R J) (Pring R I) Pring_inc" + unfolding Pring_inc_def + apply(rule ring_hom_ringI) + apply (simp add: Pring_is_ring) + using Pring_is_ring apply blast + using index_subset_Pring_subring[of I J] assms index_subset_Pring_subring subringE(1) + apply blast + using Pring_mult_eq[of I J] + apply (simp add: Pring_mult) + using Pring_add_eq[of I J] + apply (simp add: Pring_add) + using Pring_one_eq + by (simp add: Pring_one_eq) + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Restricting a Multiset to a Subset of Variables\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition restrict_to_indices :: "'c monomial \ 'c set \ 'c monomial" where +"restrict_to_indices m S = filter_mset (\i. i \ S) m" + +lemma restrict_to_indicesE: + assumes "i \# restrict_to_indices m S" + shows "i \ S" + using assms + unfolding restrict_to_indices_def + by simp + +lemma restrict_to_indicesI[simp]: + assumes "i \# m" + assumes "i \ S" + shows "i \# restrict_to_indices m S" + using assms + unfolding restrict_to_indices_def + by simp + +lemma restrict_to_indices_not_in[simp]: + assumes "i \# m" + assumes "i \ S" + shows "i \# restrict_to_indices m S" + using assms + unfolding restrict_to_indices_def + by (meson count_eq_zero_iff count_filter_mset) + +lemma restrict_to_indices_submultiset[simp]: +"restrict_to_indices m S \# m" + unfolding restrict_to_indices_def + using multiset_filter_subset + by blast + +lemma restrict_to_indices_add_element: +"restrict_to_indices (add_mset x m) S = (if x \ S then (add_mset x (restrict_to_indices m S) ) + else (restrict_to_indices m S) )" + unfolding restrict_to_indices_def + by (metis filter_mset_add_mset) + +lemma restrict_to_indices_count[simp]: +"count (restrict_to_indices m S) i = (if (i \ S) then (count m i) else 0)" + unfolding restrict_to_indices_def + by (metis count_filter_mset) + +lemma restrict_to_indices_subset: +"restrict_to_indices m S = restrict_to_indices m (set_mset m \ S)" +proof(induction m) + case empty + then show ?case unfolding restrict_to_indices_def + by (metis filter_empty_mset) +next + case (add x m) + assume IH: "restrict_to_indices m S = restrict_to_indices m (set_mset m \ S)" + show "restrict_to_indices (add_mset x m) S = restrict_to_indices (add_mset x m) (set_mset (add_mset x m) \ S)" + proof- + have "\i. count (restrict_to_indices (add_mset x m) S) i = + count (restrict_to_indices (add_mset x m) (set_mset (add_mset x m) \ S)) i " + proof- + fix i + show "count (restrict_to_indices (add_mset x m) S) i = count (restrict_to_indices (add_mset x m) (set_mset (add_mset x m) \ S)) i" + apply(cases "i \ S") + using restrict_to_indices_count + apply (metis IntI count_inI) + by (metis restrict_to_indices_count Int_iff) + qed + then show ?thesis + using multiset_eq_iff by blast + qed +qed + +text\\texttt{Restrict\_to\_indices} only depends on the intersection + of the index set with the set of indices in m:\ + +lemma restrict_to_indices_subset': + assumes "(set_mset m) \ S = (set_mset m) \ S'" + shows "restrict_to_indices m S = restrict_to_indices m S'" + by (metis restrict_to_indices_subset assms) + +lemma mset_add_plus: + assumes "m = n + k" + shows "add_mset x m = (add_mset x n) + k" + using assms + by simp + +text\Restricting to \S\ and the complement of \S\ partitions \m\:\ + +lemma restrict_to_indices_decomp: + "m = (restrict_to_indices m S) + (restrict_to_indices m ((set_mset m) - S))" + apply(induction m) + apply (metis add.right_neutral empty_Diff restrict_to_indices_submultiset set_mset_empty subset_mset.le_zero_eq) +proof- + fix x + fix m + assume A: "m = restrict_to_indices m S + restrict_to_indices m (set_mset m - S)" + show "add_mset x m = restrict_to_indices (add_mset x m) S + restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S)" + proof(cases "x \ S") + case True + then have T0: "restrict_to_indices (add_mset x m) S = (add_mset x (restrict_to_indices m S) ) " + by (simp add: True restrict_to_indices_add_element) + have T1: "restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) = + restrict_to_indices m (set_mset (add_mset x m) - S)" + using True by (metis DiffD2 restrict_to_indices_add_element) + have T2: "restrict_to_indices (add_mset x m) S + restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) + = (add_mset x (restrict_to_indices m S) ) + restrict_to_indices m (set_mset (add_mset x m) - S)" + using T0 T1 + by presburger + have T3: " add_mset x m = add_mset x (restrict_to_indices m S) + restrict_to_indices m (set_mset m - S)" + using T2 A using mset_add_plus[of m "restrict_to_indices m S" " restrict_to_indices m (set_mset m - S)" x] + by blast + have T4: "set_mset m \ (set_mset (add_mset x m) - S) = set_mset m \ (set_mset m - S)" + proof + show "set_mset m \ (set_mset (add_mset x m) - S) \ set_mset m \ (set_mset m - S)" + by blast + show "set_mset m \ (set_mset m - S) \ set_mset m \ (set_mset (add_mset x m) - S)" + by (simp add: True) + qed + have T5: "restrict_to_indices m (set_mset (add_mset x m) - S) = restrict_to_indices m (set_mset m - S)" + using T4 restrict_to_indices_subset'[of m "(set_mset (add_mset x m) - S)" " (set_mset m - S)" ] + by blast + then show ?thesis using T4 + by (metis T0 T1 T3) + next + case False + then have F0: "restrict_to_indices (add_mset x m) S = (restrict_to_indices m S) " + by (simp add: False restrict_to_indices_add_element) + have F1: "restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) = + (add_mset x (restrict_to_indices m (set_mset (add_mset x m) - S)))" + using False + by (meson DiffI restrict_to_indices_add_element union_single_eq_member) + + have F2: "restrict_to_indices (add_mset x m) S + restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) + = (restrict_to_indices m S) + (add_mset x (restrict_to_indices m (set_mset (add_mset x m) - S)))" + using F0 F1 + by presburger + have F3: " add_mset x m = (restrict_to_indices m S) + (add_mset x (restrict_to_indices m (set_mset m - S)))" + using F2 A mset_add_plus[of m "restrict_to_indices m (set_mset m - S)" "restrict_to_indices m S" x] + by (metis union_mset_add_mset_right) + have F4: " add_mset x m = restrict_to_indices (add_mset x m) S+ (add_mset x (restrict_to_indices m (set_mset m - S)))" + using F0 F3 by auto + have F5: "add_mset x (restrict_to_indices m (set_mset m - S)) = restrict_to_indices (add_mset x m) (set_mset (add_mset x m) - S) " + proof(cases "x \ set_mset m") + case True + then show ?thesis + by (metis F1 add_mset_remove_trivial more_than_one_mset_mset_diff) + next + case F: False + have "set_mset m \ (set_mset (add_mset x m) - S) = set_mset m \(set_mset m - S)" + proof + show "set_mset m \ (set_mset (add_mset x m) - S) \ set_mset m \ (set_mset m - S)" + using F False + by blast + show "set_mset m \ (set_mset m - S) \ set_mset m \ (set_mset (add_mset x m) - S)" + using F False + by (metis Diff_mono mset_add_plus Int_mono add_cancel_right_left + set_eq_subset subsetI subset_iff union_commute union_iff) + qed + then show ?thesis + by (metis F1 restrict_to_indices_subset') + qed + then show ?thesis + using False F4 + by presburger + qed +qed + +definition remove_indices :: "'c monomial \ 'c set \ 'c monomial" where +"remove_indices m S = (restrict_to_indices m (set_mset m - S))" + +lemma remove_indices_decomp: +"m = (restrict_to_indices m S) + (remove_indices m S)" + unfolding remove_indices_def + using restrict_to_indices_decomp + by blast + +lemma remove_indices_indices[simp]: + assumes "set_mset m \ I" + shows "set_mset (remove_indices m S) \ I - S" + unfolding remove_indices_def using assms + by (meson Diff_iff restrict_to_indicesE subsetD subsetI) + +subsubsection\Total evaluation of a monomial\ + +text\ + We define total evaluation of a monomial first, and then define the partial evaluation of a + monomial in terms of this. +\ + +abbreviation(input) closed_fun where +"closed_fun R g \ g \ UNIV \ carrier R" + +definition monom_eval :: "('a, 'b) ring_scheme \ 'c monomial \ ('c \ 'a) \ 'a" where +"monom_eval R (m:: 'c monomial) g = fold_mset (\ x . \ y. if y \ carrier R then (g x) \\<^bsub>R\<^esub> y else \\<^bsub>R\<^esub>) \\<^bsub>R\<^esub> m" + +context cring +begin + +lemma closed_fun_simp: + assumes "closed_fun R g" + shows "g n \ carrier R" + using assms + by blast + +lemma closed_funI: + assumes "\x. g x \ carrier R" + shows "closed_fun R g" + by (meson Pi_I assms) + + +text\The following are necessary technical lemmas to prove properties of about folds over multisets:\ + +lemma monom_eval_comp_fun: + fixes g:: "'c \ 'a" + assumes "closed_fun R g" + shows "comp_fun_commute (\ x . \y. if y \ carrier R then (g x) \ y else \)" + unfolding comp_fun_commute_def +proof- + have "\ x y. (\ya. if ya \ carrier R then g y \ ya else \) \ (\y. if y \ carrier R then g x \ y else \) = + (\y. if y \ carrier R then g x \ y else \) \ (\ya. if ya \ carrier R then g y \ ya else \)" + proof + fix x y a + show "((\ya. if ya \ carrier R then g y \ ya else \) \ (\y. if y \ carrier R then g x \ y else \)) a = + ((\y. if y \ carrier R then g x \ y else \) \ (\ya. if ya \ carrier R then g y \ ya else \)) a" + proof(cases "a \ carrier R") + case True + then show ?thesis + proof- + have LHS: "((\ya. if ya \ carrier R then g y \ ya else \) \ (\y. if y \ carrier R then g x \ y else \)) a = + ((\ya. if ya \ carrier R then g y \ ya else \) (g x \ a))" + using True assms(1) m_closed m_lcomm + unfolding o_def + by presburger + + then have LHS': "((\ya. if ya \ carrier R then g y \ ya else \) \ (\y. if y \ carrier R then g x \ y else \)) a = g y \ (g x \ a) " + using True assms m_closed + by (meson PiE UNIV_I) + then show ?thesis + unfolding o_def + using True assms m_closed m_lcomm + by (smt PiE UNIV_I) + qed + next + case False + then show ?thesis + unfolding o_def + using assms r_null closed_fun_simp + by smt + qed + qed + then show " \y x. (\ya. if ya \ carrier R then g y \ ya else \) \ (\y. if y \ carrier R then g x \ y else \) = + (\y. if y \ carrier R then g x \ y else \) \ (\ya. if ya \ carrier R then g y \ ya else \)" + by blast +qed + +lemma monom_eval_car: + assumes "closed_fun R g" + shows "monom_eval R (m:: 'c monomial) g \ carrier R" +proof(induction m) +case empty + then show ?case + unfolding monom_eval_def + by (metis fold_mset_empty one_closed) +next + case (add x m) + fix x m + assume A: "monom_eval R m g \ carrier R" + obtain f where f_def: "f = (\ x . \y. if y \ carrier R then (g x) \ y else \)" + by blast + have 0: "comp_fun_commute f" + using assms monom_eval_comp_fun[of g] f_def + by blast + have 1: "\m. monom_eval R m g = fold_mset f \ m" + using f_def monom_eval_def + by blast + have 2: "monom_eval R (add_mset x m) g = fold_mset f \ (add_mset x m)" + using 1 by blast + have 3: "g x \ carrier R" + using assms by blast + then show "monom_eval R (add_mset x m) g \ carrier R" + using assms 0 1 2 3 + by (metis A comp_fun_commute.fold_mset_add_mset f_def m_closed) +qed + +text\Formula for recursive (total) evaluation of a monomial:\ +lemma monom_eval_add: + assumes "closed_fun R g" + shows "monom_eval R (add_mset x M) g = (g x) \ (monom_eval R M g)" +proof- + obtain f where f_def: "f = (\ x . \y. if y \ carrier R then (g x) \ y else \)" + by blast + have 0: "comp_fun_commute f" + using assms monom_eval_comp_fun f_def + by blast + have 1: "\m. monom_eval R m g = fold_mset f \ m" + using f_def monom_eval_def + by blast + have 2: "monom_eval R (add_mset x M) g = fold_mset f \ (add_mset x M)" + using 1 by blast + have 3: "g x \ carrier R" + using assms by blast + have 4: "(g x) \ (monom_eval R M g) = f x (monom_eval R M g)" + using f_def 3 + by (meson assms monom_eval_car) + then show ?thesis + by (metis "0" "1" comp_fun_commute.fold_mset_add_mset) +qed + +end + +text\ + This function maps a polynomial $P$ to the set of monomials in $P$ which, after evaluating all + variables in the set $S$ to values in the ring $R$, reduce to the monomial $n$. +\ + +definition monomials_reducing_to :: +"('a,'b) ring_scheme \ 'c monomial \ ('a,'c) mvar_poly \ 'c set \ ('c monomial) set" where +"monomials_reducing_to R n P S = {m \ monomials_of R P. remove_indices m S = n}" + +lemma monomials_reducing_to_subset[simp]: + "monomials_reducing_to R n P s \ monomials_of R P" + unfolding monomials_reducing_to_def + by blast + +context cring +begin + +lemma monomials_reducing_to_finite: + assumes "P \ Pring_set R I" + shows "finite (monomials_reducing_to R n P s)" + by (meson assms monomials_finite monomials_reducing_to_subset rev_finite_subset) + +lemma monomials_reducing_to_disjoint: + assumes "n1 \ n2" + shows "monomials_reducing_to R n1 P S \ monomials_reducing_to R n2 P S = {}" + unfolding monomials_reducing_to_def + using assms + by blast + +lemma monomials_reducing_to_submset: +assumes "n \# m" +shows "n \ monomials_reducing_to R m P S" +proof(rule ccontr) + assume C: "\ n \ monomials_reducing_to R m P S " + then have "n \ monomials_reducing_to R m P S " + by blast + then have "remove_indices n S = m" + unfolding monomials_reducing_to_def + by blast + then show False + by (metis (full_types) remove_indices_def restrict_to_indices_submultiset assms subset_mset.less_asym' subset_mset.less_irrefl subset_mset_def) +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Partial Evaluation of a Polynomial\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + This function takes as input a set $S$ of variables, an evaluation function $g$, and a polynomial + to evaluate $P$. The output is a polynomial which is the result of evaluating the variables from + the set $S$ which occur in $P$, according to the evaluation function $g$. +\ + +definition poly_eval :: + "('a,'b) ring_scheme \ 'c set \ ('c \ 'a) \ ('a, 'c) mvar_poly \ ('a, 'c) mvar_poly"where +"poly_eval R S g P m = (finsum R (\n. monom_eval R (restrict_to_indices n S) g \\<^bsub>R\<^esub> (P n)) (monomials_reducing_to R m P S))" + +context cring +begin + +lemma finsum_singleton: + assumes "S = {s}" + assumes "f s \ carrier R" + shows "finsum R f S = f s" +proof- + have "finsum R f S = finsum R f (insert s {})" + using assms(1) + by blast + then show ?thesis using finsum_insert[of "{}" s f] assms + by (metis Pi_I empty_iff finite.emptyI finsum_empty r_zero) +qed + +lemma poly_eval_constant: + assumes "k \ carrier R" + shows "poly_eval R S g (indexed_const k) = (indexed_const k)" +proof + have S: "monomials_of R (indexed_const k) \ {{#}}" + unfolding indexed_const_def monomials_of_def + by (metis (mono_tags, lifting) mem_Collect_eq singletonI subset_iff) + fix x + show "poly_eval R S g (indexed_const k) x = indexed_const k x " + proof(cases "x = {#}") + case True + have "(monomials_reducing_to R x (indexed_const k) S) \ {{#}}" + using S monomials_reducing_to_subset + by blast + then show ?thesis + proof(cases "k = \") + case True + then have "(monomials_reducing_to R x (indexed_const k) S) = {}" + by (metis S \monomials_reducing_to R x (indexed_const k) S \ {{#}}\ + monomials_reducing_to_subset monoms_of_const subset_antisym subset_singletonD) + then show ?thesis + unfolding poly_eval_def + by (metis True finsum_empty indexed_const_def) + next + case False + then have "monomials_of R (indexed_const k) = {{#}}" + by (meson monoms_of_const) + have "remove_indices {#} S = {#}" + using remove_indices_decomp by blast + then have "{#} \ monomials_reducing_to R x (indexed_const k) S" + using True False unfolding monomials_reducing_to_def + \monomials_of R (indexed_const k) = {{#}}\ + by blast + then have 0: "monomials_reducing_to R x (indexed_const k) S = {{#}}" + using \monomials_reducing_to R x (indexed_const k) S \ {{#}}\ + by blast + have 1: "restrict_to_indices {#} S = {#}" + using restrict_to_indices_submultiset remove_indices_decomp by blast + have 2: "monom_eval R (restrict_to_indices {#} S) g = \" + unfolding monom_eval_def + using 1 + by (metis fold_mset_empty) + have 3: "poly_eval R S g (indexed_const k) x = + (\n\{{#}}. monom_eval R (restrict_to_indices n S) g \ indexed_const k n)" + unfolding poly_eval_def + using 0 + by presburger + have 4: "(\n\{{#}}. monom_eval R (restrict_to_indices n S) g \ indexed_const k n) = monom_eval R (restrict_to_indices {#} S) g \ indexed_const k {#}" + using finsum_singleton[of "{{#}}" "{#}" "\n. monom_eval R (restrict_to_indices n S) g \ indexed_const k n" ] + by (metis "2" assms indexed_const_def l_one) + then show ?thesis unfolding poly_eval_def + using 0 1 2 + by (metis True assms indexed_const_def l_one) + qed + next + case False + then have F0: "(indexed_const k) x = \" + by (meson indexed_const_def) + have "(monomials_reducing_to R x (indexed_const k) S) = {}" + unfolding monomials_reducing_to_def + proof(rule ccontr) + assume A: "{m \ monomials_of R (indexed_const k). remove_indices m S = x} \ {}" + then obtain m where m_def: "m \ monomials_of R (indexed_const k) \ remove_indices m S = x" + by blast + then show False using A F0 + by (metis False S empty_eq_union empty_iff + remove_indices_decomp singletonD subset_singletonD) + qed + then show ?thesis + unfolding poly_eval_def + by (metis False finsum_empty indexed_const_def) + qed +qed + +lemma finsum_partition: + assumes "finite S" + assumes "f \ S \ carrier R" + assumes "T \ S" + shows "finsum R f S = finsum R f T \ finsum R f (S - T) " +proof- + have "\U. finite U \ U \ S \ finsum R f S = finsum R f U \ finsum R f (S - U) " + proof- + fix U + show "finite U \ U \ S \ finsum R f S = finsum R f U \ finsum R f (S - U) " + apply(erule finite.induct) + apply (metis Diff_empty assms(2) finsum_closed finsum_empty l_zero) + proof + fix A :: "'c set" fix a + assume A0: "finite A" + show "A \ S \ finsum R f S = finsum R f A \ finsum R f (S - A) \ insert a A \ S \ finsum R f S = finsum R f (insert a A) \ finsum R f (S - insert a A)" + + proof- + assume A1: "A \ S \ finsum R f S = finsum R f A \ finsum R f (S - A)" + assume A2: "insert a A \ S" + show "finsum R f S = finsum R f (insert a A) \ finsum R f (S - insert a A)" + apply(cases "a \ A") + apply (metis A1 A2 insert_absorb) + proof- + assume A3: "a \ A" + have A4: "f a \ carrier R" + by (metis A2 Pi_iff assms(2) insert_subset) + have A5: " finsum R f (insert a A) = f a \ finsum R f A" + + using A0 A1 A2 finsum_insert[of A a f] assms A3 + by blast + have A6: "a \ S" + using A2 by blast + have A7: "finsum R f S \ carrier R" + using assms(2) finsum_closed by blast + have A8: " finsum R f (S - A)\ carrier R" + using Diff_subset[of S A] Pi_iff assms(2) finsum_closed[of f] subsetD[of _ S ] + by (meson Pi_anti_mono in_mono) + have A9: "finsum R f A \ carrier R" + by (meson A2 Pi_anti_mono assms(2) finsum_closed insert_subset subsetD) + have A10: "finsum R f A = finsum R f S \ finsum R f (S - A)" + using A7 A8 A9 + by (metis A1 A2 add.inv_solve_right' insert_subset minus_eq) + have A11: " finsum R f (insert a A) = f a \ (finsum R f S \ finsum R f (S - A))" + using A5 A6 A1 A2 assms A10 + by presburger + then have A12: " finsum R f (insert a A) =finsum R f S \( f a \ finsum R f (S - A))" + using A4 A7 A8 add.inv_closed add.m_lcomm minus_eq by presburger + have A13: "finsum R f (insert a A) \ carrier R" + using A4 A5 A9 add.m_closed + by presburger + have A14: " finsum R f (S - insert a A) \ carrier R" + by (meson Diff_subset Pi_anti_mono assms(2) finsum_closed in_mono) + have A15: "finsum R f S = finsum R f (insert a A) \ ( f a \ finsum R f (S - A)) " + by (metis A12 A13 A4 A7 A8 add.inv_solve_right' minus_closed minus_eq) + have A16: "finsum R f S = finsum R f (insert a A) \ finsum R f (S - A) \ f a" + using A1 A2 A4 A5 A8 A9 add.inv_closed add.m_assoc add.m_comm insert_subset minus_closed minus_eq r_neg1 + unfolding a_minus_def + by (metis add.m_closed) + have A16: "finsum R f S = finsum R f (insert a A) \ (finsum R f (S - A) \ f a)" + unfolding a_minus_def + using A13 A16 A4 A8 add.inv_closed add.m_assoc minus_eq by presburger + have A17: "(finsum R f (S - A) \ f a) = finsum R f (S - insert a A) " + proof- + have A170: "(S - A) = insert a (S - insert a A)" + using A3 A6 by blast + have A171: "a \ S - insert a A" + by blast + then have "finsum R f (S - A) =(f a) \ finsum R f (S - insert a A) " + using A170 finsum_insert[of "(S - insert a A)" a f] + by (metis A4 Diff_subset Pi_anti_mono assms(1) assms(2) + rev_finite_subset subsetD ) + then show ?thesis + by (metis A1 A13 A14 A16 A2 A4 A5 A8 A9 add.l_cancel + add.m_assoc add.m_comm insert_subset minus_closed) + qed + then show "finsum R f S = finsum R f (insert a A) \ finsum R f (S - insert a A) " + using A16 + by presburger + qed + qed + qed + qed + then show ?thesis + by (meson assms(1) assms(3) rev_finite_subset) +qed + +lemma finsum_eq_parition: + assumes "finite S" + assumes "f \ S \ carrier R" + assumes "T \ S" + assumes "\x. x \ S - T \ f x = \" + shows "finsum R f S = finsum R f T" + using assms + by (metis add.finprod_mono_neutral_cong_right) + +lemma poly_eval_scalar_mult: + assumes "k \ carrier R" + assumes "closed_fun R g" + assumes "P \ Pring_set R I" + shows "poly_eval R S g (poly_scalar_mult R k P)= + (poly_scalar_mult R k (poly_eval R S g P))" +proof + fix m + show "poly_eval R S g (poly_scalar_mult R k P) m = poly_scalar_mult R k (poly_eval R S g P) m" + unfolding poly_eval_def poly_scalar_mult_def + proof- + have 0: "(\n\monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) = + (\n\monomials_reducing_to R m (\m. k \ P m) S. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) " + proof- + have 00: "monomials_reducing_to R m (\m. k \ P m) S \ monomials_reducing_to R m P S" + proof + fix x show " x \ monomials_reducing_to R m (\m. k \ P m) S \ x \ monomials_reducing_to R m P S" + unfolding monomials_reducing_to_def + using assms assms(1) monomials_ofE complement_of_monomials_of r_null[of k] + by (metis (no_types, lifting) mem_Collect_eq) + have 01: "(\n. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) \ monomials_reducing_to R m P S \ carrier R" + by (smt Pi_I Pring_cfs_closed assms(1) assms(2) assms(3) closed_fun_simp m_closed monom_eval_car) + have 02: "finite (monomials_reducing_to R m P S)" + using assms(3) monomials_reducing_to_finite + by blast + have 03: " (\x. x \ monomials_reducing_to R m P S - monomials_reducing_to R m (\m. k \ P m) S \ + monom_eval R (restrict_to_indices x S) g \ (k \ P x) = \)" + proof- + fix x + assume A: "x \ monomials_reducing_to R m P S - monomials_reducing_to R m (\m. k \ P m) S " + have "x \ monomials_of R ((\m. k \ P m))" + proof + assume "x \ monomials_of R (\m. k \ P m)" + then have "x \ monomials_reducing_to R m (\m. k \ P m) S" + using A + unfolding monomials_reducing_to_def + by blast + then show False + using A by blast + qed + then show "monom_eval R (restrict_to_indices x S) g \ (k \ P x) = \" + by (metis assms(2) complement_of_monomials_of monom_eval_car r_null) + qed + qed + have 01: " (\x. x \ monomials_reducing_to R m P S - monomials_reducing_to R m (\m. k \ P m) S \ monom_eval R (restrict_to_indices x S) g \ (k \ P x) = \)" + proof- fix x assume A: "x \ monomials_reducing_to R m P S - monomials_reducing_to R m (\m. k \ P m) S" + hence "x \ monomials_reducing_to R m (\m. k \ P m) S" + by blast + hence "(k \ P x) = \" unfolding monomials_reducing_to_def + by (metis (no_types, lifting) A DiffD1 complement_of_monomials_of mem_Collect_eq monomials_reducing_to_def) + thus "monom_eval R (restrict_to_indices x S) g \ (k \ P x) = \" + using monom_eval_car[of g] assms(2) r_null by presburger + qed + have 02: "finite (monomials_reducing_to R m P S)" + using assms(3) monomials_reducing_to_finite by blast + have 04: "(\n. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) \ monomials_reducing_to R m P S \ carrier R" + by (smt Pi_I assms(1) assms(2) assms(3) closed_fun_simp cring.axioms(1) is_cring m_closed monom_eval_car ring.Pring_cfs_closed) + show ?thesis + using 00 01 02 04 + finsum_eq_parition[of "monomials_reducing_to R m P S" + "(\n. monom_eval R (restrict_to_indices n S) g \ (k \ P n))" + "monomials_reducing_to R m (\m. k \ P m) S"] + by blast + qed + then have 1: " (\n\monomials_reducing_to R m (\m. k \ P m) S. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) + = (\n\monomials_reducing_to R m P S. k \ (monom_eval R (restrict_to_indices n S) g \ (P n))) " + proof- + have "\n. monom_eval R (restrict_to_indices n S) g \ (k \ P n) = k \ (monom_eval R (restrict_to_indices n S) g \ (P n))" + by (metis Pring_cfs_closed assms(1) assms(2) assms(3) m_lcomm monom_eval_car) + then show ?thesis + using 0 + by presburger + qed + show " (\n\monomials_reducing_to R m (\m. k \ P m) S. monom_eval R (restrict_to_indices n S) g \ (k \ P n)) + = k \(\n\monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ (P n)) " + proof- + have "(\n. monom_eval R (restrict_to_indices n S) g \ (P n)) \ (monomials_reducing_to R m P S) \ carrier R" + by (meson Pi_I Pring_cfs_closed assms(2) assms(3) m_closed monom_eval_car) + then have "k \ (\i\monomials_reducing_to R m P S. monom_eval R (restrict_to_indices i S) g \ P i) = + (\i\monomials_reducing_to R m P S. k \ (monom_eval R (restrict_to_indices i S) g \ P i))" + using finsum_rdistr[of "monomials_reducing_to R m P S" + k + "(\n. monom_eval R (restrict_to_indices n S) g \ (P n))"] + assms monomials_reducing_to_finite by blast + then show ?thesis + using 1 + by presburger + qed + qed +qed + +lemma poly_eval_monomial: + assumes "closed_fun R g" + assumes "\ \\" + shows "poly_eval R S g (mset_to_IP R m) + = poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices m S))" +proof + have 0: "monomials_of R (mset_to_IP R m) = {m}" + using assms monomials_of_mset_to_IP + by blast + + fix x + show "poly_eval R S g (mset_to_IP R m) x = + poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices m S)) x " + proof(cases "x = (remove_indices m S)") + case True + have "monomials_reducing_to R x (mset_to_IP R m) S = {m}" + unfolding monomials_reducing_to_def + using True 0 + by auto + then have "poly_eval R S g (mset_to_IP R m) x = monom_eval R (restrict_to_indices m S) g \ (mset_to_IP R m m)" + unfolding poly_eval_def + by (metis (mono_tags, lifting) assms(1) finsum_singleton monom_eval_car mset_to_IP_simp r_one) + then show ?thesis + by (metis True mset_to_IP_simp poly_scalar_mult_def) + next + case False + then have "monomials_reducing_to R x (mset_to_IP R m) S = {}" + unfolding monomials_reducing_to_def + using 0 + by auto + then have "poly_eval R S g (mset_to_IP R m) x = \" + unfolding poly_eval_def + by (metis finsum_empty) + then show ?thesis + using False + by (metis assms(1) monom_eval_car mset_to_IP_simp' poly_scalar_mult_def r_null) + qed +qed + + +lemma(in cring) poly_eval_monomial_closed: + assumes "closed_fun R g" + assumes "\ \\" + assumes "set_mset m \ I" + shows "poly_eval R S g (mset_to_IP R m) \ Pring_set R (I - S)" +proof- + have "(mset_to_IP R (remove_indices m S)) \ Pring_set R (I - S)" + using assms mset_to_IP_closed[of "(remove_indices m S)" "I - S"] + by (metis remove_indices_indices) + then show ?thesis + using assms poly_eval_monomial[of g S m ] + poly_scalar_mult_closed[of "(monom_eval R (restrict_to_indices m S) g)" + "(mset_to_IP R (remove_indices m S))"] + by (metis monom_eval_car) +qed + +lemma poly_scalar_mult_iter: + assumes "\ \\" + assumes "P \ Pring_set R I" + assumes "k \ carrier R" + assumes "n \ carrier R" + shows "poly_scalar_mult R k (poly_scalar_mult R n P) = poly_scalar_mult R (k \ n) P" + using assms + unfolding poly_scalar_mult_def + by (metis Pring_cfs_closed m_assoc) + +lemma poly_scalar_mult_comm: + assumes "\ \\" + assumes "P \ Pring_set R I" + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "poly_scalar_mult R a (poly_scalar_mult R b P) = poly_scalar_mult R b (poly_scalar_mult R a P)" + using assms poly_scalar_mult_iter m_comm[of a b] + by metis + +lemma poly_eval_monomial_term: + assumes "closed_fun R g" + assumes "\ \\" + assumes "set_mset m \ I" + assumes "k \ carrier R" + shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)) = poly_scalar_mult R (k\(monom_eval R (restrict_to_indices m S) g)) + (mset_to_IP R (remove_indices m S))" +proof- + have 0: "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)) = + poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m))" + using assms poly_eval_scalar_mult[of k g "mset_to_IP R m" I S] mset_to_IP_closed + by blast + have 1: "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)) = + poly_scalar_mult R k (poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices m S))) " + using 0 assms + by (metis poly_eval_monomial) + have 2: "mset_to_IP R (remove_indices m S) \ Pring_set R I" + using assms mset_to_IP_closed + by (metis Diff_subset remove_indices_def restrict_to_indicesE subset_iff) + have 3: "monom_eval R (restrict_to_indices m S) g \ carrier R" + using assms monom_eval_car by blast + show ?thesis + using 1 2 3 assms poly_scalar_mult_iter[of "mset_to_IP R (remove_indices m S)" I k "(monom_eval R (restrict_to_indices m S) g)"] + by presburger +qed + +lemma poly_eval_monomial_term_closed: + assumes "closed_fun R g" + assumes "\ \\" + assumes "set_mset m \ I" + assumes "k \ carrier R" + shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)) \ Pring_set R (I - S)" +proof- + have "(mset_to_IP R (remove_indices m S)) \ Pring_set R (I - S)" + using assms + by (meson remove_indices_indices mset_to_IP_closed) + then show ?thesis + using assms poly_eval_monomial_term remove_indices_indices[of m I S] + by (metis cring.cring_simprules(5) cring.monom_eval_car is_cring poly_scalar_mult_closed) +qed + +lemma finsum_split: + assumes "finite S" + assumes "f \ S \ carrier R" + assumes "g \ S \ carrier R" + assumes "k \ carrier R" + assumes "c \ S" + assumes "\s. s \ S \ s \ c \ f s = g s" + assumes "g c = f c \ k" + shows "finsum R g S = k \ finsum R f S" +proof- + have 0: "finsum R f S = f c \ finsum R f (S - {c})" + proof- + have "f \ S - {c} \ carrier R" + by (metis Pi_split_insert_domain assms(2) assms(5) insert_Diff) + then show ?thesis + using assms finsum_insert[of "S - {c}" c f] + by (metis DiffD2 Pi_iff finite_Diff insert_Diff singletonI) + qed + have 1: "finsum R g S = g c \ finsum R g (S - {c})" + proof- + have "g \ S - {c} \ carrier R" + by (metis Pi_split_insert_domain assms(3) assms(5) insert_Diff) + then show ?thesis + using assms finsum_insert[of "S - {c}" c g] + by (metis DiffD2 Pi_iff finite_Diff insert_Diff singletonI) + qed + have "finsum R f (S - {c}) = finsum R g (S - {c})" + using assms Diff_iff Pi_split_insert_domain finsum_cong'[of "S- {c}" "S - {c}" g f] + insert_Diff singletonI + by blast + then have "finsum R g S = f c \ k \ finsum R g (S - {c})" + using assms \finsum R g S = g c \ finsum R g (S - {c})\ + by presburger + then have 1: "finsum R g S = f c \ k \ finsum R f (S - {c})" + using \finsum R f (S - {c}) = finsum R g (S - {c})\ by presburger + have "finsum R g S = k \ ( f c \ finsum R f (S - {c}))" + proof- + have "f c \ carrier R" + by (metis PiE assms(2) assms(5)) + have "finsum R f (S - {c}) \ carrier R" + by (metis Pi_split_insert_domain assms(2) assms(5) finsum_closed insert_Diff) + then show ?thesis using assms(4) 1 + using \f c \ carrier R\ add.m_assoc add.m_lcomm + by presburger + qed + then show ?thesis + using "0" + by presburger +qed + +lemma poly_monom_induction: +assumes "P (indexed_const \)" +assumes "\m k. set_mset m \ I \ k \ carrier R \ P (poly_scalar_mult R k (mset_to_IP R m))" +assumes "\Q m k. Q \ Pring_set R I \ (P Q) \ set_mset m \ I \ k \ carrier R\ P (Q \ (poly_scalar_mult R k (mset_to_IP R m)))" +shows "\Q. Q \ Pring_set R I \ P Q" +proof- + have 0: "\Ms. finite Ms \ (\ Q \ Pring_set R I. monomials_of R Q = Ms \ P Q)" + proof- + fix Ms + show " finite Ms \ (\ Q \ Pring_set R I. monomials_of R Q = Ms \ P Q)" + proof(erule finite.induct) + show "\Q\Pring_set R I. monomials_of R Q = {} \ P Q" + proof + fix Q + assume "Q \ Pring_set R I" + show "monomials_of R Q = {} \ P Q " + using assms + by (metis \Q \ Pring_set R I\ card_0_eq monomials_finite monomials_of_card_zero) + qed + show "\A a. finite A \ \Q\Pring_set R I. monomials_of R Q = A \ P Q \ + \Q\Pring_set R I. monomials_of R Q = insert a A \ P Q" + proof + fix A :: "'c monomial set" fix a fix Q + assume A0: "finite A" + assume A1: "\Q\Pring_set R I. monomials_of R Q = A \ P Q" + assume A2: " Q \ Pring_set R I " + show "monomials_of R Q = insert a A \ P Q" + proof + assume A3: "monomials_of R Q = insert a A" + show "P Q" + apply(cases "a \ A") + apply (metis A1 A2 A3 insert_absorb) + proof- + assume A4: "a \ A" + show "P Q" + proof- + have A5: "set_mset a \ I" + by (metis A2 A3 insert_iff mset_to_IP_indices) + have A6: "set_mset a \ I \ Q a \ carrier R" + using A2 A5 Pring_cfs_closed by blast + obtain Q' where Q'_def: "Q' = remove_monom R a Q" + by simp + then have "Q = Q' \ poly_scalar_mult R (Q a) (mset_to_IP R a)" + using A2 cring.remove_monom_eq is_cring by blast + then show ?thesis using A6 assms + by (metis A1 A2 A3 A4 Diff_empty Diff_insert0 Q'_def insert_Diff1 + remove_monom_closed remove_monom_monomials singletonI) + qed + qed + qed + qed + qed + qed + show "\Q. Q \ Pring_set R I \ P Q" + proof- + fix Q + assume "Q \ Pring_set R I" + show "P Q" + using 0[of "monomials_of R Q"] \Q \ Pring_set R I\ monomials_finite + by blast + qed +qed + +lemma Pring_car_induct: + assumes "q \ carrier (Pring R I)" + assumes "P \\<^bsub>Pring R I\<^esub>" + assumes "\m k. set_mset m \ I \ k \ carrier R \ P (k \\<^bsub>Pring R I\<^esub>(mset_to_IP R m))" + assumes "\Q m k. Q \ carrier (Pring R I) \ (P Q) \ set_mset m \ I \ k \ carrier R\ + P (Q \ (k \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)))" + shows "P q" + using poly_monom_induction[of P I q] assms Pring_smult[of I] Pring_car[of I] Pring_zero + by metis + +lemma poly_monom_induction2: +assumes "P (indexed_const \)" +assumes "\m k. set_mset m \ I \ k \ carrier R \ P (poly_scalar_mult R k (mset_to_IP R m))" +assumes "\Q m k. Q \ Pring_set R I \ (P Q) \ set_mset m \ I \ k \ carrier R \ P (Q \ (poly_scalar_mult R k (mset_to_IP R m)))" +assumes "Q \ Pring_set R I" +shows "P Q" + using assms(1) assms(2) assms(3) assms(4) poly_monom_induction by blast + +lemma poly_monom_induction3: +assumes "Q \ Pring_set R I" +assumes "P (indexed_const \)" +assumes "\m k. set_mset m \ I \ k \ carrier R \ P (poly_scalar_mult R k (mset_to_IP R m))" +assumes "\p q. p \ Pring_set R I \ (P p) \ q \ Pring_set R I \ (P q) \ P (p \ q)" +shows "P Q" + apply(rule poly_monom_induction2[of _ I]) + using assms(2) apply blast + using assms(3) apply blast + apply (meson assms(3) assms(4) mset_to_IP_closed poly_scalar_mult_closed) + using assms(1) by blast + +lemma Pring_car_induct': +assumes "Q \ carrier (Pring R I)" +assumes "P \\<^bsub>Pring R I\<^esub>" +assumes "\m k. set_mset m \ I \ k \ carrier R \ P (k \\<^bsub>Pring R I\<^esub> mset_to_IP R m)" +assumes "\p q. p \ carrier (Pring R I) \ (P p) \ q \ carrier (Pring R I) \ (P q) \ P (p \\<^bsub>Pring R I\<^esub> q)" +shows "P Q" + using poly_monom_induction3[of Q I P] assms Pring_smult Pring_add Pring_zero Pring_car + by metis + +lemma poly_eval_mono: + assumes "P \ Pring_set R I" + assumes "closed_fun R g" + assumes "finite F" + assumes "monomials_reducing_to R m P S \ F" + assumes "\n. n \ F \ remove_indices n S = m" + shows "poly_eval R S g P m = (\n\ F. monom_eval R (restrict_to_indices n S) g \ P n)" +proof- + have 0: "(\n\ F. monom_eval R (restrict_to_indices n S) g \ P n)= + (\n\ F - (monomials_reducing_to R m P S). monom_eval R (restrict_to_indices n S) g \ P n) \ poly_eval R S g P m" + proof- + have 00: " (\n. monom_eval R (restrict_to_indices n S) g \ P n) \ F \ carrier R" + by (meson Pi_I Pring_cfs_closed assms(1) assms(2) m_closed monom_eval_car) + have 01: "monomials_reducing_to R m P S \ F" + by (simp add: assms(4)) + have 02: "(\n\F. monom_eval R (restrict_to_indices n S) g \ P n) = + (\n\monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ P n) \ + (\n\F - monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ P n)" + using "00" "01" assms(3) finsum_partition by blast + have 03: " (\n\F - monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ P n)\ carrier R" + by (metis (mono_tags, lifting) "00" DiffD1 PiE Pi_I finsum_closed) + have " (\n\monomials_reducing_to R m P S. monom_eval R (restrict_to_indices n S) g \ P n) \ carrier R" + proof - + have "(\m. monom_eval R (restrict_to_indices m S) g \ P m) \ monomials_reducing_to R m P S \ carrier R" + using "00" "01" by blast + then show ?thesis + using finsum_closed by blast + qed + then show ?thesis + unfolding poly_eval_def + using 00 01 02 03 add.m_comm + by presburger + qed + have "(\n\ F - (monomials_reducing_to R m P S). monom_eval R (restrict_to_indices n S) g \ P n) = \" + proof- + have "\n. n \ F - (monomials_reducing_to R m P S) \ monom_eval R (restrict_to_indices n S) g \ P n = \" + proof- + fix n + assume A: "n \ F - (monomials_reducing_to R m P S)" + have "n \ monomials_of R P" + proof + assume "n \ monomials_of R P" + then have "n \ (monomials_reducing_to R m P S)" + unfolding monomials_reducing_to_def + using assms(5) A + by blast + then show False using A by blast + qed + then show " monom_eval R (restrict_to_indices n S) g \ P n = \" + by (metis assms(2) monom_eval_car complement_of_monomials_of r_null) + qed + then show ?thesis + by (meson add.finprod_one_eqI) + qed + then have 1: "(\n\F. monom_eval R (restrict_to_indices n S) g \ P n) = + \ \ poly_eval R S g P m" + using 0 Pi_I Pring_cfs_closed add.r_cancel_one assms(1) assms(2) + by presburger + have " (\n. monom_eval R (restrict_to_indices n S) g \ P n) \ monomials_reducing_to R m P S \ carrier R" + by (meson Pi_I Pring_cfs_closed assms(1) assms(2) m_closed monom_eval_car) + hence "poly_eval R S g P m \ carrier R" + using assms + unfolding poly_eval_def + using finsum_closed[of "\ n. monom_eval R (restrict_to_indices n S) g \ P n" + "monomials_reducing_to R m P S"] + by (meson Pi_I Pring_cfs_closed m_closed monom_eval_car) + then have "(\n\F. monom_eval R (restrict_to_indices n S) g \ P n) = poly_eval R S g P m" + using "1" add.r_cancel_one zero_closed + by presburger + then show ?thesis + by presburger +qed + +lemma finsum_group: + assumes "\n. f n \ carrier R" + assumes "\n. g n \ carrier R" + shows "finite S \ finsum R f S \ finsum R g S = finsum R (\n. f n \ g n) S " + apply(erule finite.induct) + apply (metis finsum_empty r_zero zero_closed) +proof- + fix A :: "'c set " + fix a + assume A0: "finite A" + assume A1: "finsum R f A \ finsum R g A = (\n\A. f n \ g n)" + show "finsum R f (insert a A) \ finsum R g (insert a A) = (\n\insert a A. f n \ g n)" + proof(cases "a \ A") + case True + then show ?thesis + by (metis A1 insert_absorb) + next + case False + have LHS: "finsum R f (insert a A) \ finsum R g (insert a A) = + (f a \ finsum R f A) \ (g a \ finsum R g A)" + using assms finsum_insert[of A a f] finsum_insert[of A a g] + by (metis A0 False Pi_I) + have F0: "(\n. f n \ g n) \ A \ carrier R" + using assms + by blast + have F1: "(f a \ g a) \ carrier R" + using assms + by blast + have RHS: " (\n\insert a A. f n \ g n) = (f a \ g a) \ (\n\A. f n \ g n)" + using F0 F1 assms finsum_insert[of A a " (\n. f n \ g n)"] False A0 + by blast + have F2: " f a \ finsum R f A \ (g a \ finsum R g A) = (f a \ g a) \ (finsum R f A \ finsum R g A)" + proof- + have F20: "f a \ carrier R" + using assms(1) by blast + have F21: "g a \ carrier R" + using assms(2) by blast + have F22: "finsum R f A \ carrier R" + by (metis Pi_iff assms(1) finsum_closed) + have F23: "finsum R g A \ carrier R" + by (metis Pi_I assms(2) finsum_closed) + show ?thesis using F21 F20 F22 F23 + using add.m_assoc add.m_closed add.m_lcomm + by presburger + qed + show ?thesis + using RHS LHS assms A1 F2 + by presburger + qed +qed + +lemma poly_eval_add: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + assumes "closed_fun R g" + shows "poly_eval R S g (P \ Q) = poly_eval R S g P \ poly_eval R S g Q " +proof + fix m + show "poly_eval R S g (P \ Q) m = (poly_eval R S g P \ poly_eval R S g Q) m" + proof- + obtain F where F_def: "F = monomials_reducing_to R m (P \ Q) S \ monomials_reducing_to R m P S \ + monomials_reducing_to R m Q S" + by simp + have 0: "finite F" + proof- + have 00: "finite (monomials_reducing_to R m (P \ Q) S)" + using assms + by (meson finite_subset monomials_finite monomials_of_add_finite monomials_reducing_to_subset) + have 01: "finite (monomials_reducing_to R m P S)" + using assms(1) monomials_reducing_to_finite by blast + have 02: "finite (monomials_reducing_to R m Q S)" + using assms(2) monomials_reducing_to_finite by blast + show ?thesis + using F_def 00 01 02 + by blast + qed + have 1: "\n. n \ F \ remove_indices n S = m" + proof- + fix n + assume A: "n \ F" + show "remove_indices n S = m" + using F_def + unfolding monomials_reducing_to_def + using A + by blast + qed + have 2: "poly_eval R S g (P \ Q) m = (\n\ F. monom_eval R (restrict_to_indices n S) g \ (P \ Q) n)" + using assms 0 1 poly_eval_mono[of "P \ Q" I g F m] F_def indexed_pset.indexed_padd + by blast + have 3: "poly_eval R S g P m = (\n\ F. monom_eval R (restrict_to_indices n S) g \ P n)" + using assms 0 1 poly_eval_mono[of "P" I g F m] F_def indexed_pset.indexed_padd + by blast + have 4: "poly_eval R S g Q m = (\n\ F. monom_eval R (restrict_to_indices n S) g \ Q n)" + using assms 0 1 poly_eval_mono[of "Q" I g F m] F_def indexed_pset.indexed_padd + by blast + have 5: "poly_eval R S g P m \ poly_eval R S g Q m = (\n\ F. monom_eval R (restrict_to_indices n S) g \ P n + \ monom_eval R (restrict_to_indices n S) g \ Q n)" + proof- + have 50: "(\n. monom_eval R (restrict_to_indices n S) g \ P n) \ F \ carrier R" + by (meson Pi_I Pring_cfs_closed assms(1) assms(3) m_closed monom_eval_car) + have 51: "(\n. monom_eval R (restrict_to_indices n S) g \ Q n) \ F \ carrier R" + by (meson Pi_I Pring_cfs_closed assms(2) assms(3) m_closed monom_eval_car) + then show ?thesis + using 0 2 3 50 finsum_group[of "(\n. monom_eval R (restrict_to_indices n S) g \ P n)" + "(\n. monom_eval R (restrict_to_indices n S) g \ Q n)" F] + by (metis (mono_tags, lifting) "4" Pring_cfs_closed assms(1) assms(2) assms(3) m_closed monom_eval_car) + qed + have 6: "poly_eval R S g P m \ poly_eval R S g Q m + = (\n\ F. monom_eval R (restrict_to_indices n S) g \( P n \ Q n))" + proof- + have 0 : "(\n. monom_eval R (restrict_to_indices n S) g \ P n \ monom_eval R (restrict_to_indices n S) g \ Q n) \ F \ carrier R" + apply(rule Pi_I) + by (meson Pring_cfs_closed add.m_closed assms(1) assms(2) assms(3) m_closed monom_eval_car) + have 1: "(\n. monom_eval R (restrict_to_indices n S) g \ (P n \ Q n)) \ F \ carrier R" + apply(rule Pi_I) + by (meson Pring_cfs_closed add.m_closed assms(1) assms(2) assms(3) m_closed monom_eval_car) + have "\n. n \ F \ monom_eval R (restrict_to_indices n S) g \( P n \ Q n) = monom_eval R (restrict_to_indices n S) g \ P n + \ monom_eval R (restrict_to_indices n S) g \ Q n" + using assms Pring_cfs_closed cring.monom_eval_car is_cring r_distr + by metis + then have "(\x\F. monom_eval R (restrict_to_indices x S) g \ P x \ monom_eval R (restrict_to_indices x S) g \ Q x) + = (\x\F. monom_eval R (restrict_to_indices x S) g \ (P x \ Q x))" + by (metis (no_types, lifting) restrict_ext) + then show ?thesis + using 5 finsum_eq[of "(\n. monom_eval R (restrict_to_indices n S) g \ P n \ monom_eval R (restrict_to_indices n S) g \ Q n)" + F "(\n. monom_eval R (restrict_to_indices n S) g \( P n \ Q n))" ] 0 1 + by presburger + qed + have 7: "monomials_reducing_to R m (P \ Q) S \ F" + using F_def + by blast + have 8: "poly_eval R S g (P \ Q) m = (\n\F. monom_eval R (restrict_to_indices n S) g \ (P \ Q) n)" + using 7 0 1 "2" + by blast + obtain f where f_def: "f = (\n. monom_eval R (restrict_to_indices n S) g \ (P \ Q) n)" + by blast + obtain h where h_def: "h = (\n. monom_eval R (restrict_to_indices n S) g \( P n \ Q n))" + by blast + have 9: "f \ F \ carrier R" + using f_def + by (metis (mono_tags, lifting) Pi_I Pring_cfs_closed add.m_closed + assms(1) assms(2) assms(3) indexed_padd_def m_closed monom_eval_car) + have 10: "h \ F \ carrier R" + using h_def + by (metis (mono_tags, lifting) "9" Pi_cong f_def indexed_padd_def) + have 11: "restrict f F = restrict h F" + using f_def h_def + by (metis indexed_padd_def) + have "finsum R f F = finsum R h F" + using 9 10 11 finsum_eq[of f F h] + by blast + then show ?thesis + using f_def h_def + by (metis (no_types, lifting) "6" "8" indexed_padd_def) + qed +qed + +lemma poly_eval_Pring_add: + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + assumes "closed_fun R g" + shows "poly_eval R S g (P \\<^bsub>Pring R I\<^esub> Q) = poly_eval R S g P \\<^bsub>Pring R I\<^esub> poly_eval R S g Q " + using assms poly_eval_add[of P I Q g S] + by (metis Pring_add Pring_car) + +text\Closure of partial evaluation maps:\ +lemma(in cring) poly_eval_closed: + assumes "closed_fun R g" + assumes "P \ Pring_set R I" + shows "poly_eval R S g P \ Pring_set R (I - S)" +proof- + obtain Pr where Pr_def[simp]: "Pr = (\Q. poly_eval R S g Q \ Pring_set R (I - S))" + by blast + have "Pr P" + apply(rule poly_monom_induction2[of _ I _ ]) + apply (metis Pr_def indexed_pset.indexed_const poly_eval_constant zero_closed) + apply (metis Pr_def assms(1) indexed_pset.indexed_const mset_to_IP_closed + poly_eval_constant poly_eval_monomial_term_closed poly_scalar_mult_zero r_null r_one) + proof- + show "P \ Pring_set R I" using assms by blast + show "\Q m k. Q \ Pring_set R I \ Pr Q \ set_mset m \ I \ k \ carrier R\ Pr (Q \ poly_scalar_mult R k (mset_to_IP R m))" + proof- + fix Q + fix m + fix k + assume A: "Q \ Pring_set R I \ Pr Q \ set_mset m \ I\ k \ carrier R" + then have 0: "poly_eval R S g Q \ Pring_set R (I - S)" + using Pr_def by blast + have 1: "poly_scalar_mult R k (mset_to_IP R m) \ Pring_set R I" + using A mset_to_IP_closed poly_scalar_mult_closed + by blast + have "poly_eval R S g (Q \ poly_scalar_mult R k (mset_to_IP R m)) = + poly_eval R S g Q \ poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)) " + using assms poly_eval_add[of Q I " (poly_scalar_mult R k (mset_to_IP R m))" g] "1" A + by blast + then show "Pr (Q \ poly_scalar_mult R k (mset_to_IP R m))" + using Pr_def + by (metis "1" A assms(1) indexed_pset.indexed_padd poly_eval_monomial_term_closed poly_scalar_mult_one poly_scalar_mult_zero) + qed + qed + then show ?thesis + using Pr_def + by blast +qed + +lemma poly_scalar_mult_indexed_pmult: + assumes "P \ Pring_set R I" + assumes "k \ carrier R" + shows " poly_scalar_mult R k (P \ i) = (poly_scalar_mult R k P) \ i" +proof- + have 0: "mset_to_IP R {#i#} \ Pring_set R (I \ {i})" + by (metis Un_upper2 mset_to_IP_closed set_mset_add_mset_insert set_mset_empty) + have 1: "P \ Pring_set R (I \ {i})" + using Pring_carrier_subset Un_upper1 assms(1) by blast + show ?thesis + using 0 1 poly_scalar_mult_times[of "mset_to_IP R {#i#}" "I \ {i}" P k] + poly_index_mult[of P "I \ {i}" i] assms + by (metis Un_upper2 insert_subset poly_index_mult poly_scalar_mult_closed) +qed + +lemma remove_indices_add_mset: + assumes "i \ S" + shows "remove_indices (add_mset i m) S = add_mset i (remove_indices m S)" + apply(induction m) + apply (smt assms empty_eq_union remove_indices_decomp restrict_to_indicesE single_is_union union_single_eq_member) + by (metis assms multi_union_self_other_eq remove_indices_decomp restrict_to_indices_add_element union_mset_add_mset_right) + +lemma poly_eval_monom_insert: + assumes "closed_fun R g" + assumes "\ \\" + assumes "i \ S" + shows "poly_eval R S g (mset_to_IP R (add_mset i m)) + = poly_scalar_mult R (g i) + (poly_eval R S g (mset_to_IP R m))" +proof- + have 0: "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (monom_eval R (restrict_to_indices (add_mset i m) S) g) + (mset_to_IP R (remove_indices (add_mset i m) S))" + using assms(1) assms(2) poly_eval_monomial by blast + have 1: "(mset_to_IP R (remove_indices (add_mset i m) S)) = + (mset_to_IP R (remove_indices m S))" + using assms + by (metis (full_types) Diff_iff insert_Diff1 remove_indices_def restrict_to_indices_add_element set_mset_add_mset_insert) + have 2: "(monom_eval R (restrict_to_indices (add_mset i m) S) g) = + (g i) \ ((monom_eval R (restrict_to_indices m S) g))" + using assms + by (metis monom_eval_add restrict_to_indices_add_element) + have 3: "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R ((g i) \ ((monom_eval R (restrict_to_indices m S) g))) + (mset_to_IP R (remove_indices m S))" + using 0 1 2 assms + by presburger + hence "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (g i) + (poly_scalar_mult R ((monom_eval R (restrict_to_indices m S) g)) + (mset_to_IP R (remove_indices m S)))" + using assms poly_scalar_mult_iter[of "(mset_to_IP R (remove_indices m S))" + UNIV "g i" "monom_eval R (restrict_to_indices m S) g"] + mset_to_IP_closed[of "remove_indices m S" UNIV] + by (metis PiE UNIV_I monom_eval_car subsetI) + thus ?thesis using assms + by (metis poly_eval_monomial) +qed + +lemma poly_eval_monom_insert': + assumes "closed_fun R g" + assumes "\ \\" + assumes "i \ S" + shows "poly_eval R S g (mset_to_IP R (add_mset i m)) + = (poly_eval R S g (mset_to_IP R m)) \ i" +proof- + have 0: "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (monom_eval R (restrict_to_indices (add_mset i m) S) g) + (mset_to_IP R (remove_indices (add_mset i m) S))" + using assms(1) assms(2) poly_eval_monomial by blast + hence "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices (add_mset i m) S))" + by (metis assms(3) restrict_to_indices_add_element) + hence "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (add_mset i (remove_indices m S)))" + by (metis assms(3) remove_indices_add_mset) + hence "poly_eval R S g (mset_to_IP R (add_mset i m)) = + poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices m S) \ i)" + by (metis monom_add_mset) + hence "poly_eval R S g (mset_to_IP R (add_mset i m)) = + (poly_scalar_mult R (monom_eval R (restrict_to_indices m S) g) + (mset_to_IP R (remove_indices m S))) \ i" + by (metis assms(1) cring.monom_eval_car is_cring local.ring_axioms + poly_scalar_mult_indexed_pmult ring.mset_to_IP_closed set_eq_subset) + thus ?thesis + by (metis assms(1) assms(2) poly_eval_monomial) +qed + +lemma poly_eval_indexed_pmult_monomial: + assumes "closed_fun R g" + assumes "k \ carrier R" + assumes "i \ S" + assumes "\ \ \" + shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m)))" +proof- + have 0: "poly_scalar_mult R k (mset_to_IP R m) \ i = + poly_scalar_mult R k (mset_to_IP R (add_mset i m ))" + using monom_add_mset[of i m] poly_scalar_mult_indexed_pmult + by (metis (no_types, opaque_lifting) assms(2) local.ring_axioms ring.mset_to_IP_closed subsetD subset_refl) + hence 1: "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + poly_scalar_mult R k (poly_eval R S g (mset_to_IP R (add_mset i m )))" + by (metis assms(1) assms(2) cring.poly_eval_scalar_mult is_cring local.ring_axioms ring.mset_to_IP_closed subsetI) + have 2: "poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m))) + = poly_scalar_mult R (g i) + (poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m))) " + by (smt assms(1) assms(2) local.ring_axioms poly_eval_scalar_mult ring.mset_to_IP_closed subsetD subset_refl) + hence 3: "poly_scalar_mult R (g i) (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m))) + = poly_scalar_mult R k + (poly_scalar_mult R (g i) (poly_eval R S g (mset_to_IP R m))) " + using assms poly_scalar_mult_comm[of "(poly_eval R S g (mset_to_IP R m))" UNIV] + poly_eval_closed[of g "(mset_to_IP R m)" UNIV] + by (metis UNIV_I closed_fun_simp cring.poly_scalar_mult_comm + is_cring local.ring_axioms ring.mset_to_IP_closed subsetI) + have 4: "(poly_eval R S g (mset_to_IP R (add_mset i m))) = (poly_scalar_mult R (g i) (poly_eval R S g (mset_to_IP R m)))" + using assms poly_eval_monom_insert[of g i S m] + by blast + thus ?thesis + using 1 3 poly_scalar_mult_comm assms + by presburger +qed + +lemma poly_eval_indexed_pmult_monomial': + assumes "closed_fun R g" + assumes "k \ carrier R" + assumes "i \ S" + assumes "\ \ \" + shows "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + (poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m))) \ i" +proof- + have "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + poly_scalar_mult R k (poly_eval R S g ( (mset_to_IP R m) \ i))" + using poly_eval_scalar_mult[of k g] + by (metis UNIV_I assms(1) assms(2) local.ring_axioms poly_scalar_mult_indexed_pmult ring.monom_add_mset ring.mset_to_IP_closed subsetI) + hence "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + poly_scalar_mult R k (poly_eval R S g (mset_to_IP R (add_mset i m)))" + by (simp add: monom_add_mset) + hence "poly_eval R S g (poly_scalar_mult R k (mset_to_IP R m) \ i) = + (poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m))) \ i" + using poly_eval_monom_insert'[of g i S m] + by (smt assms(1) assms(2) assms(3) assms(4) poly_eval_monomial_closed poly_scalar_mult_indexed_pmult subsetD subset_refl) + thus ?thesis using assms poly_eval_scalar_mult[of k g _ UNIV S] + by (metis UNIV_I mset_to_IP_closed subsetI) +qed + +lemma indexed_pmult_add: + assumes "p \ Pring_set R I" + assumes "q \ Pring_set R I" + shows "p \ q \ i = (p \ i) \ (q \ i)" + using assms poly_index_mult[of _ "I \ {i}"] + by (smt Pring_carrier_subset Set.basic_monos(1) Un_upper1 Un_upper2 cring.axioms(1) + insert_subset is_cring mk_disjoint_insert mset_to_IP_closed ring.P_ring_ldistr + ring.indexed_pset.intros(2) ring.indexed_pset_in_carrier set_mset_add_mset_insert + set_mset_empty) + +lemma poly_eval_indexed_pmult: + assumes "P \ Pring_set R I" + assumes "closed_fun R g" + shows "poly_eval R S g (P \ i) = (if i \ S then poly_scalar_mult R (g i) (poly_eval R S g P) else (poly_eval R S g P)\ i )" +proof(cases "i \ S") + case True + have "poly_eval R S g (P \ i) = poly_scalar_mult R (g i) (poly_eval R S g P)" + apply(rule poly_monom_induction3[of P I]) + using assms apply blast + apply (metis PiE UNIV_I assms(2) indexed_pmult_zero poly_eval_constant poly_scalar_mult_const r_null zero_closed) + apply (meson True assms(2) poly_eval_indexed_pmult_monomial) + apply (smt PiE True UNIV_I assms(2) genideal_one genideal_zero indexed_pmult_zero mset_to_IP_closed poly_eval_constant poly_eval_indexed_pmult_monomial poly_scalar_mult_one poly_scalar_mult_zero singletonD) + proof- fix p q assume A: + "p \ Pring_set R I" + "poly_eval R S g (p \ i) = poly_scalar_mult R (g i) (poly_eval R S g p)" + "q \ Pring_set R I" + "poly_eval R S g (q \ i) = poly_scalar_mult R (g i) (poly_eval R S g q)" + have "poly_eval R S g (p \ q \ i) = poly_eval R S g (p \ i) \ poly_eval R S g (q \ i)" + using assms poly_eval_add[of "p \ i" "I \ {i}" "q \ i" g S] + indexed_pmult_add[of p I q i] + by (smt A(1) A(3) Pring_carrier_subset Un_insert_right Un_upper1 indexed_pset.indexed_pmult insert_iff insert_subset mk_disjoint_insert) + hence "poly_eval R S g (p \ q \ i) = poly_scalar_mult R (g i) (poly_eval R S g p) \ + poly_scalar_mult R (g i) (poly_eval R S g q)" + using A + by presburger + hence "poly_eval R S g (p \ q \ i) = poly_scalar_mult R (g i) + ((poly_eval R S g p) \ (poly_eval R S g q))" + using Pring_smult poly_eval_closed[of g] A Pring_add Pring_car \poly_eval R S g (p \ q \ i) = poly_eval R S g (p \ i) \ poly_eval R S g (q \ i)\ assms(1) assms(2) closed_fun_simp + Pring_smult_r_distr[of "g i" "poly_eval R S g p" _ "poly_eval R S g q"] Pring_add Pring_car + by metis + thus "poly_eval R S g (p \ q \ i) = poly_scalar_mult R (g i) (poly_eval R S g (p \ q))" + by (metis A(1) A(3) assms(2) poly_eval_add) + qed + then show ?thesis + using True by presburger +next + case False + have "poly_eval R S g (P \ i) = (poly_eval R S g P) \ i" + apply(rule poly_monom_induction3[of P I]) + apply (simp add: assms(1)) + apply (metis indexed_pmult_zero poly_eval_constant zero_closed) + apply (metis False assms(2) indexed_pmult_zero inv_unique l_null mset_to_IP_closed + one_closed one_mset_to_IP poly_eval_constant poly_eval_indexed_pmult_monomial' + poly_scalar_mult_zero zero_closed) + proof-fix p q assume A: + "p \ Pring_set R I" + "poly_eval R S g (p \ i) = poly_eval R S g p \ i" + "q \ Pring_set R I" + "poly_eval R S g (q \ i) = poly_eval R S g q \ i" + have "poly_eval R S g (p \ q \ i) = poly_eval R S g (p \ i) \ poly_eval R S g (q \ i)" + using assms poly_eval_add[of "p \ i" "I \ {i}" "q \ i" g S] + indexed_pmult_add[of p I q i] + by (smt A(1) A(3) Pring_carrier_subset Un_insert_right Un_upper1 indexed_pset.indexed_pmult insert_iff insert_subset mk_disjoint_insert) + thus "poly_eval R S g (p \ q \ i) = poly_eval R S g (p \ q) \ i" + by (metis A(1) A(2) A(3) A(4) assms(2) indexed_pmult_add poly_eval_add poly_eval_closed) + qed + then show ?thesis + by (simp add: False) +qed + +lemma poly_eval_index: + assumes "\ \\" + assumes "closed_fun R g" + shows "poly_eval R S g (mset_to_IP R {#i#})= (if i \ S then (indexed_const (g i)) else mset_to_IP R {#i#})" +proof- + have 0: "poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (monom_eval R (restrict_to_indices {#i#} S) g) + (mset_to_IP R (remove_indices {#i#} S))" + using poly_eval_monomial[of g S "{#i#}" ] assms(1) assms(2) by blast + show ?thesis proof(cases "i \ S") + case True + then have T0: "(restrict_to_indices {#i#} S) = {#i#}" + by (metis restrict_to_indices_add_element restrict_to_indices_submultiset add_mset_subseteq_single_iff) + then have T1: "(monom_eval R (restrict_to_indices {#i#} S) g) = (monom_eval R {#i#} g)" + by presburger + then have "(monom_eval R (restrict_to_indices {#i#} S) g) = g i \ monom_eval R {#} g" + using assms monom_eval_add[of g i "{#}"] + by presburger + then have T2: "(monom_eval R (restrict_to_indices {#i#} S) g) = g i" + unfolding monom_eval_def + using T0 + by (metis PiE UNIV_I assms(2) fold_mset_empty r_one) + have T3: "(remove_indices {#i#} S) = {#}" + by (metis Diff_iff remove_indices_indices restrict_to_indicesE + T0 multiset_cases subset_iff union_single_eq_member) + then have T4: " (mset_to_IP R (remove_indices {#i#} S)) = indexed_const \" + by (metis one_mset_to_IP) + have T5: "poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (g i) (indexed_const \)" + using "0" T2 T4 + by presburger + then show ?thesis using True poly_scalar_mult_const + by (metis T2 assms(2) monom_eval_car one_closed r_one) + next + case False + have F0: "(restrict_to_indices {#i#} S) = {#}" + using False restrict_to_indices_def + by (metis restrict_to_indices_add_element filter_empty_mset) + have F1: "(monom_eval R (restrict_to_indices {#i#} S) g) = \" + using F0 + unfolding monom_eval_def + by (metis fold_mset_empty) + have F2: "(remove_indices {#i#} S) = {#i#}" + using False + by (metis Diff_iff remove_indices_def restrict_to_indices_add_element + restrict_to_indices_def filter_empty_mset set_mset_single singletonI) + have F3: "(mset_to_IP R (remove_indices {#i#} S)) = mset_to_IP R {#i#}" + by (simp add: F2) + have F4: " poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R \ (mset_to_IP R {#i#})" + using "0" F1 F3 + by presburger + show ?thesis using False + by (metis F4 mset_to_IP_closed poly_scalar_mult_one subset_iff) + qed +qed + +lemma poly_eval_indexed_pmult': + assumes "P \ Pring_set R I" + assumes "closed_fun R g" + assumes "i \ I" + shows "poly_eval R S g (P \\<^sub>p (mset_to_IP R {#i#})) = poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R {#i#})" +proof(cases "i \ S") + case True + have "(P \\<^sub>p mset_to_IP R {#i#})=(P \ i) " + using assms poly_index_mult + by metis + then have 0: " poly_eval R S g (P \\<^sub>p mset_to_IP R {#i#}) = poly_scalar_mult R (g i) (poly_eval R S g P)" + using True assms poly_eval_indexed_pmult[of P I g S i] + by presburger + have 1: "poly_eval R S g (mset_to_IP R {#i#}) = indexed_const (g i)" + using assms True + by (smt PiE UNIV_I genideal_one genideal_zero indexed_pmult_zero monom_add_mset one_mset_to_IP poly_eval_constant poly_eval_index singletonD) + then have "poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R {#i#})= poly_eval R S g P\\<^sub>p indexed_const (g i) " + by presburger + then have "poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R {#i#})= indexed_const (g i) \\<^sub>p poly_eval R S g P" + using assms P_ring_mult_comm[of "indexed_const (g i)" "poly_eval R S g P"] + unfolding carrier_coeff_def + by (metis "1" Pring_cfs_closed cring.closed_fun_simp indexed_pset.indexed_const is_cring poly_eval_closed) + then have "poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R {#i#})= poly_scalar_mult R (g i) (poly_eval R S g P)" + using assms + by (metis "0" "1" \P \\<^sub>p mset_to_IP R {#i#} = P \ i\ cring.closed_fun_simp is_cring poly_eval_closed poly_scalar_mult_eq) + then show ?thesis using 0 + by presburger +next + case False + then have 0: "poly_eval R S g (P \\<^sub>p (mset_to_IP R {#i#})) = (poly_eval R S g P)\ i " + using assms + by (metis poly_eval_indexed_pmult poly_index_mult) + have 1: "poly_eval R S g (mset_to_IP R {#i#})= mset_to_IP R {#i#}" + using False + by (metis assms(2) indexed_pmult_zero monom_add_mset one_closed one_mset_to_IP poly_eval_constant poly_eval_index) + then show ?thesis + using 0 False assms poly_eval_index[of g] + by (metis UNIV_I cring.Pring_set_restrict is_cring local.ring_axioms poly_eval_closed ring.poly_index_mult subsetI) +qed + +lemma poly_eval_monom_mult: + assumes "P \ Pring_set R I" + assumes "closed_fun R g" + shows "poly_eval R S g (P \\<^sub>p (mset_to_IP R m)) = poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R m) " +proof(induct m) + case empty + have 0: "mset_to_IP R {#} = indexed_const \" + using one_mset_to_IP by blast + then have 1: "(P \\<^sub>p mset_to_IP R {#}) = P" + using assms + by (metis P_ring_mult_comm Pring_cfs_closed carrier_coeff_def + mset_to_IP_simp mset_to_IP_simp' one_closed one_mult_left zero_closed) + have 2: "poly_eval R S g (mset_to_IP R {#}) = indexed_const \" + by (metis "0" one_closed poly_eval_constant) + show ?case using 0 1 2 + by (metis assms(1) assms(2) cring.P_ring_mult_comm indexed_pset.indexed_const + is_cring local.ring_axioms one_closed one_mult_left poly_eval_closed + ring.indexed_pset_in_carrier set_eq_subset) +next + case (add x m) + fix x + fix m + assume A: "poly_eval R S g (P \\<^sub>p mset_to_IP R m) = poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R m)" + show "poly_eval R S g (P \\<^sub>p mset_to_IP R (add_mset x m))= poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R (add_mset x m))" + proof- + obtain J where J_def: "J = I \ set_mset m \{x}" + by blast + have I0: "P \ Pring_set R J" + using J_def assms + by (meson Pring_carrier_subset Un_upper1 subsetD) + have I1: "set_mset m \ J" + using J_def by blast + have I2: "x \ J" + using J_def by blast + have "mset_to_IP R (add_mset x m)= (mset_to_IP R m) \ x" + by (simp add: monom_add_mset) + then have "(P \\<^sub>p mset_to_IP R (add_mset x m)) = P \\<^sub>p((mset_to_IP R m) \ x)" + by simp + then have "(P \\<^sub>p mset_to_IP R (add_mset x m)) = P \\<^sub>p((mset_to_IP R m) \\<^sub>p (mset_to_IP R {#x#}))" + by (metis add_mset_add_single monom_mult) + then have I3: "(P \\<^sub>p mset_to_IP R (add_mset x m)) = (P \\<^sub>p(mset_to_IP R m)) \\<^sub>p (mset_to_IP R {#x#})" + by (metis I0 P_ring_mult_assoc indexed_pset_in_carrier mset_to_IP_closed set_eq_subset) + have "poly_eval R S g (P \\<^sub>p mset_to_IP R m \\<^sub>p mset_to_IP R {#x#}) = + poly_eval R S g (P \\<^sub>p mset_to_IP R m) \\<^sub>p poly_eval R S g (mset_to_IP R {#x#})" + using poly_eval_indexed_pmult'[of "(P \\<^sub>p(mset_to_IP R m))" J g x] + I0 I1 I2 assms(2) assms(1) mset_to_IP_mult_closed + by blast + then have "poly_eval R S g (P \\<^sub>p mset_to_IP R (add_mset x m)) = + poly_eval R S g (P \\<^sub>p(mset_to_IP R m)) \\<^sub>p poly_eval R S g (mset_to_IP R {#x#}) " + using I3 + by simp + then have "poly_eval R S g (P \\<^sub>p mset_to_IP R (add_mset x m)) = + poly_eval R S g P \\<^sub>p poly_eval R S g (mset_to_IP R m) \\<^sub>p poly_eval R S g (mset_to_IP R {#x#}) " + by (simp add: A) + then have "poly_eval R S g (P \\<^sub>p mset_to_IP R (add_mset x m)) = + poly_eval R S g P \\<^sub>p (poly_eval R S g (mset_to_IP R m) \\<^sub>p poly_eval R S g (mset_to_IP R {#x#}))" + using P_ring_mult_assoc[of "poly_eval R S g P " "poly_eval R S g (mset_to_IP R m)" " poly_eval R S g (mset_to_IP R {#x#})"] + by (metis assms(1) assms(2) indexed_pset_in_carrier mset_to_IP_closed poly_eval_closed set_eq_subset) + then have "poly_eval R S g (P \\<^sub>p mset_to_IP R (add_mset x m)) = + poly_eval R S g P \\<^sub>p (poly_eval R S g ((mset_to_IP R m) \\<^sub>p (mset_to_IP R {#x#})))" + by (metis I1 I2 assms(2) mset_to_IP_closed poly_eval_indexed_pmult') + then show ?thesis + by (metis add_mset_add_single monom_mult) + qed +qed + +abbreviation mon_term ("Mt") where +"Mt k m \ poly_scalar_mult R k (mset_to_IP R m)" + +lemma poly_eval_monom_term_mult: + assumes "P \ Pring_set R I" + assumes "closed_fun R g" + assumes "k \ carrier R" + shows "poly_eval R S g (P \\<^sub>p (Mt k m)) = poly_eval R S g P \\<^sub>p poly_eval R S g (Mt k m) " +proof- + obtain J where J_def: "J = I \ (set_mset m)" + by blast + have J0: "P \ Pring_set R J" + using J_def Pring_carrier_subset Un_upper1 assms(1) by blast + have J1: "mset_to_IP R m \ Pring_set R J" + by (metis J_def Un_upper2 mset_to_IP_closed) + have 0: "(P \\<^sub>p (Mt k m)) = poly_scalar_mult R k (P \\<^sub>p mset_to_IP R m)" + using times_poly_scalar_mult[of P J "mset_to_IP R m" k ] J0 J1 assms(3) + by blast + have 1: "poly_eval R S g (P \\<^sub>p (Mt k m)) = poly_scalar_mult R k (poly_eval R S g (P \\<^sub>p mset_to_IP R m))" + by (metis "0" J0 J1 P_ring_mult_closed' assms(2) assms(3) cring.poly_eval_scalar_mult is_cring) + have 2: "poly_eval R S g (P \\<^sub>p (Mt k m)) = poly_scalar_mult R k ((poly_eval R S g P) \\<^sub>p (poly_eval R S g (mset_to_IP R m)))" + by (metis "1" assms(1) assms(2) poly_eval_monom_mult) + have 3: "poly_eval R S g (P \\<^sub>p (Mt k m)) = (poly_eval R S g P) \\<^sub>p poly_scalar_mult R k (poly_eval R S g (mset_to_IP R m))" + by (metis "0" "2" J0 J1 assms(2) assms(3) poly_eval_closed times_poly_scalar_mult) + then show ?thesis + by (metis J1 assms(3) assms(2) poly_eval_scalar_mult) +qed + +lemma poly_eval_mult: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + assumes "closed_fun R g" + shows "poly_eval R S g (P \\<^sub>p Q) = poly_eval R S g P \\<^sub>p poly_eval R S g Q " +proof- + obtain Pr where Pr_def: "Pr = (\Q. poly_eval R S g (P \\<^sub>p Q) = poly_eval R S g P \\<^sub>p poly_eval R S g Q )" + by blast + have "Pr Q" + proof(rule poly_monom_induction2[of _ I]) + show "Q \ Pring_set R I" + by (simp add: assms(2)) + show "Pr (indexed_const \)" + proof- + have 0: "(P \\<^sub>p indexed_const \) = indexed_const \" + by (metis assms(1) cring.P_ring_mult_comm indexed_pset.indexed_const + indexed_pset_in_carrier is_cring poly_scalar_mult_eq poly_scalar_mult_zero + set_eq_subset zero_closed) + have 1: "poly_eval R S g (indexed_const \) = indexed_const \" + using poly_eval_constant by blast + have 2: "poly_eval R S g P \\<^sub>p poly_eval R S g (indexed_const \) = indexed_const \" + using 1 + by (metis "0" assms(1) assms(3) local.ring_axioms one_closed poly_eval_monom_term_mult + poly_scalar_mult_const r_one ring.one_mset_to_IP zero_closed) + have 3: "poly_eval R S g (P \\<^sub>p indexed_const \) = indexed_const \" + using "0" "1" by presburger + show ?thesis + using 2 3 Pr_def + by presburger + qed + show "\m k. set_mset m \ I \ k \ carrier R \ Pr (poly_scalar_mult R k (mset_to_IP R m))" + using Pr_def assms(1) assms(3) poly_eval_monom_term_mult by blast + show " \Q m k. Q \ Pring_set R I \ Pr Q \ set_mset m \ I \ k \ carrier R \ Pr (Q \ Mt k m)" + proof- + fix Q + fix m + fix k + assume A: "Q \ Pring_set R I \ Pr Q \ set_mset m \ I \ k \ carrier R" + show "Pr (Q \ Mt k m)" + proof- + have 0: "poly_eval R S g (P \\<^sub>p Q) = poly_eval R S g P \\<^sub>p poly_eval R S g Q " + using A Pr_def by blast + have 1: "poly_eval R S g (P \\<^sub>p (Mt k m)) = poly_eval R S g P \\<^sub>p poly_eval R S g (Mt k m)" + using A assms(1) assms(3) poly_eval_monom_term_mult + by blast + have 2: "P \\<^sub>p (Q \ Mt k m) =(P \\<^sub>p Q)\ (P \\<^sub>p (Mt k m)) " + by (meson A assms(1) local.ring_axioms mset_to_IP_closed poly_scalar_mult_closed + ring.P_ring_rdistr ring.indexed_pset_in_carrier set_eq_subset) + have 3: "poly_eval R S g (P \\<^sub>p (Q \ Mt k m))= poly_eval R S g (P \\<^sub>p Q) \ poly_eval R S g (P \\<^sub>p (Mt k m)) " + by (metis "2" A P_ring_mult_closed' assms(1) assms(3) + mset_to_IP_closed poly_eval_add poly_scalar_mult_closed) + have 4: "poly_eval R S g (P \\<^sub>p (Q \ Mt k m))= + (poly_eval R S g P \\<^sub>p poly_eval R S g Q) \ ( poly_eval R S g P \\<^sub>p poly_eval R S g (Mt k m) )" + by (simp add: "0" "1" "3") + have 5: " poly_eval R S g P \\<^sub>p (poly_eval R S g Q \ poly_eval R S g (Mt k m)) = poly_eval R S g P \\<^sub>p poly_eval R S g Q \ (poly_eval R S g P \\<^sub>p poly_eval R S g (Mt k m))" + using ring.P_ring_rdistr[of R "(poly_eval R S g P)" "(poly_eval R S g Q)" "( poly_eval R S g (Mt k m) )"] + by (meson A assms(1) assms(3) indexed_pset_in_carrier local.ring_axioms + poly_eval_closed poly_scalar_mult_closed ring.mset_to_IP_closed subset_refl) + have 6: "poly_eval R S g (P \\<^sub>p (Q \ Mt k m))= + (poly_eval R S g P) \\<^sub>p ((poly_eval R S g Q) \ ( poly_eval R S g (Mt k m) ))" + using 4 5 + by simp + have 7: "poly_eval R S g (P \\<^sub>p (Q \ Mt k m))= + (poly_eval R S g P) \\<^sub>p (poly_eval R S g (Q\(Mt k m)))" + using 6 poly_eval_add[of "(poly_eval R S g Q)" I "(Mt k m)" g ] + by (metis A assms(3) mset_to_IP_closed poly_eval_add poly_scalar_mult_closed) + show ?thesis using 7 + using Pr_def by blast + qed + qed + qed + then show ?thesis + using Pr_def by blast +qed + +lemma poly_eval_Pring_mult: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + assumes "closed_fun R g" + shows "poly_eval R S g (P \\<^bsub>Pring R I\<^esub> Q) = poly_eval R S g P \\<^bsub>Pring R I\<^esub> poly_eval R S g Q " + by (metis Pring_mult assms(1) assms(2) assms(3) poly_eval_mult) + +lemma poly_eval_smult: + assumes "P \ Pring_set R I" + assumes "a \ carrier R" + assumes "closed_fun R g" + shows "poly_eval R S g (a \\<^bsub>Pring R I\<^esub> P) =a \\<^bsub>Pring R I\<^esub> poly_eval R S g P" + using poly_eval_scalar_mult assms + by (metis Pring_smult) + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Partial Evaluation is a Homomorphism\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma poly_eval_ring_hom: + assumes "I \ J" + assumes "closed_fun R g" + assumes "J - S \ I" + shows "ring_hom_ring (Pring R J) (Pring R I) (poly_eval R S g)" + apply(rule ring_hom_ringI) + apply (simp add: Pring_is_ring) + apply (simp add: Pring_is_ring) + apply (metis (full_types) Pring_car Pring_carrier_subset assms(2) assms(3) poly_eval_closed subset_iff) + apply (metis Pring_car assms(2) local.ring_axioms poly_eval_mult ring.Pring_mult) + apply (metis Pring_add Pring_car assms(2) poly_eval_add) + by (metis Pring_one one_closed poly_eval_constant) + +text\\texttt{poly\_eval} R at the zero function is an inverse to the inclusion of polynomial rings\ + +lemma poly_eval_zero_function: + assumes "g = (\n. \)" + assumes "J - S = I" + shows "P \ Pring_set R I \ poly_eval R S g P = P" + apply(erule indexed_pset.induct) + using poly_eval_constant apply blast + using assms poly_eval_add[of _ I _ g S] zero_closed + apply (metis Pi_I) + using Diff_iff assms(1) assms(2) Pi_I[of UNIV g ] + poly_eval_indexed_pmult set_eq_subset subsetD zero_closed + by smt + +lemma poly_eval_eval_function_eq: + assumes "closed_fun R g" + assumes "closed_fun R g'" + assumes "restrict g S = restrict g' S" + assumes "P \ Pring_set R I" + shows "poly_eval R S g P = poly_eval R S g' P" + apply(rule indexed_pset.induct[of P I "carrier R"]) + apply (simp add: assms(4)) + apply (metis poly_eval_constant) + apply (metis assms(1) assms(2) poly_eval_add) +proof- fix P i assume A: "P \ Pring_set R I" "poly_eval R S g P = poly_eval R S g' P" "i \ I " + show "poly_eval R S g (P \ i) = poly_eval R S g' (P \ i)" + proof(cases "i \ S") + case True + then have "g i = g' i" + using assms + unfolding restrict_def + by meson + then show ?thesis + using assms A poly_eval_indexed_pmult[of P I g S i] poly_eval_indexed_pmult[of P I g' S i] + by presburger + next + case False + then show ?thesis + using assms A poly_eval_indexed_pmult[of P I g S i] poly_eval_indexed_pmult[of P I g' S i] + by presburger + qed +qed + +lemma poly_eval_eval_set_eq: + assumes "closed_fun R g" + assumes "S \ I = S' \ I" + assumes "P \ Pring_set R I" + assumes "\ \\" + shows "poly_eval R S g P = poly_eval R S' g P" + apply(rule indexed_pset.induct[of P I "carrier R"]) + apply (simp add: assms(3)) + apply (metis poly_eval_constant) + apply (metis assms(1) poly_eval_add) +proof- fix P i + assume A: " P \ Pring_set R I" "poly_eval R S g P = poly_eval R S' g P" "i \ I " + show "poly_eval R S g (P \ i) = poly_eval R S' g (P \ i)" + using assms poly_eval_index[of g _ i] A + by (metis Diff_Diff_Int Diff_iff poly_eval_indexed_pmult) +qed + +lemma poly_eval_trivial: + assumes "closed_fun R g" + assumes "P \ Pring_set R (I - S)" + shows "poly_eval R S g P = P" + apply(rule indexed_pset.induct[of P "I - S" "carrier R"]) + apply (simp add: assms(2)) + using poly_eval_constant apply blast + apply (metis assms(1) poly_eval_add) + by (metis Diff_iff assms(1) poly_eval_indexed_pmult) + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Total Evaluation of a Polynomial\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma zero_fun_closed: +"closed_fun R (\n. \)" + by blast + +lemma deg_zero_cf_eval: + shows "P \ Pring_set R I \ poly_eval R I (\n. \) P = indexed_const (P {#})" + apply(erule indexed_pset.induct) + apply (metis indexed_const_def poly_eval_constant) +proof- + show " \P Q. P \ Pring_set R I \ + poly_eval R I (\n. \) P = indexed_const (P {#}) \ + Q \ Pring_set R I \ poly_eval R I (\n. \) Q = indexed_const (Q {#}) \ poly_eval R I (\n. \) (P \ Q) = indexed_const ((P \ Q) {#})" + proof- + fix P + fix Q + assume A: "P \ Pring_set R I" + assume B: " poly_eval R I (\n. \) P = indexed_const (P {#})" + assume C: "Q \ Pring_set R I " + assume D: " poly_eval R I (\n. \) Q = indexed_const (Q {#})" + have "indexed_const ((P \ Q) {#}) = indexed_const (P {#}) \ indexed_const (Q {#})" + by (metis indexed_padd_const indexed_padd_def) + thus "poly_eval R I (\n. \) (P \ Q) = indexed_const ((P \ Q) {#})" + using A B C D poly_eval_add[of P I Q "\n. \" I] + by (smt zero_fun_closed) + qed + show "\P i. P \ Pring_set R I \ poly_eval R I (\n. \) P = indexed_const (P {#}) \ i \ I \ poly_eval R I (\n. \) (P \ i) = indexed_const ((P \ i) {#})" + proof- + fix P + fix i + assume A: "P \ Pring_set R I" "poly_eval R I (\n. \) P = indexed_const (P {#})" "i \ I" + show "poly_eval R I (\n. \) (P \ i) = indexed_const ((P \ i) {#})" + proof- + have "poly_eval R I (\n. \) (P \ i) = indexed_const \" + using A(1) A(3) cring.poly_eval_constant is_cring poly_eval_indexed_pmult + poly_eval_scalar_mult poly_scalar_mult_zero zero_closed + by (metis Pi_I) + have "(P \ i) {#} = \" + using indexed_pmult_def + by (metis empty_iff set_mset_empty) + then show ?thesis + using \poly_eval R I (\n. \) (P \ i) = indexed_const \\ + by presburger + qed + qed +qed + +lemma deg_zero_cf_mult: + assumes "P \ Pring_set R I" + assumes "Q \ Pring_set R I" + shows " (P \\<^sub>p Q) {#} = P {#} \ Q {#}" +proof- + have "poly_eval R I (\n. \) (P \\<^sub>p Q) = poly_eval R I (\n. \) P \\<^sub>p poly_eval R I (\n. \) Q" + using zero_fun_closed assms + by (metis poly_eval_mult) + then have 0: "indexed_const ((P \\<^sub>p Q) {#}) = (indexed_const (P {#})) \\<^sub>p (indexed_const (Q {#}))" + by (metis P_ring_mult_closed' Pring_cfs_closed assms(1) assms(2) deg_zero_cf_eval indexed_const_P_mult_eq indexed_const_def) + have "indexed_const ((P \\<^sub>p Q) {#}) = (indexed_const ((P {#}) \(Q {#})))" + apply(rule ccontr) + using 0 + by (metis Pring_cfs_closed assms(1) assms(2) indexed_const_P_mult_eq) + then show ?thesis + proof - + show ?thesis + by (metis (no_types) \indexed_const ((P \\<^sub>p Q) {#}) = indexed_const (P {#} \ Q {#})\ + local.ring_axioms ring.indexed_const_def) + qed +qed + +definition deg_zero_cf :: "('a, 'c) mvar_poly \ 'a" where +"deg_zero_cf P = P {#}" + +lemma deg_zero_cf_ring_hom: + shows "ring_hom_ring (Pring R I) R (deg_zero_cf)" + apply(rule ring_hom_ringI) + using Pring_is_ring apply blast + apply (simp add: local.ring_axioms) + apply (metis deg_zero_cf_def Pring_car Pring_cfs_closed) + apply (metis deg_zero_cf_def Pring_car Pring_mult deg_zero_cf_mult) + apply (metis deg_zero_cf_def Pring_add indexed_padd_def) + by (metis deg_zero_cf_def Pring_one indexed_const_def) + +end + +definition eval_in_ring :: + "('a,'b) ring_scheme \ 'c set \ ('c \ 'a) \ ('a, 'c) mvar_poly \ 'a" where +"eval_in_ring R S g P = (poly_eval R S g P) {#}" + +definition total_eval :: +"('a,'b) ring_scheme \ ('c \ 'a) \ ('a, 'c) mvar_poly \ 'a" where +"total_eval R g P = eval_in_ring R UNIV g P" + +context cring +begin + +lemma eval_in_ring_ring_hom: + assumes "closed_fun R g" + shows "ring_hom_ring (Pring R I) R (eval_in_ring R S g)" + unfolding eval_in_ring_def + apply(rule ring_hom_ringI) + apply (simp add: Pring_is_ring) + apply (simp add: local.ring_axioms) + using Pring_car Pring_carrier_coeff' assms poly_eval_closed + apply (metis ) + using Pring_mult[of I] poly_eval_mult[of _ I _ g S] poly_eval_closed[of g _ I S] deg_zero_cf_mult[of _ I] assms + Pring_car[of I] + apply (metis deg_zero_cf_mult) + + using Pring_add[of I ] poly_eval_add[of _ I _ g S] Pring_car[of I] assms indexed_padd_def + apply metis + unfolding deg_zero_cf_def + by (metis Pring_one indexed_const_def one_closed poly_eval_constant) + +lemma eval_in_ring_smult: + assumes "P \ carrier (Pring R I)" + assumes "a \ carrier R" + assumes "closed_fun R g" + shows "eval_in_ring R S g (a \\<^bsub>Pring R I\<^esub> P) = a \ eval_in_ring R S g P " + using assms unfolding eval_in_ring_def + by (smt Pring_car Pring_smult poly_eval_scalar_mult poly_scalar_mult_def) + + +lemma total_eval_ring_hom: + assumes "closed_fun R g" + shows "ring_hom_ring (Pring R I) R (total_eval R g)" + using assms unfolding total_eval_def + using eval_in_ring_ring_hom by blast + +lemma total_eval_smult: + assumes "P \ carrier (Pring R I)" + assumes "a \ carrier R" + assumes "closed_fun R g" + shows "total_eval R g (a \\<^bsub>Pring R I\<^esub> P) = a \ total_eval R g P" + using assms unfolding total_eval_def + using eval_in_ring_smult by blast + +lemma total_eval_const: + assumes "k \ carrier R" + shows "total_eval R g (indexed_const k) = k" + unfolding total_eval_def eval_in_ring_def + using assms + by (metis indexed_const_def poly_eval_constant) + +lemma total_eval_var: + assumes "closed_fun R g" + shows "(total_eval R g (mset_to_IP R {#i#})) = g i" + unfolding total_eval_def eval_in_ring_def + using UNIV_I assms indexed_const_def indexed_pmult_zero monom_add_mset one_closed + one_mset_to_IP one_zeroD poly_eval_constant poly_eval_index singletonD + by (smt PiE iso_tuple_UNIV_I monom_eval_add monom_eval_car mset_to_IP_simp poly_scalar_mult_const r_one) + +lemma total_eval_indexed_pmult: + assumes "P \ carrier (Pring R I)" + assumes "i \ I" + assumes "closed_fun R g" + shows "total_eval R g (P \ i) = total_eval R g P \\<^bsub>R\<^esub> g i" +proof- + have "P \ i = P \\<^sub>p mset_to_IP R ((add_mset i) {#})" + using assms poly_index_mult[of P I i] Pring_car + by blast + then have 0: "(P \ i) = P \\<^bsub>Pring R I\<^esub> (mset_to_IP R {#i#})" + by (simp add: Pring_mult) + then have "total_eval R g (P \ i) = (total_eval R g P) \\<^bsub>R\<^esub> (total_eval R g (mset_to_IP R {#i#}))" + proof- + have "(mset_to_IP R {#i#}) \ carrier (Pring R I)" + by (metis Pring_car assms(2) mset_to_IP_closed set_mset_single singletonD subset_iff) + then show ?thesis + using assms total_eval_ring_hom[of g I] ring_hom_mult + by (metis "0" ring_hom_ring.homh) + qed + then show ?thesis + by (metis assms(3) cring.total_eval_var is_cring) +qed + +lemma total_eval_mult: + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + assumes "closed_fun R g" + shows "total_eval R g (P \\<^bsub>Pring R I\<^esub> Q) = (total_eval R g P) \\<^bsub>R\<^esub>(total_eval R g Q) " + by (metis assms(1) assms(2) assms(3) ring_hom_mult ring_hom_ring.homh total_eval_ring_hom) + +lemma total_eval_add: + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + assumes "closed_fun R g" + shows "total_eval R g (P \\<^bsub>Pring R I\<^esub> Q) = (total_eval R g P) \\<^bsub>R\<^esub>(total_eval R g Q) " + by (metis assms(1) assms(2) assms(3) ring_hom_add ring_hom_ring.homh total_eval_ring_hom) + +lemma total_eval_one: + assumes "closed_fun R g" + shows "total_eval R g \\<^bsub>Pring R I\<^esub> = \" + by (metis Pring_one one_closed total_eval_const) + +lemma total_eval_zero: + assumes "closed_fun R g" + shows "total_eval R g \\<^bsub>Pring R I\<^esub> = \" + by (metis Pring_zero total_eval_const zero_closed) + +lemma total_eval_closed: + assumes "P \ carrier (Pring R I)" + assumes "closed_fun R g" + shows "total_eval R g P \ carrier R" + using assms total_eval_ring_hom[of g] + by (metis ring_hom_closed ring_hom_ring.homh) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Constructing Homomorphisms from Indexed Polynomial Rings and a Universal Property\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\The inclusion of \R\ into its polynomial ring\ + +lemma indexed_const_ring_hom: +"ring_hom_ring R (Pring R I) (indexed_const)" + apply(rule ring_hom_ringI) + apply (simp add: local.ring_axioms) + apply (simp add: Pring_is_ring) + using Pring_car indexed_pset.indexed_const apply blast + apply (metis Pring_mult indexed_const_P_mult_eq) + apply (metis indexed_padd_const local.ring_axioms ring.Pring_add) + by (metis Pring_one) + +lemma indexed_const_inj_on: +"inj_on (indexed_const) (carrier R)" + by (metis cring.total_eval_const inj_onI is_cring) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Mapping $R[x] \to S[x]$ along a homomorphism $R \to S$\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition ring_hom_to_IP_ring_hom :: +"('a, 'e) ring_hom \ ('a, 'c) mvar_poly \ 'c monomial \ 'e" where +"ring_hom_to_IP_ring_hom \ P m = \ (P m)" + +context cring +begin + +lemma ring_hom_to_IP_ring_hom_one: + assumes "cring S" + assumes "ring_hom_ring R S \" + shows "ring_hom_to_IP_ring_hom \ \\<^bsub>Pring R I\<^esub> = \\<^bsub>Pring S I\<^esub>" + unfolding ring_hom_to_IP_ring_hom_def +proof + fix m + show " \ (\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>Pring S I\<^esub> m" + proof(cases "m = {#}") + case True + then have "(\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>R\<^esub>" + by (metis Pring_one indexed_const_def) + then have "\ (\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>S\<^esub>" + using assms + unfolding ring_hom_ring_def + by (metis assms(2) ring_hom_one ring_hom_ring.homh) + then show ?thesis using assms cring True ring.indexed_const_def ring_hom_ring_def + by (metis ring.Pring_one) + next + case False + then have "(\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>R\<^esub>" + by (metis indexed_const_def local.ring_axioms ring.Pring_one) + + then have "\ (\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>S\<^esub>" + using assms + unfolding ring_hom_ring_def cring_def + by (metis assms(2) ring_hom_zero ring_hom_ring.homh) + then show ?thesis using assms False ring.indexed_const_def ring_hom_ring_def + by (metis ring.Pring_one) + qed +qed + +lemma ring_hom_to_IP_ring_hom_constant: + assumes "cring S" + assumes "ring_hom_ring R S \" + assumes "a \ carrier R" + shows "ring_hom_to_IP_ring_hom \ ((indexed_const a):: 'c monomial \ 'a) = ring.indexed_const S (\ a)" + unfolding ring_hom_to_IP_ring_hom_def indexed_const_def +proof + fix m:: "'c monomial" + show "\ (if m = {#} then a else \) = ring.indexed_const S (\ a) m" + apply(cases "m = {#}") + apply (simp add: assms(1) cring.axioms(1) ring.indexed_const_def) + proof- + assume "m \ {#} " + then have "\ (if m = {#} then a else \) = \\<^bsub>S\<^esub>" + by (metis (full_types) assms(1) assms(2) cring_def local.ring_axioms + ring_hom_ring.homh ring_hom_zero) + then show " \ (if m = {#} then a else \) = ring.indexed_const S (\ a) m" + using assms + by (metis \m \ {#}\ cring.axioms(1) ring.indexed_const_def) + qed +qed + +lemma ring_hom_to_IP_ring_hom_add: + assumes "cring S" + assumes "ring_hom_ring R S \" + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + shows "ring_hom_to_IP_ring_hom \ (P \\<^bsub>Pring R I\<^esub> Q) = + (ring_hom_to_IP_ring_hom \ P) \\<^bsub>Pring S I\<^esub> (ring_hom_to_IP_ring_hom \ Q)" + unfolding ring_hom_to_IP_ring_hom_def +proof + fix m + show " \ ((P \\<^bsub>Pring R I\<^esub> Q) m) = ((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (Q m))) m" + proof- + have RHS: "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (Q m))) m = \ (P m) \\<^bsub>S\<^esub> \ (Q m)" + using ring.Pring_add[of S I _ _ ] assms + by (metis cring.axioms(1) ring.indexed_padd_def) + have LHS: "\ ((P \\<^bsub>Pring R I\<^esub> Q) m) = \ ((P m)\\<^bsub>R\<^esub> (Q m))" + by (metis Pring_add indexed_padd_def) + then show ?thesis + using assms unfolding ring_hom_ring_def + using ring_hom_add[of \ R S "P m" "Q m"] + by (metis Pring_carrier_coeff' RHS assms(2) ring_hom_ring.homh) + qed +qed + +lemma ring_hom_to_IP_ring_hom_closed: + assumes "cring S" + assumes "ring_hom_ring R S \" + assumes "P \ carrier (Pring R I)" + shows "ring_hom_to_IP_ring_hom \ P \ carrier (Pring S I)" + apply(rule indexed_pset.induct[of P I "carrier R"]) + using Pring_car assms(3) apply blast +proof- + show "\k. k \ carrier R \ ring_hom_to_IP_ring_hom \ (indexed_const k) \ carrier (Pring S I)" + proof- + fix k + show " k \ carrier R \ ring_hom_to_IP_ring_hom \ (indexed_const k) \ carrier (Pring S I)" + proof- + assume A: "k \ carrier R" + have "(\ k) \ carrier S" + by (meson A assms(2) ring_hom_closed ring_hom_ring.homh) + then have 0: "ring.indexed_const S (\ k) \ carrier (Pring S I)" + by (metis assms(1) cring.axioms(1) ring.Pring_car ring.indexed_pset.indexed_const) + then show ?thesis + using assms(2) + by (simp add: "0" A ring_hom_to_IP_ring_hom_constant[of S \ k] assms(1)) + qed + qed + show "\P Q. P \ Pring_set R I \ + ring_hom_to_IP_ring_hom \ P \ carrier (Pring S I) \ + Q \ Pring_set R I \ + ring_hom_to_IP_ring_hom \ Q \ carrier (Pring S I) \ ring_hom_to_IP_ring_hom \ (P \ Q) \ carrier (Pring S I)" + using ring_hom_to_IP_ring_hom_add[of S \ _ I] Pring_car[of I] ring.Pring_car[of S I] + by (metis Pring_add assms(1) assms(2) cring.axioms(1) ring.Pring_add_closed ) + show "\P i. P \ Pring_set R I \ + ring_hom_to_IP_ring_hom \ P \ carrier (Pring S I) \ i \ I \ ring_hom_to_IP_ring_hom \ (P \ i) \ carrier (Pring S I)" + proof- + fix P + fix i + assume A: "P \ Pring_set R I " "ring_hom_to_IP_ring_hom \ P \ carrier (Pring S I) " "i \ I" + have 0: "(\m. \ ((P \ i) m)) = ring.indexed_pmult S (ring_hom_to_IP_ring_hom \ P) i" + proof + fix m + show "\ ((P \ i) m) = ring.indexed_pmult S (ring_hom_to_IP_ring_hom \ P) i m" + proof(cases "i \# m") + case True + then have LHS: "\ ((P \ i) m) = \ (P (m - {#i#}))" + by (metis indexed_pmult_def) + then show ?thesis + using True + by (metis assms(1) cring.axioms(1) ring.indexed_pmult_def ring_hom_to_IP_ring_hom_def) + next + case False + then have "\ ((P \ i) m) = \ \\<^bsub>R\<^esub>" + by (metis indexed_pmult_def) + then have LHS: "\ ((P \ i) m) = \\<^bsub>S\<^esub>" + by (metis assms(1) assms(2) cring.axioms(1) local.ring_axioms + ring_hom_ring.homh ring_hom_zero) + then show ?thesis + using False assms ring.indexed_pmult_def + by (metis cring.axioms(1)) + qed + qed + then show "ring_hom_to_IP_ring_hom \ (P \ i) \ carrier (Pring S I)" + using assms + unfolding ring_hom_to_IP_ring_hom_def + by (metis "0" A(2) A(3) cring.axioms(1) ring.Pring_car ring.indexed_pset.simps) + qed +qed + +lemma ring_hom_to_IP_ring_hom_monom: + assumes "cring S" + assumes "ring_hom_ring R S \" + shows "ring_hom_to_IP_ring_hom \ (mset_to_IP R m) = mset_to_IP S m" +proof + fix x + show "ring_hom_to_IP_ring_hom \ (mset_to_IP R m) x = mset_to_IP S m x" + unfolding ring_hom_to_IP_ring_hom_def mset_to_IP_def apply( cases "x = m" ) + apply (metis (mono_tags, lifting) assms(2) ring_hom_one ring_hom_ring.homh) + by (metis (full_types) assms(1) assms(2) cring.axioms(1) local.ring_axioms + ring_hom_ring.homh ring_hom_zero) +qed + +lemma Pring_morphism: + assumes "cring S" + assumes "\ \ (carrier (Pring R I)) \ (carrier S)" + assumes "\ \\<^bsub>Pring R I\<^esub> = \\<^bsub>S\<^esub>" + assumes "\ \\<^bsub>Pring R I\<^esub> = \\<^bsub>S\<^esub>" + assumes "\P Q. P \ carrier (Pring R I) \ Q \ carrier (Pring R I) \ + \ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> (\ Q)" + assumes "\ i . \ P. i \ I \ P \ carrier (Pring R I) \ \ (P \ i) = (\ P) \\<^bsub>S\<^esub> (\ (mset_to_IP R {#i#}))" + assumes "\k Q. k \ carrier R \ Q \ carrier (Pring R I) \ \ (poly_scalar_mult R k Q) = + (\ (indexed_const k)) \\<^bsub>S\<^esub> (\ Q)" + shows "ring_hom_ring (Pring R I) S \" + apply(rule ring_hom_ringI) + apply (simp add: Pring_is_ring; fail) + apply (simp add: assms(1) cring.axioms(1); fail) + using assms(2) apply blast +proof- + + show "\x y. x \ carrier (Pring R I) \ y \ carrier (Pring R I) \ \ (x \\<^bsub>Pring R I\<^esub> y) = \ x \\<^bsub>S\<^esub> \ y" + proof- + fix P Q + assume A0: "P \ carrier (Pring R I)" + assume A1: "Q \ carrier (Pring R I)" + show "\ (P \\<^bsub>Pring R I\<^esub> Q) = \ P \\<^bsub>S\<^esub> \ Q" + proof(rule mpoly_induct'[of ]) + show "\ (P \\<^bsub>Pring R I\<^esub> indexed_const \) = \ P \\<^bsub>S\<^esub> \ (indexed_const \)" + proof- + have 0: "(P \\<^bsub>Pring R I\<^esub> indexed_const \) = \\<^bsub>Pring R I\<^esub>" + using assms + by (metis A0 Pring_is_cring Pring_zero cring.cring_simprules(27) is_cring) + have 1: " \ P \\<^bsub>S\<^esub> \ (indexed_const \) = \\<^bsub>S\<^esub>" + proof- + have "\ P \ carrier S" + using assms(2) A0 + by blast + then show ?thesis + using assms(4) Pring_zero[of I] + by (metis assms(1) cring.cring_simprules(27)) + qed + then show ?thesis using 0 1 + using assms(4) by presburger + qed + show "\n Q. (\Q. Q \ Pring_set R I \ card (monomials_of R Q) \ n \ \ (P \\<^bsub>Pring R I\<^esub> Q) = \ P \\<^bsub>S\<^esub> \ Q) \ + Q \ Pring_set R I \ card (monomials_of R Q) = Suc n \ \ (P \\<^bsub>Pring R I\<^esub> Q) = \ P \\<^bsub>S\<^esub> \ Q" + proof- fix n fix Q + assume IH: " (\Q. Q \ Pring_set R I \ card (monomials_of R Q) \ n \ \ (P \\<^bsub>Pring R I\<^esub> Q) = \ P \\<^bsub>S\<^esub> \ Q)" + assume A: "Q \ Pring_set R I \ card (monomials_of R Q) = Suc n" + show "\ (P \\<^bsub>Pring R I\<^esub> Q) = \ P \\<^bsub>S\<^esub> \ Q" + proof- + obtain m M where m_M_def: "monomials_of R Q = insert m M \ m \ M" + using A + by (metis card_0_eq ex_in_conv finite.emptyI mk_disjoint_insert nat.distinct(1)) + have "Q = (restrict_poly_to_monom_set R Q M) \\<^bsub>Pring R I\<^esub> (poly_scalar_mult R (Q m) (mset_to_IP R m))" + by (metis A Pring_add m_M_def remove_monom_eq remove_monom_restrict_poly_to_monom_set) + obtain Q' where Q'_def: "Q' = (restrict_poly_to_monom_set R Q M)" + by simp + have Q'_fact: " Q' \ Pring_set R I \ card (monomials_of R Q') \ n" + proof- + have 0: "Q' \ Pring_set R I" + using Q'_def A restriction_closed + by blast + have "monomials_of R Q' = M" + using A m_M_def Q'_def + by (metis restrict_poly_to_monom_set_monoms subset_insertI) + then have "card (monomials_of R Q') = n" using A m_M_def + by (metis "0" card_insert_disjoint diff_Suc_1 monomials_finite) + then show ?thesis + by (simp add: "0") + qed + have 0:"(P \\<^bsub>Pring R I\<^esub> Q) = (P \\<^bsub>Pring R I\<^esub> (Q' \\<^bsub>Pring R I\<^esub> (poly_scalar_mult R (Q m) (mset_to_IP R m))))" + using Q'_def \Q = restrict_poly_to_monom_set R Q M \\<^bsub>Pring R I\<^esub> Mt (Q m) m\ by presburger + have 1: "(P \\<^bsub>Pring R I\<^esub> Q) = (P \\<^bsub>Pring R I\<^esub> Q') \\<^bsub>Pring R I\<^esub> (P \\<^bsub>Pring R I\<^esub> (poly_scalar_mult R (Q m) (mset_to_IP R m)))" + proof- + have 10: "P \ carrier (Pring R I)" + by (simp add: A0) + have 11: "Q \ carrier (Pring R I)" + by (simp add: A Pring_car) + have 12: "Q' \ carrier (Pring R I)" + using A Pring_car Q'_def restriction_closed by blast + have 13: "(poly_scalar_mult R (Q m) (mset_to_IP R m)) \ carrier (Pring R I)" + by (metis A Pring_car Pring_carrier_coeff' insert_iff m_M_def mset_to_IP_closed' poly_scalar_mult_closed) + then show ?thesis + using 0 10 11 12 13 + by (metis Pring_mult_rdistr) + qed + then have "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ (P \\<^bsub>Pring R I\<^esub> Q')) \\<^bsub>S\<^esub> \ (P \\<^bsub>Pring R I\<^esub> (poly_scalar_mult R (Q m) (mset_to_IP R m)))" + by (metis A A0 Pring_car Pring_mult_closed Pring_cfs_closed Q'_def assms(5) insert_iff + local.ring_axioms m_M_def poly_scalar_mult_closed ring.mset_to_IP_closed' + ring.restriction_closed) + then have "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> (\ Q') \\<^bsub>S\<^esub> \ (P \\<^bsub>Pring R I\<^esub> (poly_scalar_mult R (Q m) (mset_to_IP R m)))" + using IH[of Q'] Q'_fact + by presburger + then have 2: "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> (\ Q') \\<^bsub>S\<^esub> \ (poly_scalar_mult R (Q m) (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)))" + by (metis A A0 Pring_car Pring_mult Pring_cfs_closed insert_iff m_M_def mset_to_IP_closed' times_poly_scalar_mult) + have 3: "\k. k \ carrier R \ set_mset m \ I \ \ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m))) = + \ (indexed_const k) \\<^bsub>S\<^esub> (\ P) \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + proof(induction m) + case empty + assume A: "set_mset {#} \ I" + then have E0: "(P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#}) = P" + by (metis A0 Pring_mult_one Pring_one one_mset_to_IP) + have E1: "k \ carrier R" + using A Pring_cfs_closed empty.prems(1) + by linarith + have E2: "\ (mset_to_IP R {#}) = \\<^bsub>S\<^esub>" + by (metis Pring_one assms(3) one_mset_to_IP) + have E3: " \ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#})) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ (P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#})" + using E1 E2 E0 assms(7)[of "k" "P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#}"] A0 + by (simp add: A0 E1) + have E4: " \ (P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#}) = (\ P) " + using E0 + by auto + have E5: " \ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R {#})) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ P" + using E0 E3 by presburger + have E6: " \ (indexed_const k) \\<^bsub>S\<^esub> \ P = \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> \\<^bsub>S\<^esub>" + proof- + have 0: "\ P \ carrier S" + using assms A0 + by blast + have "indexed_const k \ carrier (Pring R I)" + using assms(2) E1 Pring_car indexed_pset.indexed_const + by blast + then have 1: "\ (indexed_const k) \ carrier S" + using assms(2) + by blast + show ?thesis + using assms(1) 0 1 + by (metis cring.cring_simprules(12) cring.cring_simprules(14) + cring.cring_simprules(5) cring.cring_simprules(6)) + qed + then show ?case + using E0 E2 E3 by presburger + next + case (add x m) + assume AA: "\k. k \ carrier R \ set_mset m \ I \ \ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R m)) = \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + assume P0: "set_mset (add_mset x m) \ I" + then have IA: "set_mset m \I" + by (metis insert_subset set_mset_add_mset_insert) + have IH: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R m)) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + using AA IA add.prems(1) + by blast + then have x_mem: "x \ I" + by (meson P0 subsetD union_single_eq_member) + have 0: " mset_to_IP R (add_mset x m) = mset_to_IP R m \ x" + by (simp add: monom_add_mset) + then have 1: "P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m) = P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m \ x)" + by simp + have 2: "P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m) =( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)) \ x" + proof- + have "mset_to_IP R (add_mset x m) = (mset_to_IP R m) \ x" + using "0" by blast + then have "P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m) = P \\<^bsub>Pring R I\<^esub> ((mset_to_IP R m) \ x)" + by simp + then have "P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m) = P \\<^bsub>Pring R I\<^esub> ((mset_to_IP R m) \\<^bsub>Pring R I\<^esub> mset_to_IP R {#x#})" + by (metis IA Pring_mult mset_to_IP_closed poly_index_mult x_mem) + then have "P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m) = (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)) \\<^bsub>Pring R I\<^esub> mset_to_IP R {#x#}" + by (metis A0 IA Pring_car Pring_mult_assoc Pring_one Pring_one_closed + indexed_pset.indexed_pmult monom_add_mset mset_to_IP_closed one_mset_to_IP x_mem) + then show ?thesis + by (metis A0 IA Pring_car Pring_mult Pring_mult_closed mset_to_IP_closed poly_index_mult x_mem) + qed + have 3: "poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m)) = + poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m) \ x)" + using "2" by presburger + have 4: "poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m)) = + poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)) \ x" + proof- + have "poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m)) = + poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m) \\<^bsub>Pring R I\<^esub> mset_to_IP R {#x#})" + by (metis "2" A0 IA Pring_car Pring_mult Pring_mult_closed mset_to_IP_closed poly_index_mult x_mem) + have "poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m)) = + (poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> mset_to_IP R m)) \\<^bsub>Pring R I\<^esub> mset_to_IP R {#x#}" + by (metis A0 IA Pring_car Pring_mult Pring_mult_closed Pring_one Pring_one_closed + \poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m)) = poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R m \\<^bsub>Pring R I\<^esub> mset_to_IP R {#x#})\ + add.prems(1) indexed_pset.indexed_pmult monom_add_mset mset_to_IP_closed + one_mset_to_IP poly_scalar_mult_times x_mem) + then show ?thesis + by (metis A0 IA Pring_car Pring_mult Pring_mult_closed add.prems(1) + cring.poly_scalar_mult_closed is_cring mset_to_IP_closed + poly_index_mult x_mem) + qed + have 5: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m)) \ x)" + using 4 by metis + have 6: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (poly_scalar_mult R k ( P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m))) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#})" + using assms + by (metis "5" A0 IA Pring_car Pring_mult_closed add.prems(1) + cring.poly_scalar_mult_closed is_cring mset_to_IP_closed x_mem) + have 7: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> \ (mset_to_IP R m) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#})" + using assms 6 IH + by presburger + have 8: " \ (mset_to_IP R m) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#}) = \ (mset_to_IP R (add_mset x m))" + proof- + have "\ (mset_to_IP R m) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#}) = + \ ((mset_to_IP R m) \\<^bsub>Pring R I\<^esub> (mset_to_IP R {#x#}))" + by (metis IA Pring_car Pring_mult assms(6) mset_to_IP_closed poly_index_mult x_mem) + then show ?thesis + by (metis "0" IA Pring_car assms(6) mset_to_IP_closed x_mem) + qed + have 9: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> ( \ (mset_to_IP R m) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#}))" + proof- + have 0: "\ (indexed_const k) \ carrier S" + proof- + have "(indexed_const k) \ carrier (Pring R I)" + using A Pring_car Pring_cfs_closed indexed_pset.indexed_const add.prems(1) + by blast + then show ?thesis + using assms(2) + by blast + qed + have 1: "\ P \ carrier S" + using A0 assms(2) + by blast + have 2: "\ (mset_to_IP R m) \ carrier S" + proof- + have "(mset_to_IP R m) \ carrier (Pring R I)" + using IA Pring_car mset_to_IP_closed by blast + then show ?thesis using assms(2) + by blast + qed + have 3: " \ (mset_to_IP R {#x#}) \ carrier S" + proof- + have "(mset_to_IP R {#x#}) \ carrier (Pring R I)" + by (metis Pring_car Pring_one Pring_one_closed indexed_pset.indexed_pmult + monom_add_mset one_mset_to_IP x_mem) + then show ?thesis + using assms(2) + by blast + qed + have "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (indexed_const k) \\<^bsub>S\<^esub> (\ P \\<^bsub>S\<^esub> \ (mset_to_IP R m) \\<^bsub>S\<^esub> \ (mset_to_IP R {#x#}))" + using 1 2 3 7 + by (metis "0" \\ (mset_to_IP R m) \ carrier S\ assms(1) + cring.cring_simprules(11) cring.cring_simprules(5)) + then show ?thesis + by (metis "0" "1" "2" "3" "8" Pring_mult assms(1) cring.cring_simprules(11) + cring.cring_simprules(5)) + qed + have 10: "\ (poly_scalar_mult R k (P \\<^bsub>Pring R I\<^esub> mset_to_IP R (add_mset x m))) = + \ (indexed_const k) \\<^bsub>S\<^esub> \ P \\<^bsub>S\<^esub> \ (mset_to_IP R (add_mset x m))" + using 8 9 + by presburger + then show ?case + by blast + qed + have 4: "\ (poly_scalar_mult R (Q m) (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R m))) = + \ (indexed_const (Q m)) \\<^bsub>S\<^esub> (\ P) \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + using 3 + by (metis A Pring_mult insert_iff local.ring_axioms m_M_def mset_to_IP_indices + ring.Pring_cfs_closed) + have 5: "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> (\ Q') \\<^bsub>S\<^esub> \ (indexed_const (Q m)) \\<^bsub>S\<^esub> (\ P) \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + using 2 4 + by presburger + have " \ (indexed_const (Q m)) \\<^bsub>S\<^esub> (\ P) \\<^bsub>S\<^esub> \ (mset_to_IP R m) = + (\ P) \\<^bsub>S\<^esub> \ (indexed_const (Q m)) \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + proof- + have 0: "(\ P) \ carrier S" + using A0 assms(2) + by blast + have 1: "\ (indexed_const (Q m)) \ carrier S" + proof- + have "indexed_const (Q m ) \ carrier (Pring R I)" + using A Pring_car Pring_cfs_closed indexed_pset.indexed_const by blast + then show ?thesis + using assms(2) + by blast + qed + have 2: "\ (mset_to_IP R m) \ carrier S" + proof- + have "(mset_to_IP R m) \ carrier (Pring R I)" + by (metis A Pring_car insert_iff m_M_def mset_to_IP_closed') + then show ?thesis using assms(2) + by blast + qed + show ?thesis using assms(1) 0 1 2 + by (metis cring.cring_simprules(14)) + qed + then have 6: "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> (\ Q') \\<^bsub>S\<^esub> (\ P) \\<^bsub>S\<^esub> \ (indexed_const (Q m)) \\<^bsub>S\<^esub> \ (mset_to_IP R m)" + using 5 + by presburger + then have 7: "\ (P \\<^bsub>Pring R I\<^esub> Q) = (\ P) \\<^bsub>S\<^esub> ((\ Q') \\<^bsub>S\<^esub> \ (indexed_const (Q m)) \\<^bsub>S\<^esub> \ (mset_to_IP R m))" + proof- + have 0: "(\ P) \ carrier S" + using A0 assms(2) + by blast + have 1: "\ (indexed_const (Q m)) \ carrier S" + proof- + have "indexed_const (Q m ) \ carrier (Pring R I)" + using A Pring_car Pring_cfs_closed indexed_pset.indexed_const by blast + then show ?thesis + using assms(2) + by blast + qed + have 2: "\ (mset_to_IP R m) \ carrier S" + proof- + have "(mset_to_IP R m) \ carrier (Pring R I)" + by (metis A Pring_car insert_iff m_M_def mset_to_IP_closed') + then show ?thesis using assms(2) + by blast + qed + have 3: "\ Q' \ carrier S" + proof- + have "Q' \ carrier (Pring R I)" + using Q'_fact Pring_car + by blast + then show ?thesis + using assms(2) + by blast + qed + show ?thesis using 0 1 2 3 6 + by (metis assms(1) cring.cring_simprules(11) cring.cring_simprules(25) cring.cring_simprules(5)) + qed + then show ?thesis + by (metis A Pring_car Pring_cfs_closed Q'_def Q'_fact + \Q = restrict_poly_to_monom_set R Q M \\<^bsub>Pring R I\<^esub> Mt (Q m) m\ assms(5) assms(7) + cring.poly_scalar_mult_closed insert_iff is_cring m_M_def mset_to_IP_closed') + qed + qed + show "Q \ Pring_set R I " + using A1 Pring_car + by blast + qed + qed + show "\x y. x \ carrier (Pring R I) \ y \ carrier (Pring R I) \ \ (x \\<^bsub>Pring R I\<^esub> y) = \ x \\<^bsub>S\<^esub> \ y" + using assms(5) + by blast + show "\ \\<^bsub>Pring R I\<^esub> = \\<^bsub>S\<^esub>" + by (simp add: assms(3)) +qed + +lemma(in cring) indexed_const_Pring_mult: + assumes "k \ carrier R" + assumes "P \ carrier (Pring R I)" + shows "(indexed_const k \\<^bsub>Pring R I\<^esub> P) m = k \\<^bsub>R\<^esub> (P m)" + "(P \\<^bsub>Pring R I\<^esub> indexed_const k) m = k \\<^bsub>R\<^esub> (P m)" + apply (metis Pring_car Pring_mult assms(1) assms(2) poly_scalar_mult_def poly_scalar_mult_eq) + by (metis Pring_car Pring_carrier_coeff' Pring_mult assms(1) assms(2) indexed_const_P_multE m_comm) + +lemma(in cring) ring_hom_to_IP_ring_hom_is_hom: + assumes "cring S" + assumes "ring_hom_ring R S \" + shows "ring_hom_ring (Pring R I) (Pring S I) (ring_hom_to_IP_ring_hom \)" +proof(rule Pring_morphism) + show 0: "cring (Pring S I)" + by (simp add: assms(1) cring.axioms(1) ring.Pring_is_cring) + show 1: "ring_hom_to_IP_ring_hom \ \ carrier (Pring R I) \ carrier (Pring S I)" + by (meson Pi_I assms(1) assms(2) ring_hom_to_IP_ring_hom_closed) + show 2: "ring_hom_to_IP_ring_hom \ \\<^bsub>Pring R I\<^esub> = \\<^bsub>Pring S I\<^esub>" + by (simp add: assms(1) assms(2) ring_hom_to_IP_ring_hom_one) + show 3: "ring_hom_to_IP_ring_hom \ \\<^bsub>Pring R I\<^esub> = \\<^bsub>Pring S I\<^esub>" + proof- + have "\m. \ (\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>S\<^esub>" + using assms + by (metis Pring_carrier_coeff' Pring_zero Pring_zero_closed cring.axioms(1) + cring.cring_simprules(26) indexed_const_Pring_mult(2) is_cring ring.Pring_is_cring + ring_hom_ring.homh ring_hom_zero zero_closed) + then have "\m. \ (\\<^bsub>Pring R I\<^esub> m) = \\<^bsub>Pring S I\<^esub> m" + by (metis assms(1) cring.axioms(1) ring.Pring_zero ring.indexed_const_def) + then show ?thesis unfolding ring_hom_to_IP_ring_hom_def + by blast + qed + show 4: " \P Q. P \ carrier (Pring R I) \ + Q \ carrier (Pring R I) \ ring_hom_to_IP_ring_hom \ (P \\<^bsub>Pring R I\<^esub> Q) = ring_hom_to_IP_ring_hom \ P \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q" + using assms(1) assms(2) ring_hom_to_IP_ring_hom_add by blast + show 5: " \i P. i \ I \ + P \ carrier (Pring R I) \ + ring_hom_to_IP_ring_hom \ (P \ i) = ring_hom_to_IP_ring_hom \ P \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ (mset_to_IP R {#i#})" + proof- + fix i P + assume i0: "i \ I" + assume P0: "P \ carrier (Pring R I)" + show "ring_hom_to_IP_ring_hom \ (P \ i) = ring_hom_to_IP_ring_hom \ P \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ (mset_to_IP R {#i#})" + unfolding ring_hom_to_IP_ring_hom_def + proof + fix m + show " \ ((P \ i) m) = ((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m" + proof(cases "i \# m") + case True + have "(\m. \ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}" + proof + fix m + show "\ (mset_to_IP R {#i#} m) = mset_to_IP S {#i#} m" + apply(cases "{#i#} = m") + apply (metis (mono_tags, lifting) assms(1) assms(2) cring.axioms(1) local.ring_axioms + mset_to_IP_def one_mset_to_IP ring.Pring_one ring.Pring_one ring.one_mset_to_IP + ring_hom_to_IP_ring_hom_def ring_hom_to_IP_ring_hom_one) + proof- + assume "{#i#} \ m" + then have LHS: "(mset_to_IP R {#i#} m) = \\<^bsub>R\<^esub>" + by (metis mset_to_IP_simp') + have RHS :"mset_to_IP S {#i#} m = \\<^bsub>S\<^esub>" + by (metis \{#i#} \ m\ mset_to_IP_def) + have "\ \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" + using assms(1) assms(2) cring.axioms(1) local.ring_axioms + ring_hom_ring.homh ring_hom_zero by blast + then show ?thesis + using LHS RHS + by presburger + qed + qed + then have RHS: "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + ((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> mset_to_IP S {#i#}) m" + by presburger + then have RHS': "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + (ring.indexed_pmult S (\m. \ (P m)) i) m" + proof- + have 0: "(\m. \ (P m)) \ Pring_set S I" + using ring_hom_to_IP_ring_hom_closed[of S \ P I] ring.Pring_car[of S I] assms + unfolding ring_hom_to_IP_ring_hom_def + using P0 cring.axioms(1) by blast + show ?thesis using ring.poly_index_mult[of S "(\m. \ (P m))" I i] + by (metis "0" \(\m. \ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}\ + assms(1) cring.axioms(1) i0 ring.Pring_mult) + qed + then have RHS'': "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + (\m. \ (P m)) (m - {#i#})" using ring.indexed_pmult_def[of S "(\m. \ (P m))" i] True + by (metis assms(1) cring.axioms(1)) + then show ?thesis + by (metis True indexed_pmult_def) + next + case False + have LHS: "((P \ i) m) = \\<^bsub>R\<^esub>" + using False + by (meson indexed_pmult_def) + have "(\m. \ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}" + proof + fix m + show "\ (mset_to_IP R {#i#} m) = mset_to_IP S {#i#} m" + apply(cases "{#i#} = m") + apply (metis (mono_tags, lifting) assms(1) assms(2) cring.axioms(1) local.ring_axioms + mset_to_IP_def one_mset_to_IP ring.Pring_one ring.Pring_one ring.one_mset_to_IP + ring_hom_to_IP_ring_hom_def ring_hom_to_IP_ring_hom_one) + proof- + assume "{#i#} \ m" + then have LHS: "(mset_to_IP R {#i#} m) = \\<^bsub>R\<^esub>" + by (metis mset_to_IP_simp') + have RHS :"mset_to_IP S {#i#} m = \\<^bsub>S\<^esub>" + by (metis \{#i#} \ m\ mset_to_IP_def) + have "\ \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" + using assms(1) assms(2) cring.axioms(1) local.ring_axioms + ring_hom_ring.homh ring_hom_zero by blast + then show ?thesis + using LHS RHS + by presburger + qed + qed + then have RHS: "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + ((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> mset_to_IP S {#i#}) m" + by presburger + then have RHS': "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + (ring.indexed_pmult S (\m. \ (P m)) i) m" + proof- + have 0: "(\m. \ (P m)) \ Pring_set S I" + using ring_hom_to_IP_ring_hom_closed[of S \ P I] ring.Pring_car[of S I] assms + unfolding ring_hom_to_IP_ring_hom_def + using P0 cring.axioms(1) by blast + show ?thesis using ring.poly_index_mult[of S "(\m. \ (P m))" I i] + by (metis "0" \(\m. \ (mset_to_IP R {#i#} m)) = mset_to_IP S {#i#}\ + assms(1) cring.axioms(1) i0 ring.Pring_mult) + qed + then have RHS'': "((\m. \ (P m)) \\<^bsub>Pring S I\<^esub> (\m. \ (mset_to_IP R {#i#} m))) m = + \\<^bsub>S\<^esub>" + using False + by (metis assms(1) cring.axioms(1) ring.indexed_pmult_def) + then show ?thesis + using LHS False assms ring_hom_zero[of \ R S] + by (metis cring.axioms(1) local.ring_axioms ring_hom_ring.homh) + qed + qed + qed + show "\k Q. k \ carrier R \ + Q \ carrier (Pring R I) \ + ring_hom_to_IP_ring_hom \ (poly_scalar_mult R k Q) = ring_hom_to_IP_ring_hom \ (indexed_const k) \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q" + proof- + fix k Q + assume A0: "k \ carrier R" + assume A1: " Q \ carrier (Pring R I)" + show "ring_hom_to_IP_ring_hom \ (poly_scalar_mult R k Q) = + ring_hom_to_IP_ring_hom \ (indexed_const k) \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q" + proof + fix x + show "ring_hom_to_IP_ring_hom \ (poly_scalar_mult R k Q) x = (ring_hom_to_IP_ring_hom \ (indexed_const k) \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q) x" + proof- + have LHS: "ring_hom_to_IP_ring_hom \ (poly_scalar_mult R k Q) x = \ (k \\<^bsub>R\<^esub> Q x)" + by (metis poly_scalar_mult_def ring_hom_to_IP_ring_hom_def) + then have LHS': "ring_hom_to_IP_ring_hom \ (poly_scalar_mult R k Q) x = (\ k) \\<^bsub>S\<^esub> \ (Q x)" + proof- + have "Q x \ carrier R" + using A1 Pring_car local.ring_axioms ring.Pring_cfs_closed by blast + then show ?thesis + using LHS assms ring_hom_mult[of \ R S k "Q x"] + by (metis A0 ring_hom_ring.homh) + qed + have RHS: "(ring_hom_to_IP_ring_hom \ (indexed_const k) \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q) x = + (\ k) \\<^bsub>S\<^esub> (\ (Q x))" + proof- + have 0: "ring_hom_to_IP_ring_hom \ (indexed_const k)= ring.indexed_const S (\ k)" + by (simp add: A0 assms(1) assms(2) ring_hom_to_IP_ring_hom_constant) + have 1: "(\ k) \ carrier S" + by (meson A0 assms(2) ring_hom_closed ring_hom_ring.homh) + have 2: "ring_hom_to_IP_ring_hom \ Q \ carrier (Pring S I)" + using A1 assms ring_hom_to_IP_ring_hom_closed + by blast + have 3: "(ring.indexed_const S (\ k) \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ Q) x = (\ k) \\<^bsub>S\<^esub> (\ (Q x))" + using assms(1) 1 2 + cring.indexed_const_Pring_mult(1)[of S "\ k" "ring_hom_to_IP_ring_hom \ Q" I x] + ring_hom_to_IP_ring_hom_def[of \ Q x] + by presburger + then show ?thesis + by (metis "0") + qed + then show ?thesis + using LHS' + by metis + qed + qed + qed +qed + +lemma ring_hom_to_IP_ring_hom_smult: + assumes "cring S" + assumes "ring_hom_ring R S \" + assumes "P \ carrier (Pring R I)" + assumes "a \ carrier R" + shows "ring_hom_to_IP_ring_hom \ (a \\<^bsub>Pring R I\<^esub>P) = + \ a \\<^bsub>Pring S I\<^esub> (ring_hom_to_IP_ring_hom \ P)" +proof fix m + have 0: "\ ((a \\<^bsub>Pring R I\<^esub> P) m) = \ (a \ (P m))" + using assms + by (metis Pring_smult_cfs) + hence 1: "\ ((a \\<^bsub>Pring R I\<^esub> P) m) = \ a \\<^bsub>S\<^esub> \ (P m)" + using assms ring_hom_mult[of \ R S] + by (metis Pring_carrier_coeff' ring_hom_ring.homh) + have 2: "(\ a \\<^bsub>Pring S I\<^esub> (\m. \ (P m))) m = \ a \\<^bsub>S\<^esub> \ (P m)" + using assms ring.Pring_smult[of S I] + unfolding poly_scalar_mult_def cring_def + by presburger + show "ring_hom_to_IP_ring_hom \ (a \\<^bsub>Pring R I\<^esub> P) m = + (\ a \\<^bsub>Pring S I\<^esub> ring_hom_to_IP_ring_hom \ P) m" + unfolding ring_hom_to_IP_ring_hom_def using assms 1 2 + by presburger +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\A Universal Property for Indexed Polynomial Rings\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma Pring_universal_prop_0: + assumes a_cring: "cring S" + assumes index_map: "closed_fun S g" + assumes ring_hom: "ring_hom_ring R S \" + assumes "\ = (total_eval S g) \ (ring_hom_to_IP_ring_hom \)" + shows "(ring_hom_ring (Pring R I) S \)" + "(\i \ I. \ (mset_to_IP R {#i#}) = g i)" + "(\a \ carrier R. \ (indexed_const a) = \ a)" + "\ \. (ring_hom_ring (Pring R I) S \) \ + (\i \ I. \ (mset_to_IP R {#i#}) = g i) \ + (\a \ carrier R. \ (indexed_const a) = \ a) \ + (\x \ carrier (Pring R I). \ x = \ x)" +proof- + have 0: " (ring_hom_to_IP_ring_hom \) \ ring_hom (Pring R I) (Pring S I)" + using a_cring ring_hom ring_hom_ring.homh ring_hom_to_IP_ring_hom_is_hom + by blast + have 1: "(total_eval S g) \ ring_hom (Pring S I) S " + using a_cring cring.total_eval_ring_hom index_map ring_hom_ring.homh + by blast + show P0: "ring_hom_ring (Pring R I) S \ " + using ring_hom_trans 0 1 Pring_is_ring a_cring assms(4) cring.axioms(1) ring_hom_ringI2 + by blast + show P1: "\i\I. \ (mset_to_IP R {#i#}) = g i" + proof + fix i assume Pi: "i \ I" + show "\ (mset_to_IP R {#i#}) = g i" + proof- + have 0: "\ (mset_to_IP R {#i#}) = (total_eval S g) ( (ring_hom_to_IP_ring_hom \) (mset_to_IP R {#i#}))" + by (simp add: assms(4)) + have "( (ring_hom_to_IP_ring_hom \) (mset_to_IP R {#i#})) = (mset_to_IP S {#i#})" + by (simp add: a_cring ring_hom ring_hom_to_IP_ring_hom_monom) + then have "\ (mset_to_IP R {#i#}) = (total_eval S g) (mset_to_IP S {#i#})" + by (simp add: 0) + then show ?thesis + by (simp add: a_cring cring.total_eval_var index_map) + qed + qed + show P2: "\a\carrier R. \ (indexed_const a) = \ a" + proof + fix a assume A: "a \ carrier R" + show "\ (indexed_const a) = \ a" + proof- + have 0: "ring_hom_to_IP_ring_hom \ (indexed_const a) = ring.indexed_const S (\ a)" + by (simp add: A a_cring ring_hom ring_hom_to_IP_ring_hom_constant) + have 1: "total_eval S g (ring.indexed_const S (\ a)) = \ a" + by (meson A a_cring cring.total_eval_const ring_hom ring_hom_closed ring_hom_ring.homh) + show ?thesis + using assms 0 1 + by (simp add: "0" index_map) + qed + qed + show "\ \. (ring_hom_ring (Pring R I) S \) \ + (\i \ I. \ (mset_to_IP R {#i#}) = g i) \ + (\a \ carrier R. \ (indexed_const a) = \ a) \ + (\x \ carrier (Pring R I). \ x = \ x)" + proof + fix \ + show "ring_hom_ring (Pring R I) S \ \ (\i\I. \ (mset_to_IP R {#i#}) = g i) \ (\a\carrier R. \ (indexed_const a) = \ a) \ + (\x\carrier (Pring R I). \ x = \ x)" + proof + assume A: "(ring_hom_ring (Pring R I) S \) \ + (\i \ I. \ (mset_to_IP R {#i#}) = g i) \ + (\a \ carrier R. \ (indexed_const a) = \ a)" + show "(\x \ carrier (Pring R I). \ x = \ x)" + proof + fix x assume B: "x \ carrier (Pring R I)" + show "\ x = \ x" + apply(rule indexed_pset.induct[of x I "carrier R"]) + using B Pring_car apply blast + apply (metis A P2) + proof- + show "\P Q. P \ Pring_set R I \ \ P = \ P \ Q \ Pring_set R I \ \ Q = \ Q \ \ (P \ Q) = \ (P \ Q)" + proof- + fix P Q + assume A0: "P \ Pring_set R I " "\ P = \ P" " Q \ Pring_set R I" " \ Q = \ Q " + show "\ (P \ Q) = \ (P \ Q)" + using A A0 ring_hom_add[of \ "Pring R I" S P Q] ring_hom_add[of \ "Pring R I" S P Q] P0 + by (metis Pring_add Pring_car ring_hom_ring.homh) + qed + show "\P i. P \ Pring_set R I \ \ P = \ P \ i \ I \ \ (P \ i) = \ (P \ i)" + proof- + fix P i + assume A0: " P \ Pring_set R I" " \ P = \ P" "i \ I" + show "\ (P \ i) = \ (P \ i)" + proof- + have "\ (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R {#i#})) = \ (P \\<^bsub>Pring R I\<^esub> (mset_to_IP R {#i#}))" + using A A0 ring_hom_mult[of \ "Pring R I" S P "(mset_to_IP R {#i#})"] + ring_hom_mult[of \ "Pring R I" S P "(mset_to_IP R {#i#})"] P0 + by (metis P1 Pring_car Pring_one Pring_one_closed indexed_pset.indexed_pmult + monom_add_mset one_mset_to_IP ring_hom_ring.homh) + then show ?thesis + by (metis A0(1) A0(3) Pring_mult poly_index_mult) + qed + qed + qed + qed + qed + qed +qed + +end + +definition close_fun :: "'c set \ ('e, 'f) ring_scheme \ ('c \ 'e) \ ('c \ 'e)" where +"close_fun I S g = (\i. (if i \ I then g i else \\<^bsub>S\<^esub>))" + +context cring +begin + +lemma close_funE: + assumes "cring S" + assumes "g \ I \ carrier S" + shows "closed_fun S (close_fun I S g)" + apply(rule cring.closed_funI) + apply (simp add: assms(1)) + by (metis close_fun_def PiE assms(1) assms(2) cring.cring_simprules(2)) + +end + +definition indexed_poly_induced_morphism :: + "'c set \ ('e, 'f) ring_scheme \ ('a, 'e) ring_hom \ ('c \ 'e) \ (('a,'c) mvar_poly, 'e) ring_hom" where +"indexed_poly_induced_morphism I S \ g = (total_eval S (close_fun I S g)) \ (ring_hom_to_IP_ring_hom \)" + +context cring +begin + +lemma Pring_universal_prop: + assumes a_cring: "cring S" + assumes index_map: "g \ I \ carrier S" + assumes ring_hom: "ring_hom_ring R S \" + assumes "\ = indexed_poly_induced_morphism I S \ g" + shows "(ring_hom_ring (Pring R I) S \)" + "(\i \ I. \ (mset_to_IP R {#i#}) = g i)" + "(\a \ carrier R. \ (indexed_const a) = \ a)" + "\ \. (ring_hom_ring (Pring R I) S \) \ + (\i \ I. \ (mset_to_IP R {#i#}) = g i) \ + (\a \ carrier R. \ (indexed_const a) = \ a) \ + (\x \ carrier (Pring R I). \ x = \ x)" +proof- + obtain g' where g'_def: "g' = (close_fun I S g)" + by simp + have 0: "closed_fun S g'" + using close_funE a_cring g'_def index_map + by blast + show "(ring_hom_ring (Pring R I) S \)" using assms 0 + using cring.Pring_universal_prop_0(1) indexed_poly_induced_morphism_def g'_def is_cring + by blast + show "(\i \ I. \ (mset_to_IP R {#i#}) = g i)" + proof- + have "(\i \ I. \ (mset_to_IP R {#i#}) = g' i)" + apply(intro Pring_universal_prop_0[of S _ \] assms) + unfolding assms indexed_poly_induced_morphism_def g'_def + using assms 0 g'_def apply fastforce + by auto + thus ?thesis unfolding g'_def using assms + by (simp add: close_fun_def) + qed + show "\a\carrier R. \ (indexed_const a) = \ a" + using 0 indexed_poly_induced_morphism_def Pring_universal_prop_0(3) a_cring assms(4) g'_def ring_hom + by blast + show "\\. ring_hom_ring (Pring R I) S \ \ (\i\I. \ (mset_to_IP R {#i#}) = g i) \ (\a\carrier R. \ (indexed_const a) = \ a) \ + (\x\carrier (Pring R I). \ x = \ x)" + proof- + have "\\. ring_hom_ring (Pring R I) S \ \ (\i\I. \ (mset_to_IP R {#i#}) = g' i) \ (\a\carrier R. \ (indexed_const a) = \ a) \ + (\x\carrier (Pring R I). \ x = \ x)" + using assms 0 Pring_universal_prop_0(4) g'_def + unfolding indexed_poly_induced_morphism_def + by blast + then show ?thesis + using g'_def + unfolding close_fun_def + by meson + qed +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Mapping Mulitvariate Polynomials over a Single Variable to Univariate Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\Constructor for multisets which have one distinct element\ + +definition nat_to_mset :: "'c \ nat \ 'c monomial" where +"nat_to_mset i n = Abs_multiset (\j. if (j = i) then n else 0)" + +lemma nat_to_msetE: "count (nat_to_mset i n) i = n" + unfolding nat_to_mset_def by simp + +lemma nat_to_msetE': + assumes "j \ i" + shows "count (nat_to_mset i n) j = 0" + unfolding nat_to_mset_def using assms by simp + +lemma nat_to_mset_add: "nat_to_mset i (n + m) = (nat_to_mset i n) + (nat_to_mset i m)" + apply(rule multiset_eqI) + by (metis add.right_neutral nat_to_msetE nat_to_msetE' count_union) + +lemma nat_to_mset_inj: + assumes "n \ m" + shows "(nat_to_mset i n) \ (nat_to_mset i m)" + using assms + by (metis nat_to_msetE) + +lemma nat_to_mset_zero: "nat_to_mset i 0 = {#}" + by (metis add.right_neutral add_cancel_right_right nat_to_mset_add) + +lemma nat_to_mset_Suc: "nat_to_mset i (Suc n) = add_mset i (nat_to_mset i n)" + using nat_to_mset_add[of i n 1] + by (simp add: multiset_eqI nat_to_msetE nat_to_msetE') + +lemma nat_to_mset_Pring_singleton: + assumes "cring R" + assumes "P \ carrier (Pring R {i})" + assumes "m \ monomials_of R P" + shows "m = nat_to_mset i (count m i)" +proof- + have "\ j. count m j = count (nat_to_mset i (count m i)) j" + proof- + fix j + show "count m j = count (nat_to_mset i (count m i)) j" + apply(cases "j = i") + apply (simp add: nat_to_msetE; fail) + proof- + assume A: "j \i" + have P0: "set_mset m \ {i}" + using assms + by (metis cring.axioms(1) ring.Pring_car ring.mset_to_IP_indices) + then have "count m j = 0" + using A assms + by (metis count_inI empty_iff singletonD subset_singletonD) + then show " count m j = count (nat_to_mset i (count m i)) j" + by (simp add: A nat_to_msetE') + qed + qed + then show ?thesis + using multiset_eqI by blast +qed + +definition IP_to_UP :: "'d \ ('e, 'd) mvar_poly \ 'e u_poly" where +"IP_to_UP i P = (\ (n::nat). P (nat_to_mset i n))" + +lemma IP_to_UP_closed: + assumes "cring R" + assumes "P \ carrier (Pring R {i::'c})" + shows "IP_to_UP i P \ carrier (UP R)" +proof- + have "IP_to_UP i P \ up R" + apply(rule mem_upI) + using assms + apply (metis cring_def IP_to_UP_def ring.Pring_carrier_coeff') + unfolding bound_def IP_to_UP_def + apply(rule ccontr) + proof- + assume A: "\n. \m>n. P (nat_to_mset i m) = \\<^bsub>R\<^esub>" + then have 0: "\n. \m> n. (nat_to_mset i m) \ monomials_of R P" + by (meson assms(1) cring_def ring.complement_of_monomials_of) + have "\ finite {m. (nat_to_mset i m) \ monomials_of R P }" + proof + assume "finite {m. nat_to_mset i m \ monomials_of R P}" + then have "\n. \m>n. m \ {m. nat_to_mset i m \ monomials_of R P}" + by (meson finite_nat_set_iff_bounded nat_less_le order.strict_trans) + then have "\n. \m>n. nat_to_mset i m \ monomials_of R P" + by (simp add: monomials_of_def) + then show False using 0 + by blast + qed + then obtain S where S_def: "infinite (S::nat set) \ (\m \ S. (nat_to_mset i m) \ monomials_of R P)" + by blast + have "inj_on (nat_to_mset i) S" + using inj_onI[of S "nat_to_mset i"] + by (meson nat_to_mset_inj) + then have 1: "infinite (nat_to_mset i ` S)" + using S_def finite_imageD + by blast + have 2: "(nat_to_mset i ` S) \ (monomials_of R P)" + using S_def + by blast + then have "infinite (monomials_of R P)" + using S_def "1" finite_subset + by blast + then show False using assms + by (metis Pring_def cring_def partial_object.select_convs(1) ring.monomials_finite) + qed + then show ?thesis + by (metis (no_types, lifting) UP_def partial_object.select_convs(1)) +qed + +lemma IP_to_UP_var: + shows "IP_to_UP i (mset_to_IP R {#i#}) = X_poly R" +proof + have UP: "UP_cring R" + by (simp add: UP_cring_def cring_axioms) + fix x + show "IP_to_UP i (mset_to_IP R {#i#}) x = X_poly R x" + proof(cases "x = 1") + case True + then have RHS: "X_poly R x = \\<^bsub>R\<^esub>" + unfolding X_poly_def using UP UP_ring.cfs_monom[of R] + unfolding UP_ring_def + using local.ring_axioms one_closed by presburger + have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x + = mset_to_IP R ((add_mset i) {#}) (nat_to_mset i x)" + unfolding IP_to_UP_def + by blast + have "(nat_to_mset i x) = {#i#}" + proof- + have "\n. count (nat_to_mset i x) n = count {#i#} n" + by (simp add: True nat_to_msetE nat_to_msetE') + then show ?thesis + using multi_count_eq + by blast + qed + then have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x + = mset_to_IP R ((add_mset i) {#}) {#i#}" + using LHS by presburger + then show ?thesis + unfolding deg_zero_cf_def + by (metis RHS mset_to_IP_simp) + next + case False + then have RHS: "X_poly R x = \\<^bsub>R\<^esub>" + unfolding X_poly_def + using UP UP_ring.cfs_monom[of R] + unfolding UP_ring_def + using local.ring_axioms one_closed by presburger + have LHS: "IP_to_UP i (mset_to_IP R ((add_mset i) {#})) x + = mset_to_IP R ((add_mset i) {#}) (nat_to_mset i x)" + unfolding IP_to_UP_def + by blast + then show ?thesis + unfolding deg_zero_cf_def + by (metis False RHS count_single mset_to_IP_def nat_to_msetE) + qed +qed + +end + +context UP_cring +begin + +lemma IP_to_UP_monom: + shows "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) = ((X_poly R)[^]\<^bsub>UP R\<^esub>n) " +proof + fix x + show "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = (X_poly R [^]\<^bsub>UP R\<^esub> n) x" + proof(cases "x = n") + case True + have RHS: "(X_poly R [^]\<^bsub>UP R\<^esub> n) x = \\<^bsub>R\<^esub>" + unfolding X_poly_def + by (metis P.nat_pow_closed P.nat_pow_eone P_def R.one_closed True UP_cring.X_closed + UP_cring.monom_coeff UP_one_closed UP_r_one deg_one is_UP_cring monom_one monom_rep_X_pow + to_poly_inverse to_poly_mult_simp(2)) + have LHS: "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = \\<^bsub>R\<^esub>" + by (metis R.mset_to_IP_simp True IP_to_UP_def) + then show ?thesis + using RHS by presburger + next + case False + have 0: "\x y::nat. nat_to_mset x = nat_to_mset y \ x = y" + proof- + fix a b::nat assume A: "nat_to_mset a = nat_to_mset b" + then show "a = b" unfolding nat_to_mset_def + by (metis A nat_to_msetE nat_to_msetE' zero_neq_one) + qed + have 1: "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = (if nat_to_mset i x = nat_to_mset i n then \ else \) " + unfolding IP_to_UP_def mset_to_IP_def + by blast + hence 2: "IP_to_UP i (mset_to_IP R (nat_to_mset i n)) x = (if x = n then \ else \) " + using 0 + by (meson False nat_to_mset_inj) + have 3: "(X_poly R [^]\<^bsub>UP R\<^esub> n) x = \" + unfolding X_poly_def using False + by (smt ctrm_degree P.nat_pow_closed P.nat_pow_eone P.r_null P_def R.one_closed + UP_cring.ltrm_of_X UP_cring.ltrm_rep_X_pow UP_cring.X_closed UP_cring.monom_coeff + UP_r_one UP_zero_closed X_mult_cf cfs_closed cfs_monom deg_nzero_nzero is_UP_cring + monom_closed monom_one to_poly_inverse to_poly_mult_simp(2)) + thus ?thesis using 2 1 + using False by presburger + qed +qed + +lemma IP_to_UP_one: + "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> = \\<^bsub>UP R\<^esub>" +proof + fix x + show "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> x = \\<^bsub>UP R\<^esub> x" + proof(cases "x = 0") + case True + have RHS: "\\<^bsub>UP R\<^esub> x = \\<^bsub>R\<^esub>" + using P_def True cfs_one by presburger + have "\\<^bsub>Pring R {i}\<^esub> = (\ m. if m = {#} then \\<^bsub>R\<^esub> else \\<^bsub>R\<^esub>)" + by (metis R.Pring_one R.indexed_const_def) + then have "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> = IP_to_UP i (\ m. if m = {#} then \\<^bsub>R\<^esub> else \\<^bsub>R\<^esub>)" + by presburger + then have LHS: "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> x = \\<^bsub>R\<^esub>" + by (smt True count_empty IP_to_UP_def multi_count_eq nat_to_msetE nat_to_msetE') + then show ?thesis + using RHS by presburger + next + case False + have RHS: "\\<^bsub>UP R\<^esub> x = \\<^bsub>R\<^esub>" + by (smt False UP_def monoid.simps(2)) + show ?thesis + using False count_empty + nat_to_msetE + ring.indexed_const_def + unfolding IP_to_UP_def + by (metis R.Pring_one R.ring_axioms RHS) + qed +qed + +lemma IP_to_UP_zero: + "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> = \\<^bsub>UP R\<^esub>" +proof + fix x + show "IP_to_UP i \\<^bsub>Pring R {i}\<^esub> x = \\<^bsub>UP R\<^esub> x" + unfolding IP_to_UP_def using R.Pring_zero + by (metis P_def R.indexed_zero_def cfs_zero) +qed + +lemma IP_to_UP_add: + assumes " x \ carrier (Pring R {i})" + assumes " y \ carrier (Pring R {i})" + shows " IP_to_UP i (x \\<^bsub>Pring R {i}\<^esub> y) = + IP_to_UP i x \\<^bsub>UP R\<^esub> IP_to_UP i y" +proof + fix n + have LHS: "IP_to_UP i (x \\<^bsub>Pring R {i}\<^esub> y) n = (x \\<^bsub>Pring R {i}\<^esub> y) (nat_to_mset i n)" + by (meson IP_to_UP_def) + then have LHS: "IP_to_UP i (x \\<^bsub>Pring R {i}\<^esub> y) n = x (nat_to_mset i n) \\<^bsub>R\<^esub> y (nat_to_mset i n)" + using assms unfolding IP_to_UP_def + by (metis R.Pring_add R.indexed_padd_def) + have RHS: "(IP_to_UP i x \\<^bsub>UP R\<^esub> IP_to_UP i y) n = + (IP_to_UP i x) n \\<^bsub>R\<^esub> (IP_to_UP i y) n" + using assms UP_ring.cfs_add IP_to_UP_closed + by (simp add: UP_ring.cfs_add R_cring cring.IP_to_UP_closed is_UP_ring) + then show "IP_to_UP i (x \\<^bsub>Pring R {i}\<^esub> y) n = (IP_to_UP i x \\<^bsub>UP R\<^esub> IP_to_UP i y) n" + using assms + by (metis LHS IP_to_UP_def) +qed + +lemma IP_to_UP_indexed_const: + assumes "k \ carrier R" + shows "IP_to_UP i (ring.indexed_const R k) = to_polynomial R k" +proof + fix x + show "IP_to_UP i (ring.indexed_const R k) x = to_polynomial R k x" + proof(cases "x = 0") + case True + have LHS: "IP_to_UP i (ring.indexed_const R k) x = k" + using True unfolding IP_to_UP_def + by (metis R.indexed_const_def nat_to_mset_zero) + then show ?thesis + using assms + unfolding to_polynomial_def + using True to_polynomial_def + by (metis UP_ring.cfs_monom is_UP_ring) + next + case False + have LHS: "IP_to_UP i (ring.indexed_const R k) x = \\<^bsub>R\<^esub>" + using False unfolding IP_to_UP_def + by (metis R.indexed_const_def nat_to_mset_inj nat_to_mset_zero) + then show ?thesis + using assms + unfolding to_polynomial_def + using False UP_cring.intro UP_cring.monom_coeff UP_cring.monom_rep_X_pow + using P_def cfs_monom by presburger + qed +qed + +lemma IP_to_UP_indexed_pmult: + assumes "p \ carrier (Pring R {i})" + shows "IP_to_UP i (ring.indexed_pmult R p i) = (IP_to_UP i p) \\<^bsub>UP R\<^esub> (X_poly R)" +proof + fix n + have 0: "IP_to_UP i p \ carrier (UP R)" + by (simp add: R_cring assms cring.IP_to_UP_closed) + show "IP_to_UP i (ring.indexed_pmult R p i) n = (IP_to_UP i p \\<^bsub>UP R\<^esub> X_poly R) n" + proof(cases "n = 0") + case True + then have RHS: "(IP_to_UP i p \\<^bsub>UP R\<^esub> X_poly R) n = \\<^bsub>R\<^esub>" + by (metis (no_types, lifting) "0" lcf_closed One_nat_def P.r_null P_def R.r_null + UP_cring.ltrm_of_X UP_cring.cfs_monom_mult UP_cring.cfs_monom_mult_l UP_zero_closed + X_closed cfs_times_X deg_leE deg_nzero_nzero is_UP_cring lessI neq0_conv plus_1_eq_Suc to_poly_inverse) + have LHS: "IP_to_UP i (ring.indexed_pmult R p i) n = ring.indexed_pmult R p i (nat_to_mset i n)" + unfolding IP_to_UP_def + by blast + then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n = + (p \\<^bsub>Pring R {i}\<^esub> (mset_to_IP R {#i#})) (nat_to_mset i n)" + using assms(1) ring.Pring_car ring.Pring_mult + ring.poly_index_mult singletonI + by (metis R.ring_axioms) + then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n = + (p \\<^bsub>Pring R {i}\<^esub> (mset_to_IP R {#i#})) {#}" + using True + by (metis nat_to_mset_zero) + then show ?thesis using RHS LHS True assms(1) nat_to_mset_zero ring.indexed_pmult_def + by (metis R.ring_axioms empty_iff set_mset_empty) + next + case False + then have RHS: "(IP_to_UP i p \\<^bsub>UP R\<^esub> X_poly R) n = (IP_to_UP i p ) (n -1)" + using "0" Suc_diff_1 Suc_eq_plus1 + assms(1) bot_nat_def IP_to_UP_def nat_neq_iff not_less0 + by (metis (no_types, lifting) P_def UP_cring X_closed cfs_times_X cring.cring_simprules(14)) + have LHS: "IP_to_UP i (ring.indexed_pmult R p i) n = ring.indexed_pmult R p i (nat_to_mset i n)" + unfolding IP_to_UP_def + by blast + then have LHS': "IP_to_UP i (ring.indexed_pmult R p i) n = + (p \\<^bsub>Pring R {i}\<^esub> (mset_to_IP R {#i#})) (nat_to_mset i n)" + using assms(1) ring.Pring_car ring.Pring_mult + ring.poly_index_mult singletonI + by (metis R.ring_axioms) + then have LHS'': "IP_to_UP i (ring.indexed_pmult R p i) n = + (p \\<^bsub>Pring R {i}\<^esub> (mset_to_IP R {#i#})) (add_mset i (nat_to_mset i (n-1))) " + by (metis False Suc_diff_1 nat_to_mset_Suc neq0_conv) + then show ?thesis using RHS unfolding IP_to_UP_def + by (metis (no_types, lifting) False R.indexed_pmult_def Suc_diff_1 add_mset_remove_trivial add_mset_remove_trivial_If multi_self_add_other_not_self nat_to_mset_Suc neq0_conv) + qed +qed + +lemma IP_to_UP_ring_hom: + shows "ring_hom_ring (Pring R {i}) (UP R) (IP_to_UP i)" + apply(rule cring.Pring_morphism) + apply (simp add: R_cring; fail) + using P_def UP_cring apply blast + apply (simp add: R.IP_to_UP_closed R_cring; fail) + apply (meson IP_to_UP_one) + apply (meson IP_to_UP_zero) + apply (meson IP_to_UP_add) + apply (metis R.IP_to_UP_var IP_to_UP_indexed_pmult singletonD) +proof- + fix k Q + assume A0: "k \ carrier R" + assume A1: "Q \ carrier (Pring R {i})" + show "IP_to_UP i (poly_scalar_mult R k Q) = + IP_to_UP i (ring.indexed_const R k) \\<^bsub>UP R\<^esub> IP_to_UP i Q" + unfolding poly_scalar_mult_def + proof + fix x + show "IP_to_UP i (\m. k \\<^bsub>R\<^esub> Q m) x = + (IP_to_UP i (ring.indexed_const R k) \\<^bsub>UP R\<^esub> IP_to_UP i Q) x" + proof- + have LHS: "IP_to_UP i (\m. k \\<^bsub>R\<^esub> Q m) x = k \\<^bsub>R\<^esub> Q (nat_to_mset i x)" + unfolding IP_to_UP_def + by blast + have RHS: "(IP_to_UP i (ring.indexed_const R k) \\<^bsub>UP R\<^esub> IP_to_UP i Q) x = + (to_polynomial R k \\<^bsub>UP R\<^esub> IP_to_UP i Q) x" + by (metis A0 IP_to_UP_indexed_const) + have RHS': "(IP_to_UP i (ring.indexed_const R k) \\<^bsub>UP R\<^esub> IP_to_UP i Q) x = + k \\<^bsub>R\<^esub> ((IP_to_UP i Q) x)" + proof- + have 0: "deg R (to_polynomial R k) = 0" + using A0 degree_to_poly by blast + have 1: "(IP_to_UP i Q) \ carrier (UP R)" + using IP_to_UP_closed unfolding P_def + by (simp add: A1 R.IP_to_UP_closed R_cring) + then show ?thesis + proof - + have "UP_cring R \ IP_to_UP i Q \ carrier (UP R)" + using "1" is_UP_cring by blast + then show ?thesis + by (metis A0 UP_cring.to_poly_mult_simp(1) UP_ring.UP_mult_closed UP_ring.coeff_simp UP_ring.coeff_smult UP_ring.monom_closed IP_to_UP_indexed_const is_UP_ring to_polynomial_def) + qed + qed + then show ?thesis + by (metis IP_to_UP_def) + qed + qed +qed + +lemma IP_to_UP_ring_hom_inj: + shows "inj_on (IP_to_UP i) (carrier (Pring R {i}))" +proof + fix x y + assume A: "x \ carrier (Pring R {i})" "y \ carrier (Pring R {i}) " + assume B: "IP_to_UP i x = IP_to_UP i y" + show "x = y" + proof + fix a + show "x a = y a" + proof(cases "set_mset a \ {i}") + case True + then obtain n where "a = (nat_to_mset i n)" + by (metis count_eq_zero_iff insert_subset multiset_eqI nat_to_msetE nat_to_msetE' + set_eq_subset singletonD singleton_insert_inj_eq' subset_insertI subset_refl) + then have LHS: "x a = IP_to_UP i x n" + by (metis IP_to_UP_def) + then show ?thesis + by (metis B \a = nat_to_mset i n\ IP_to_UP_def) + next + case False + then show ?thesis + using ring.Pring_set_zero[of R y "{i}" a] ring.Pring_set_zero[of R x "{i}" a] A + by (metis R.Pring_car R.ring_axioms) + qed + qed +qed + +lemma IP_to_UP_scalar_mult: + assumes "a \ carrier R" + assumes "p \ carrier (Pring R {i})" + shows "(IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> p)) = a\\<^bsub>UP R\<^esub> (IP_to_UP i p)" + apply(rule ring.indexed_pset.induct[of R p "{i}" "carrier R"]) + apply (simp add: R.ring_axioms; fail) + using R.Pring_car assms(2) apply blast + apply (metis IP_to_UP_indexed_const P_def R.m_closed R.poly_scalar_mult_const R.ring_axioms assms(1) ring.Pring_smult to_poly_closed to_poly_mult to_poly_mult_simp(1)) +proof- + show "\P Q. P \ Pring_set R {i} \ + IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> P) = a \\<^bsub>UP R\<^esub> IP_to_UP i P \ + Q \ Pring_set R {i} \ IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> Q) = a \\<^bsub>UP R\<^esub> IP_to_UP i Q \ IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> (P \ Q)) = a \\<^bsub>UP R\<^esub> IP_to_UP i (P \ Q)" + proof- + fix p Q + assume A0: "p \ Pring_set R {i}" + "IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> p) = a \\<^bsub>UP R\<^esub> IP_to_UP i p" + " Q \ Pring_set R {i}" + "IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> Q) = a \\<^bsub>UP R\<^esub> IP_to_UP i Q" + show "IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> (p \ Q)) = a \\<^bsub>UP R\<^esub> IP_to_UP i (p \ Q)" + proof- + have "(a \\<^bsub>Pring R {i}\<^esub> (p \ Q)) = a \\<^bsub>Pring R {i}\<^esub> p \\<^bsub>Pring R {i}\<^esub> a \\<^bsub>Pring R {i}\<^esub> Q" + by (metis A0(1) A0(3) R.Pring_add R.Pring_car R.Pring_smult_r_distr assms(1)) + then show ?thesis using A0 + by (metis IP_to_UP_add P_def R.IP_to_UP_closed R.Pring_add R.Pring_car R.Pring_smult_closed R_cring UP_smult_r_distr assms(1)) + qed + qed + show " \P ia. P \ Pring_set R {i} \ IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> P) = a \\<^bsub>UP R\<^esub> IP_to_UP i P \ ia \ {i} \ IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> (P \ ia)) = a \\<^bsub>UP R\<^esub> IP_to_UP i (P \ ia)" + proof + fix P j x + assume A0: "P \ Pring_set R {i}" + assume A1: "IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> P) = a \\<^bsub>UP R\<^esub> IP_to_UP i P" + assume A2: "j \ {i}" + then have A3: "j = i" + by blast + have "IP_to_UP i (ring.indexed_pmult R P j) \ carrier (UP R)" + by (simp add: A0 A3 R.Pring_car R.indexed_pset.indexed_pmult R_cring cring.IP_to_UP_closed) + then have "(a \\<^bsub>UP R\<^esub> (\n. ring.indexed_pmult R P j (nat_to_mset i n))) x = a \\<^bsub>R\<^esub>((\n. ring.indexed_pmult R P j (nat_to_mset i n))) x" + using A0 A1 A3 assms unfolding IP_to_UP_def + using P_def cfs_smult by blast + then show " IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> (P \ j)) x = (a \\<^bsub>UP R\<^esub> IP_to_UP i (P \ j)) x" + by (metis A0 A2 IP_to_UP_def P_def R.Pring_car R.Pring_smult_cfs R.indexed_pset.indexed_pmult \IP_to_UP i (P \ j) \ carrier (UP R)\ assms(1) cfs_smult) + qed +qed + +end + +text\Evaluation of indexed polynomials commutes with evaluation of univariate polynomials:\ + +lemma pvar_closed: + assumes "cring R" + assumes "i \ I" + shows "(pvar R i) \ carrier (Pring R I)" + by (meson assms(1) assms(2) cring.axioms(1) ring.Pring_var_closed) + +context UP_cring +begin + +lemma pvar_mult: + assumes "i \ I" + assumes "j \ I" + shows "(pvar R i) \\<^bsub>Pring R I\<^esub> (pvar R j) = mset_to_IP R {#i, j#}" +proof- + have "{#i#} + {#j#} = {# i, j#}" + by auto + then show ?thesis + unfolding var_to_IP_def + by (metis R.Pring_mult R.monom_mult) +qed + +lemma pvar_pow: + + assumes "i \ I" + shows "(pvar R i)[^]\<^bsub>Pring R I\<^esub>(n::nat) = mset_to_IP R (nat_to_mset i n)" + apply(induction n) + apply (metis Group.nat_pow_0 R.one_mset_to_IP R.ring_axioms nat_to_mset_zero ring.Pring_one) +proof- + fix n + assume IH: "pvar R i [^]\<^bsub>Pring R I\<^esub> n = mset_to_IP R (nat_to_mset i n)" + show "pvar R i [^]\<^bsub>Pring R I\<^esub> Suc n = mset_to_IP R (nat_to_mset i (Suc n)) " + proof- + have "mset_to_IP R (nat_to_mset i (Suc n)) = mset_to_IP R (nat_to_mset i n) \\<^bsub>Pring R I\<^esub> pvar R i" + using R.monom_mult[of "nat_to_mset i n" "nat_to_mset i 1"] + by (metis One_nat_def R.Pring_mult Suc_eq_plus1 nat_to_mset_Suc nat_to_mset_add nat_to_mset_zero var_to_IP_def) + then show ?thesis + using IH + by simp + qed +qed + +lemma IP_to_UP_poly_eval: + assumes "p \ Pring_set R {i}" + assumes "closed_fun R g" + shows "total_eval R g p = to_function R (IP_to_UP i p) (g i)" + apply(rule R.indexed_pset.induct[of p "{i}" "carrier R" ]) + apply (simp add: assms(1); fail) +proof- + show "\k. k \ carrier R \ total_eval R g (R.indexed_const k) = to_function R (IP_to_UP i (R.indexed_const k)) (g i)" + proof- + fix k + assume A: "k \ carrier R" + have P0: "total_eval R g (ring.indexed_const R k) = k" + unfolding total_eval_def eval_in_ring_def + using cring.poly_eval_constant[of R k UNIV g] + by (metis A R.indexed_const_def R_cring) + have P1: "(IP_to_UP i (ring.indexed_const R k)) = to_polynomial R k" + by (meson A IP_to_UP_indexed_const) + have P2: "to_function R (IP_to_UP i (ring.indexed_const R k)) (g i) = + to_function R (to_polynomial R k) (g i)" + using P1 by presburger + have P3: "to_function R (to_polynomial R k) (g i) = k" + using A assms(2) to_fun_to_poly[of k "g i"] unfolding to_fun_def by blast + then show "total_eval R g (R.indexed_const k) = to_function R (IP_to_UP i (R.indexed_const k)) (g i)" + using P0 P2 by presburger + qed + show "\P Q. P \ Pring_set R {i} \ + total_eval R g P = to_function R (IP_to_UP i P) (g i) \ + Q \ Pring_set R {i} \ + total_eval R g Q = to_function R (IP_to_UP i Q) (g i) \ total_eval R g (P \ Q) = to_function R (IP_to_UP i (P \ Q)) (g i)" + proof- + fix p Q + assume A0: "p \ Pring_set R {i}" + assume A1: " total_eval R g p = to_function R (IP_to_UP i p) (g i)" + assume A2: " Q \ Pring_set R {i}" + assume A3: "total_eval R g Q = to_function R (IP_to_UP i Q) (g i)" + have "total_eval R g (R.indexed_padd p Q) = (total_eval R g p) \\<^bsub>R\<^esub> (total_eval R g Q)" + using R.total_eval_add[of p "{i}" Q g] A0 A1 + by (metis A2 R.Pring_add R.Pring_car assms(2)) + then + have 0: "total_eval R g (p \ Q) = total_eval R g p \ total_eval R g Q " + by blast + have 1: "IP_to_UP i (p \ Q) = IP_to_UP i p \\<^bsub>UP R\<^esub> IP_to_UP i Q" + using A0 A1 A3 assms A2 R.ring_axioms R_cring IP_to_UP_add + by (metis R.Pring_add R.Pring_car) + have "g i \ carrier R" + using assms by blast + hence 2: "to_function R (IP_to_UP i (p \ Q)) (g i) = to_function R (IP_to_UP i p) (g i) \ to_function R (IP_to_UP i Q) (g i)" + using A0 A1 A3 assms A2 R.ring_axioms R_cring to_fun_plus[of "IP_to_UP i p" "IP_to_UP i Q" "g i"] + IP_to_UP_closed is_UP_cring UP_cring.to_fun_def + to_fun_def 0 1 + unfolding to_fun_def P_def + by (smt (z3) P_def R.IP_to_UP_closed R.Pring_car to_fun_plus) + show "total_eval R g (R.indexed_padd p Q) = to_function R (IP_to_UP i (ring.indexed_padd R p Q)) (g i) " + using A0 A1 A3 assms A2 R.ring_axioms R_cring is_UP_cring to_fun_def 0 1 2 + unfolding to_fun_def by metis + qed + show "\P ia. + P \ Pring_set R {i} \ + total_eval R g P = to_function R (IP_to_UP i P) (g i) \ ia \ {i} \ total_eval R g (P \ ia) = to_function R (IP_to_UP i (P \ ia)) (g i)" + proof- + fix P + fix j + assume A0: "P \ Pring_set R {i}" + assume A1: "total_eval R g P = to_function R (IP_to_UP i P) (g i)" + assume A2: "j \ {i}" + then have A3: "j = i" + by blast + show "total_eval R g (P \ j) = to_function R (IP_to_UP i (P \ j)) (g i)" + proof- + have LHS: "total_eval R g (P \ j) = (total_eval R g P) \\<^bsub>R\<^esub> (g i)" + using assms A0 A3 + by (metis R.Pring_car R_cring cring.total_eval_indexed_pmult insertI1) + have RHS: "IP_to_UP i (P \ j)= IP_to_UP i P \\<^bsub> UP R\<^esub> X_poly R" + by (metis A0 A3 IP_to_UP_indexed_pmult R.Pring_car) + have "g i \ carrier R" + using assms by blast + then show ?thesis + using A0 A1 A3 X_closed to_fun_X[of "g i"] to_fun_mult[of "IP_to_UP i P" "X_poly R" "g i"] LHS RHS + assms cring.axioms(1) domain.axioms(1) + IP_to_UP_indexed_pmult IP_to_UP_closed + Pring_car unfolding to_fun_def P_def + by (smt (z3) P.m_comm P_def R.m_comm R_cring cring.IP_to_UP_closed ring.Pring_car to_fun_closed to_fun_def) + qed + qed +qed +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Mapping Univariate Polynomials to Multivariate Polynomials over a Singleton Variable Set\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition UP_to_IP :: "('a,'b) ring_scheme \ 'c \ 'a u_poly \ ('a, 'c) mvar_poly" where +"UP_to_IP R i P = (\ m. if (set_mset m) \ {i} then P (count m i) else \\<^bsub>R\<^esub>)" + +context UP_cring +begin + +lemma UP_to_IP_inv: + assumes "p \ Pring_set R {i}" + shows "UP_to_IP R i (IP_to_UP i p) = p" +proof + fix x + show "UP_to_IP R i (IP_to_UP i p) x = p x" + proof(cases "(set_mset x) = {i}") + case True + have "{a. 0 < (\j. if j = i then count x i else 0) a} = {i}" + by (smt Collect_cong True count_eq_zero_iff neq0_conv singletonI singleton_conv) + then have "finite {j. (if j = i then count x i else 0) \ 0}" + by auto + have "(\j. if j = i then count x i else 0) = count x" + proof + fix j + show "(if j = i then count x i else 0) = count x j " + apply(cases "j = i") + using True + apply (simp; fail) + using True + by (metis count_inI singletonD) + qed + then have "(Abs_multiset (\j. if j = i then count x i else 0)) = x" + using count_inverse + by simp + then show ?thesis + unfolding UP_to_IP_def IP_to_UP_def nat_to_mset_def + by (metis True set_eq_subset) + next + case False + then show ?thesis + apply(cases "x = {#}") + apply (metis count_empty empty_subsetI IP_to_UP_def nat_to_mset_zero set_mset_empty UP_to_IP_def) + unfolding UP_to_IP_def IP_to_UP_def nat_to_mset_def + using False assms + by (metis R.Pring_set_zero set_mset_eq_empty_iff subset_singletonD) + qed +qed + +lemma UP_to_IP_const: + assumes "a \ carrier R" + shows "UP_to_IP R i (to_polynomial R a) = ring.indexed_const R a" +proof + fix x + show "UP_to_IP R i (to_polynomial R a) x = ring.indexed_const R a x" + apply(cases "x = {#}") + unfolding UP_to_IP_def + apply (metis R.indexed_const_def UP_ring.cfs_monom assms count_eq_zero_iff insert_absorb insert_not_empty is_UP_ring set_mset_empty subset_insert subset_refl to_polynomial_def) + by (metis R.indexed_const_def UP_ring.cfs_monom assms count_eq_zero_iff is_UP_ring set_mset_eq_empty_iff subset_empty subset_insert to_polynomial_def) +qed + +lemma UP_to_IP_add: + assumes "p \ carrier (UP R)" + assumes "Q \ carrier (UP R)" + shows "UP_to_IP R i (p \\<^bsub>UP R\<^esub> Q) = + UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q" +proof + fix x + show "UP_to_IP R i (p \\<^bsub>UP R\<^esub> Q) x = (UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q) x" + proof(cases "set_mset x \ {i}") + case True + have "(UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q) x = + (UP_to_IP R i p) x \\<^bsub>R\<^esub> (UP_to_IP R i Q) x" + using True assms + by (metis R.Pring_add R.indexed_padd_def) + then show ?thesis using assms True + unfolding UP_to_IP_def UP_def + by (smt partial_object.select_convs(1) restrict_def ring_record_simps(12)) + next + case False + have "(UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q) x = + (UP_to_IP R i p) x \\<^bsub>R\<^esub> (UP_to_IP R i Q) x" + using False assms + by (metis R.Pring_add R.indexed_padd_def) + then show ?thesis using False assms + unfolding UP_to_IP_def UP_def + using R.l_zero R.zero_closed by presburger + qed +qed + +lemma UP_to_IP_var: + shows "UP_to_IP R i (X_poly R) = pvar R i" +proof + have 0: "(count {#i#} i) = 1" + by simp + have 1: "set_mset {#i#} \ {i}" + by simp + have 2: "pvar R i {#i#} = \" + by (metis R.mset_to_IP_simp var_to_IP_def) + fix x + show "UP_to_IP R i (X_poly R) x = pvar R i x" + apply(cases "x = {#i#}") + using X_poly_def[of R] cfs_monom[of \ 1 "count x i"] 0 1 2 + unfolding UP_to_IP_def P_def + using R.one_closed apply presburger + proof- + assume A: "x \{#i#}" + then show "(if set_mset x \ {i} then X_poly R (count x i) else \\<^bsub>R\<^esub>) = pvar R i x" + proof(cases "set_mset x \ {i}") + case True + have "count x i \ 1" using True A + by (metis One_nat_def count_empty count_inI count_single empty_iff multiset_eqI + set_mset_add_mset_insert set_mset_empty set_mset_eq_empty_iff singletonD singletonI subset_singletonD) + then have 0: "X_poly R (count x i) = \\<^bsub>R\<^esub>" + using A UP_cring.X_closed UP_cring.degree_X UP_cring.intro True + unfolding X_poly_def + using P_def R.one_closed \\ \ carrier R \ up_ring.monom P \ 1 (count x i) = (if 1 = count x i then \ else \)\ by presburger + have "pvar R i x = \\<^bsub>R\<^esub>" + using A var_to_IP_def + by (metis R.mset_to_IP_simp') + then show ?thesis + using A "0" by presburger + next + case False + have "pvar R i x = \\<^bsub>R\<^esub>" using A var_to_IP_def False + by (metis "1" R.Pring_set_zero R.mset_to_IP_closed) + then show ?thesis + unfolding UP_to_IP_def + using False by presburger + qed + qed +qed + +lemma UP_to_IP_var_pow: + shows "UP_to_IP R i ((X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat)) = (pvar R i)[^]\<^bsub>Pring R {i}\<^esub>n" +proof + fix x + show "UP_to_IP R i (X_poly R [^]\<^bsub>UP R\<^esub> n) x = (pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) x " + proof(cases "set_mset x \ {i}") + case True + show ?thesis + proof(cases "count x i = n") + case T: True + then have 0: "x = nat_to_mset i n" + using True + by (metis count_inI emptyE insert_iff multiset_eqI nat_to_msetE + nat_to_msetE' subsetD) + have 1: "x = nat_to_mset i (count x i)" + using "0" T by auto + then have LHS: "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) x = \\<^bsub>R\<^esub>" + using T True 0 1 + by (metis R.mset_to_IP_simp insertI1 pvar_pow) + have 2: "UP_to_IP R i (X_poly R [^]\<^bsub>UP R\<^esub> n) x = (up_ring.monom (UP R) \ n) (count x i)" + unfolding UP_to_IP_def X_poly_def using True + by (metis ctrm_degree P.nat_pow_closed P.nat_pow_eone P_def R.one_closed UP_cring.monom_coeff + UP_one_closed UP_r_one X_closed is_UP_cring monom_one monom_rep_X_pow to_poly_inverse + to_poly_mult_simp(2)) + then show ?thesis + using True T LHS P_def R.one_closed cfs_monom + by presburger + next + case False + have "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) = mset_to_IP R (nat_to_mset i n)" + by (simp add: pvar_pow) + hence 0: "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) x = \" + by (metis False R.mset_to_IP_simp' nat_to_msetE) + have 1: "UP_to_IP R i (X_poly R [^]\<^bsub>UP R\<^esub> n) x = (up_ring.monom (UP R) \ n) (count x i)" + unfolding UP_to_IP_def X_poly_def using False True + by (metis ctrm_degree P.nat_pow_closed P.nat_pow_eone P_def R.one_closed UP_one_closed + UP_r_one X_closed cfs_monom monom_one monom_rep_X_pow to_poly_inverse to_poly_mult_simp(2)) + thus ?thesis using True False + unfolding UP_to_IP_def X_poly_def 0 + by (metis P_def R.one_closed cfs_monom) + qed + next + case False + then have 0: "UP_to_IP R i (X_poly R [^]\<^bsub>UP R\<^esub> n) x = \" + unfolding UP_to_IP_def + by meson + have "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) = mset_to_IP R (nat_to_mset i n)" + by (simp add: pvar_pow) + hence "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) x = \" + by (metis False R.mset_to_IP_simp' count_eq_zero_iff nat_to_msetE' singleton_iff subsetI) + then show ?thesis using 0 + by presburger + qed +qed + +lemma one_var_indexed_poly_monom_simp: + assumes "a \ carrier R" + shows "(a \\<^bsub>Pring R {i}\<^esub> ((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n)) x = (if x = (nat_to_mset i n) then a else \)" +proof- + have 0: "(a \\<^bsub>Pring R {i}\<^esub> ((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n)) x = + a \ (((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n) x)" + using Pring_smult_cfs Pring_var_closed assms cring_def is_cring monoid.nat_pow_closed ring.Pring_is_monoid singletonI + by (simp add: monoid.nat_pow_closed ring.Pring_is_monoid R.Pring_smult_cfs R.Pring_var_closed R.ring_axioms) + have 1: "(pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n = mset_to_IP R (nat_to_mset i n)" + using insertI1 + by (simp add: pvar_pow) + then have 1: "(a \\<^bsub>Pring R {i}\<^esub> ((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n)) x= + a \ (mset_to_IP R (nat_to_mset i n) x)" + using "0" by presburger + show ?thesis + using assms 1 unfolding mset_to_IP_def + using r_null r_one by simp +qed + + +lemma UP_to_IP_monom: + assumes "a \ carrier R" + shows "UP_to_IP R i (up_ring.monom (UP R) a n) = a \\<^bsub>Pring R {i}\<^esub> ((pvar R i)[^]\<^bsub>Pring R {i}\<^esub>n)" +proof + fix x + show "UP_to_IP R i (up_ring.monom (UP R) a n) x = (a \\<^bsub>Pring R {i}\<^esub> ((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n)) x" + proof(cases "set_mset x \ {i}") + case True + then show ?thesis + proof(cases "count x i = n") + case T: True + then have "x = nat_to_mset i n" + using True + by (metis count_inI emptyE insert_iff multiset_eqI nat_to_msetE + nat_to_msetE' subsetD) + then have LHS: " (a \\<^bsub>Pring R {i}\<^esub> ((pvar R i) [^]\<^bsub>Pring R {i}\<^esub> n)) x = a" + using assms + by (simp add: one_var_indexed_poly_monom_simp) + then show ?thesis + unfolding UP_to_IP_def + using T True assms(1) + by (metis UP_ring.cfs_monom is_UP_ring) + next + case False + then show ?thesis using True + unfolding UP_to_IP_def + by (metis INTEG.R.nat_to_msetE P_def assms cfs_monom one_var_indexed_poly_monom_simp) + qed + next + case False + then show ?thesis + unfolding UP_to_IP_def + by (metis (no_types, opaque_lifting) one_var_indexed_poly_monom_simp assms + count_eq_zero_iff equalityD2 insert_subset nat_to_msetE' subsetI subset_eq) + qed +qed + +lemma UP_to_IP_monom': + assumes "a \ carrier R" + shows "UP_to_IP R i (up_ring.monom (UP R) a n) = a \\<^bsub>Pring R {i}\<^esub> ((pvar R i)[^]\<^bsub>Pring R {i}\<^esub>n)" + by (metis R.Pring_smult UP_to_IP_monom assms) + +lemma UP_to_IP_closed: + assumes "p \ carrier P" + shows "(UP_to_IP R i p) \ carrier (Pring R {i})" + apply(rule poly_induct3[of ]) + using assms apply blast + apply (metis P_def R.Pring_add_closed UP_to_IP_add) +proof- + fix a fix n::nat + assume A0: "a \ carrier R" + have "(pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) \ carrier (Pring R {i})" + using pvar_closed[of R ] monoid.nat_pow_closed[of "Pring R {i}"] + proof - + show ?thesis + by (meson R.Pring_is_monoid R.Pring_var_closed monoid.nat_pow_closed singleton_iff) + qed + then show "a \ carrier R \ + UP_to_IP R i (up_ring.monom P a n) \ carrier (Pring R {i})" + using A0 assms(1) UP_to_IP_monom[of a i n] cring.poly_scalar_mult_closed [of R a _ "{i}"] + by (metis P_def R.Pring_smult_closed) +qed + +lemma IP_to_UP_inv: + assumes "p \ carrier P" + shows "IP_to_UP i (UP_to_IP R i p) = p" + apply(rule poly_induct3[of ]) + using assms apply linarith +proof- + show "\p q. q \ carrier P \ + p \ carrier P \ + IP_to_UP i (UP_to_IP R i p) = p \ + IP_to_UP i (UP_to_IP R i q) = q \ + IP_to_UP i (UP_to_IP R i (p \\<^bsub>P\<^esub> q)) = p \\<^bsub>P\<^esub> q" + proof- + fix p q assume A: + "q \ carrier P" + "p \ carrier P" + "IP_to_UP i (UP_to_IP R i p) = p" + "IP_to_UP i (UP_to_IP R i q) = q" + show "IP_to_UP i (UP_to_IP R i (p \\<^bsub>P\<^esub> q)) = p \\<^bsub>P\<^esub> q" + using A UP_to_IP_add[of p q i] + UP_to_IP_closed + IP_to_UP_add + unfolding P_def + by metis + qed + show "\a n. a \ carrier R \ + IP_to_UP i (UP_to_IP R i (up_ring.monom P a n)) = + up_ring.monom P a n" + proof- + fix a fix n::nat + assume A0: "a \ carrier R" + have A1: "pvar R i [^]\<^bsub>Pring R {i}\<^esub> n \ carrier (Pring R {i})" + using pvar_closed monoid.nat_pow_closed + by (metis R.Pring_is_monoid R_cring singletonI) + have "UP_to_IP R i (up_ring.monom (UP R) a n) = a \\<^bsub>Pring R {i}\<^esub> (pvar R i [^]\<^bsub>Pring R {i}\<^esub> n)" + by (meson A0 UP_to_IP_monom') + then have A2: "IP_to_UP i (UP_to_IP R i (up_ring.monom (UP R) a n)) = + IP_to_UP i (a \\<^bsub>Pring R {i}\<^esub> (pvar R i [^]\<^bsub>Pring R {i}\<^esub> n))" + by presburger + have A3: "IP_to_UP i (pvar R i [^]\<^bsub>Pring R {i}\<^esub> n) = (up_ring.monom P \ n)" + proof(induction n) + case 0 + then show ?case + by (metis Group.nat_pow_0 IP_to_UP_one P_def monom_one) + next + case (Suc n) + + then show ?case + using IP_to_UP_ring_hom[of i] + ring_hom_mult[of "IP_to_UP i" "Pring R {i}" "UP R" "pvar R i" "pvar R i [^]\<^bsub>Pring R {i}\<^esub> n"] + ring_hom_ring.homh[of "Pring R {i}" "UP R" "IP_to_UP i"] + by (metis IP_to_UP_monom P.l_one P.nat_pow_closed P_def R.one_closed UP_cring.ctrm_degree UP_cring.monom_rep_X_pow UP_one_closed X_closed cfs_monom is_UP_cring monom_one pvar_pow singletonI to_poly_inverse to_poly_mult_simp(1)) + qed + then show "IP_to_UP i (UP_to_IP R i (up_ring.monom P a n)) = + up_ring.monom P a n" + using A2 IP_to_UP_scalar_mult[of a "pvar R i [^]\<^bsub>Pring R {i}\<^esub> n" i] + A0 A1 P_def monic_monom_smult by presburger + qed +qed + +lemma UP_to_IP_mult: + assumes "p \ carrier (UP R)" + assumes "Q \ carrier (UP R)" + shows "UP_to_IP R i (p \\<^bsub>UP R\<^esub> Q) = + UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q" +proof- + have 0: "IP_to_UP i (UP_to_IP R i (p \\<^bsub>UP R\<^esub> Q)) = (p \\<^bsub>UP R\<^esub> Q)" + by (meson UP_cring.IP_to_UP_inv UP_ring.UP_mult_closed assms(1) assms(2) is_UP_cring is_UP_ring) + have 1: "IP_to_UP i (UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q) = + IP_to_UP i (UP_to_IP R i p) \\<^bsub>UP R\<^esub> IP_to_UP i ( UP_to_IP R i Q)" + using IP_to_UP_ring_hom[of i] + ring_hom_mult[of "IP_to_UP i"] + UP_to_IP_closed assms + by (smt P_def ring_hom_ring.homh) + have 2: "IP_to_UP i (UP_to_IP R i (p \\<^bsub>UP R\<^esub> Q)) = + IP_to_UP i (UP_to_IP R i p \\<^bsub>Pring R {i}\<^esub> UP_to_IP R i Q)" + using 0 1 assms + by (metis UP_cring.IP_to_UP_inv is_UP_cring) + then show ?thesis + by (metis "0" P_def R.Pring_mult_closed R.ring_axioms assms(1) assms(2) ring.Pring_car UP_to_IP_closed UP_to_IP_inv) +qed + +lemma UP_to_IP_ring_hom: +shows "ring_hom_ring (UP R) (Pring R {i}) (UP_to_IP R i)" + apply(rule ring_hom_ringI) + using P_def UP_ring apply force + apply (simp add: R.Pring_is_ring; fail) + apply (metis P_def UP_to_IP_closed) + apply (meson UP_to_IP_mult) + apply (meson UP_to_IP_add) + by (metis IP_to_UP_one R.Pring_car R.Pring_one_closed UP_to_IP_inv) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\The isomorphism $R[I\cup J] \sim R[I][J]$, where $I$ and $J$ are disjoint variable sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + Given a ring $R$ and variable sets $I$ and $J$, we'd like to construct the canonical + (iso)morphism $R[I\cup J] \to R[I][J]$. This can be done with the univeral property of the + previous section. Let $\phi: R \to R[J]$ be the inclusion of constants, and $f: J \to R[I]$ be + the map which sends the variable $i$ to the polynomial variable $i$ over the ring $R[I][J]$. + Then these are the two basic pieces of input required to give us a canonical homomoprhism + $R[I \cup J] \to R[I][J]$ with the universal property. The first map $\phi$ will be + "\texttt{dist\_varset\_morpshim}" below, and the second map will be + "\texttt{dist\_varset\_var\_ass}". The desired induced isomorphism will be + called "\texttt{var\_factor}". \ + +definition(in ring) dist_varset_morphism + :: "'d set \ 'd set \ + ('a, (('a, 'd) mvar_poly, 'd) mvar_poly) ring_hom" where + "dist_varset_morphism (I:: 'd set) (J:: 'd set) = + (ring.indexed_const (Pring R J) :: ('d multiset \ 'a) \ 'd multiset \ ('d multiset \ 'a))\ (ring.indexed_const R ::'a \ 'd multiset \ 'a)" + +definition(in ring) dist_varset_var_ass + :: "'d set \ 'd set \ 'd \ (('a, 'd) mvar_poly, 'd) mvar_poly" + where +"dist_varset_var_ass (I:: 'd set) (J:: 'd set) = (\i. if i \ J then ring.indexed_const (Pring R J) (pvar R i) else + pvar (Pring R J) i )" + +context cring +begin + +lemma dist_varset_morphism_is_morphism: + assumes "(I:: 'd set) \ J0 \ J1" + assumes "J1 \ I" + assumes "\ = dist_varset_morphism I J0" + shows "ring_hom_ring R (Pring (Pring R J0) J1) \" +proof- + have 0:"ring_hom_ring R (Pring R J0) indexed_const" + by (simp add: indexed_const_ring_hom) + have 1:"ring_hom_ring (Pring R J0) (Pring (Pring R J0) J1) (ring.indexed_const (Pring R J0))" + by (simp add: cring.indexed_const_ring_hom is_cring local.ring_axioms ring.Pring_is_cring) + show ?thesis using 0 1 assms ring_hom_trans[of "indexed_const" R "Pring R J0" "ring.indexed_const (Pring R J0)" + "(Pring (Pring R J0) J1) "] + unfolding dist_varset_morphism_def + by (meson ring_hom_ring.axioms(1) ring_hom_ring.axioms(2) ring_hom_ring.homh ring_hom_ringI2) +qed + +definition var_factor :: + "'d set \ 'd set \ 'd set \ + (('a, 'd) mvar_poly, (('a, 'd) mvar_poly, 'd) mvar_poly) ring_hom "where +"var_factor (I:: 'd set) (J0:: 'd set) (J1:: 'd set) = indexed_poly_induced_morphism I (Pring (Pring R J0) J1) + (dist_varset_morphism I J0) (dist_varset_var_ass I J0)" + +lemma indexed_const_closed: + assumes "x \ carrier R" + shows "indexed_const x \ carrier (Pring R I)" + using Pring_car assms indexed_pset.indexed_const by blast + +lemma var_factor_morphism: + assumes "(I:: 'd set) \ J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "g = dist_varset_var_ass I J0" + assumes "\ = dist_varset_morphism I J0" + assumes "\ = (var_factor I J0 J1)" + shows "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) \ " + "\i. i \ J0 \ I \ \ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)" + "\i. i \ J1 \ \ (pvar R i) = pvar (Pring R J0) i" + "\a. a \ carrier (Pring R (J0 \ I)) \ \ a = ring.indexed_const (Pring R J0) a" +proof- + have 0: "g \ I \ carrier (Pring (Pring R J0) J1)" + proof + fix x + assume A0: "x \ I" + then have A1: "x \ J0 \ x \ J1" + by (meson UnE assms(1) subsetD) + have A2: "x \ J0 \ x \ J1" + using A1 by blast + show "g x \ carrier (Pring (Pring R J0) J1)" + apply(cases "x \ J0") + using assms A0 A1 A2 pvar_closed[of R x J0] + pvar_closed[of "Pring R J0" x J1] + cring.indexed_const_closed[of "Pring R J0" ] + unfolding dist_varset_var_ass_def + apply (smt Pring_is_cring is_cring) + using assms A0 A1 A2 pvar_closed[of R x J0] + pvar_closed[of "Pring R J0" x J1] + cring.indexed_const_closed[of "Pring R J0" ] + unfolding dist_varset_var_ass_def + by (smt Pring_is_cring is_cring) + qed + have 1: " cring (Pring R J0)" + by (simp add: Pring_is_cring is_cring) + have 2: " cring (Pring (Pring R J0) J1)" + by (simp add: "1" Pring_is_ring ring.Pring_is_cring) + show C0: "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) \ " + using 0 assms Pring_universal_prop[of "Pring (Pring R J0) J1" g I \ \] + dist_varset_morphism_is_morphism[of I J0 J1] + unfolding var_factor_def + by (meson Pring_is_cring Pring_is_ring is_cring ring.Pring_is_cring) + show C1: "\i. i \ J0 \ I \ \ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)" + using 0 1 2 assms Pring_universal_prop(2)[of "Pring (Pring R J0) J1" g I \ \] + dist_varset_morphism_is_morphism[of I J0 J1 \] dist_varset_var_ass_def + dist_varset_morphism_def unfolding var_factor_def var_to_IP_def + by (smt IntE dist_varset_var_ass_def inf_commute inf_le2 ring.Pring_is_cring subsetD var_to_IP_def) + have 3: "\i. i \ J1 \ g i = mset_to_IP (Pring R J0) {#i#} " + using assms unfolding dist_varset_var_ass_def var_to_IP_def + by (meson disjoint_iff_not_equal) + show C2: "\i. i \ J1 \ \ (pvar R i) = pvar (Pring R J0) i" + using 0 1 2 3 assms Pring_universal_prop(2)[of "Pring (Pring R J0) J1" g I \ \] + dist_varset_morphism_is_morphism[of I J0 J1 \] + unfolding var_factor_def var_to_IP_def + by (metis subsetD) + have 4: "\k. k \ carrier R \ \ (indexed_const k) = ring.indexed_const (Pring R J0) (indexed_const k)" + using 0 1 2 3 assms Pring_universal_prop(3)[of "Pring (Pring R J0) J1" g I \ \] + dist_varset_morphism_is_morphism[of I J0 J1 \] comp_apply + unfolding var_factor_def var_to_IP_def dist_varset_morphism_def + by metis + show C3: "\a. a \ carrier (Pring R (J0 \ I)) \ \ a = ring.indexed_const (Pring R J0) a" + proof- fix a assume A: "a \ carrier (Pring R (J0 \ I))" + show " \ a = ring.indexed_const (Pring R J0) a" + apply(rule indexed_pset.induct[of a "J0 \ I" "carrier R"]) + using A Pring_car apply blast + using 4 + apply blast + proof- + show "\P Q. P \ Pring_set R (J0 \ I) \ + \ P = ring.indexed_const (Pring R J0) P \ + Q \ Pring_set R (J0 \ I) \ \ Q = ring.indexed_const (Pring R J0) Q \ \ (P \ Q) = ring.indexed_const (Pring R J0) (P \ Q)" + proof- fix P Q + assume A0: "P \ Pring_set R (J0 \ I)" + assume A1: "\ P = ring.indexed_const (Pring R J0) P" + assume A2: "Q \ Pring_set R (J0 \ I)" + assume A3: "\ Q = ring.indexed_const (Pring R J0) Q" + have A0': "P \ Pring_set R I" + using A0 + by (meson Int_lower2 Pring_carrier_subset subsetD) + have A1': "Q \ Pring_set R I" + by (meson A2 Int_lower2 Pring_carrier_subset subsetD) + have B: "\ (P \ Q) = \ P \\<^bsub>Pring (Pring R J0) J1\<^esub> \ Q" + using A0' A1' A2 A3 C0 assms ring_hom_add + by (metis (no_types, lifting) Pring_add local.ring_axioms ring.Pring_car ring_hom_ring.homh) + have " ring.indexed_const (Pring R J0) (P \ Q) = ring.indexed_const (Pring R J0) P + \\<^bsub>Pring (Pring R J0) J1\<^esub>ring.indexed_const (Pring R J0) Q" + by (simp add: Pring_add Pring_is_ring ring.Pring_add ring.indexed_padd_const) + then show " \ (P \ Q) = ring.indexed_const (Pring R J0) (P \ Q)" + using B + by (simp add: A1 A3) + qed + show "\P i. P \ Pring_set R (J0 \ I) \ \ P = ring.indexed_const (Pring R J0) P \ + i \ J0 \ I \ \ (P \ i) = ring.indexed_const (Pring R J0) (P \ i)" + proof- fix P i assume A0: "P \ Pring_set R (J0 \ I)" + assume A1: "\ P = ring.indexed_const (Pring R J0) P" + assume A2: "i \ J0 \ I" + have A0': "P \ carrier (Pring R I)" + using A0 Pring_carrier_subset + by (metis (no_types, opaque_lifting) Pring_car in_mono inf_commute le_inf_iff subset_refl) + have A0'': "P \ carrier (Pring R J0)" + using A0 Pring_carrier_subset + by (metis (no_types, opaque_lifting) Pring_car in_mono inf_commute le_inf_iff subset_refl) + have " \ (P \ i) = \ P \\<^bsub>Pring (Pring R J0) J1\<^esub> ring.indexed_const (Pring R J0) (pvar R i)" + proof- + have "(P \ i) = P \\<^bsub>Pring R I\<^esub> pvar R i" + using A0' A2 unfolding var_to_IP_def + by (metis A0 Pring_mult poly_index_mult) + then show ?thesis + using A0' C0 A2 C1[of i] ring_hom_ring.homh + ring_hom_mult[of \ "Pring R I" "Pring (Pring R J0) J1" P "pvar R i"] + by (metis IntE Pring_var_closed) + qed + then have" \ (P \ i) = ring.indexed_const (Pring R J0) P \\<^bsub>Pring (Pring R J0) J1\<^esub> + ring.indexed_const (Pring R J0) (pvar R i)" + by (simp add: A1) + then have " \ (P \ i) = ring.indexed_const (Pring R J0) (P \\<^bsub>Pring R J0\<^esub> (pvar R i))" + using A0'' A2 cring.indexed_const_ring_hom[of "Pring R J0" J1] ring_hom_ring.homh + ring_hom_mult[of "ring.indexed_const (Pring R J0)" "Pring R J0" _ P "pvar R i"] + by (smt "1" IntE Pring_var_closed) + then show "\ (P \ i) = ring.indexed_const (Pring R J0) (P \ i)" + using poly_index_mult[of P J0 i] unfolding var_to_IP_def + by (metis A0'' A2 IntE Pring_car Pring_mult) + qed + qed + qed +qed + +lemma var_factor_morphism': + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "\ = (var_factor I J0 J1)" + shows "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) \ " + "\i. i \ J1 \ \ (pvar R i) = pvar (Pring R J0) i" + "\a. a \ carrier (Pring R (J0 \ I)) \ \ a = ring.indexed_const (Pring R J0) a" + using assms var_factor_morphism + apply blast + using assms var_factor_morphism(3) + apply (metis subset_refl) + using assms var_factor_morphism(4) + by (metis Un_subset_iff Un_upper1) + +text\Constructing the inverse morphism for \texttt{var\_factor\_morphism} \ + + +lemma pvar_ass_closed: + assumes "J1 \ I" + shows "pvar R \ J1 \ carrier (Pring R I)" + by (meson Pi_I Pring_var_closed assms subsetD) + +text\The following function gives us the inverse morphism $R[I][J] \to R[I \cup J]$:\ +definition var_factor_inv :: "'d set \ 'd set \ 'd set \ + ((('a, 'd) mvar_poly, 'd) mvar_poly, ('a, 'd) mvar_poly) ring_hom" where +"var_factor_inv (I:: 'd set) (J0:: 'd set) (J1:: 'd set) = indexed_poly_induced_morphism J1 (Pring R I) + (id:: ('d multiset \ 'a) \ 'd multiset \ 'a) + (pvar R)" + +lemma var_factor_inv_morphism: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "\ = (var_factor_inv I J0 J1)" + shows "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I) \ " + "\i. i \ J1 \ \ (pvar (Pring R J0) i) = pvar R i" + "\a. a \ carrier (Pring R J0) \ \ (ring.indexed_const (Pring R J0) a) = a" +proof- + have 0: "ring_hom_ring (Pring R J0) (Pring R I) id" + apply(rule ring_hom_ringI) + apply (simp add: Pring_is_ring; fail) + apply (simp add: Pring_is_ring;fail ) + apply (metis Pring_car Pring_carrier_subset Un_upper1 assms(1) id_apply subsetD) + apply (metis Pring_mult_eq id_apply) + apply (metis Pring_add_eq id_apply) + by (simp add: Pring_one_eq) + then show "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I) \ " + using cring.Pring_universal_prop(1)[of "Pring R J0" "Pring R I" "pvar R" J1 id \] + pvar_ass_closed[of J0 I] + by (metis Pring_is_cring assms(2) assms(4) is_cring pvar_ass_closed var_factor_inv_def) + show "\i. i \ J1 \ \ (pvar (Pring R J0) i) = pvar R i" + using 0 assms pvar_ass_closed[of J0 I] + cring.Pring_universal_prop(2)[of "Pring R J0" "Pring R I" "pvar R" J1 id \] + by (metis Pring_is_cring is_cring pvar_ass_closed var_factor_inv_def var_to_IP_def) + show "\a. a \ carrier (Pring R J0) \ \ (ring.indexed_const (Pring R J0) a) = a" + using 0 assms pvar_ass_closed[of J0 I] + cring.Pring_universal_prop(3)[of "Pring R J0" "Pring R I" "pvar R" J1 id \] + by (smt Pi_I Pring_is_cring Pring_var_closed id_def is_cring subsetD var_factor_inv_def) +qed + +lemma var_factor_inv_inverse: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "\1 = (var_factor_inv I J0 J1)" + assumes "\0 = (var_factor I J0 J1)" + assumes "P \ carrier (Pring R I)" + shows "\1 (\0 P) = P" + apply(rule indexed_pset.induct[of P I "carrier R"]) + using Pring_car assms(6) apply blast + using var_factor_inv_morphism(3)[of I J0 J1 "\1"] var_factor_morphism'(3)[of I J0 J1 \0] assms + apply (metis indexed_const_closed inf_sup_absorb) +proof- + have 0: "ring_hom_ring (Pring (Pring R J0) J1) (Pring R I) \1" + by (simp add: assms(1) assms(3) assms(4) var_factor_inv_morphism(1)) + have 1: "ring_hom_ring (Pring R I) (Pring (Pring R J0) J1) \0" + by (simp add: assms(1) assms(3) assms(5) var_factor_morphism'(1)) + have 2: "\1 \ \0 \ ring_hom (Pring R I) (Pring R I)" + using 0 1 ring_hom_trans[of \0 "Pring R I" "Pring (Pring R J0) J1" \1 "Pring R I"] + ring_hom_ring.homh[of "Pring R I" "Pring (Pring R J0) J1" \0] + ring_hom_ring.homh[of "Pring (Pring R J0) J1" "Pring R I" \1] + by blast + show "\P Q. P \ Pring_set R I \ + \1 (\0 P) = P \ Q \ Pring_set R I \ \1 (\0 Q) = Q \ \1 (\0 (P \ Q)) = P \ Q" + proof- fix P Q assume A0: "P \ Pring_set R I" "\1 (\0 P) = P" + assume A1: "Q \ Pring_set R I" "\1 (\0 Q) = Q" + show "\1 (\0 (P \ Q)) = P \ Q" + using A0 A1 2 ring_hom_add[of "\1 \ \0" "Pring R I" "Pring R I" P Q] comp_apply[of \1 \0] + by (simp add: "2" \P \ Pring_set R I\ \Q \ Pring_set R I\ Pring_add Pring_car) + qed + show "\P i. P \ Pring_set R I \ \1 (\0 P) = P \ i \ I \ \1 (\0 (P \ i)) = P \ i" + proof- + fix P i assume A: "P \ Pring_set R I" "\1 (\0 P) = P" "i \ I" + show "\1 (\0 (P \ i)) = P \ i" + proof- + have A0: "P \ i = P \\<^bsub>Pring R I\<^esub> pvar R i" + by (metis A(1) A(3) Pring_mult local.ring_axioms ring.poly_index_mult var_to_IP_def) + have A1: "\1 (\0 (pvar R i)) = pvar R i" + by (metis A(3) Int_iff Pring_var_closed UnE Un_subset_iff Un_upper1 assms(1) assms(2) + assms(3) assms(4) assms(5) cring.var_factor_morphism(2) is_cring + var_factor_inv_morphism(2) var_factor_inv_morphism(3) var_factor_morphism'(2)) + then show ?thesis + using 2 A0 A ring_hom_mult[of "\1 \ \0" "Pring R I" _ P "pvar R i" ] + Pring_car Pring_var_closed comp_apply[of \1 \0] + by smt + qed + qed +qed + +lemma var_factor_total_eval: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "\ = (var_factor I J0 J1)" + assumes "closed_fun R g" + assumes "P \ carrier (Pring R I)" + shows "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P))" + apply(rule indexed_pset.induct[of P I "carrier R"]) + using Pring_car assms apply blast + apply (metis Pring_is_cring assms(1) assms(2) assms(3) assms(4) cring.total_eval_const + indexed_const_closed is_cring var_factor_morphism'(3)) +proof- + show "\P Q. P \ Pring_set R I \ + total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P)) \ + Q \ Pring_set R I \ + total_eval R g Q = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ Q)) \ + total_eval R g (P \ Q) = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ Q)))" + proof- fix P Q + assume A: " P \ Pring_set R I" + "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P))" + assume B: " Q \ Pring_set R I" + "total_eval R g Q = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ Q))" + have 0: "\ (P \ Q) = \ P \\<^bsub>Pring (Pring R J0) J1\<^esub> \ Q" + using A B assms var_factor_morphism'(1)[of I J0 J1 \] + Pring_add[of I P Q] Pring_car[of I] + ring_hom_ring.homh[of "Pring R I" "Pring (Pring R J0) J1" \] + ring_hom_add[of \ "Pring R I" "Pring (Pring R J0) J1" P Q] + by metis + have 1: "closed_fun (Pring R J0) (indexed_const \ g)" + using assms comp_apply + by (smt Pi_I closed_fun_simp indexed_const_closed) + have 2: "\ P \ carrier (Pring (Pring R J0) J1)" + using assms A var_factor_morphism'(1)[of I J0 J1 \] + ring_hom_ring.homh ring_hom_closed Pring_car + by metis + have 3: "\ Q \ carrier (Pring (Pring R J0) J1)" + using assms B var_factor_morphism'(1)[of I J0 J1 \] + ring_hom_ring.homh ring_hom_closed Pring_car + by metis + have 4: "(total_eval (Pring R J0) (indexed_const \ g) (\ (P \ Q))) = + (total_eval (Pring R J0) (indexed_const \ g) (\ P)) \\<^bsub>Pring R J0\<^esub> + (total_eval (Pring R J0) (indexed_const \ g) (\ Q))" + using 0 1 2 3 A B assms cring.total_eval_add[of "Pring R J0" "\ P" J1 "\ Q" "indexed_const \ g"] + by (metis Pring_car Pring_is_cring is_cring) + have 5: "(total_eval (Pring R J0) (indexed_const \ g) (\ P)) \ carrier (Pring R J0)" + using 3 assms cring.total_eval_closed "1" "2" Pring_is_cring is_cring + by blast + have 6: "(total_eval (Pring R J0) (indexed_const \ g) (\ Q)) \ carrier (Pring R J0)" + using 4 assms "1" "3" Pring_is_cring cring.total_eval_closed is_cring + by blast + show "total_eval R g (P \ Q) = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ Q)))" + using 5 6 4 assms + total_eval_add[of "(total_eval (Pring R J0) (indexed_const \ g) (\ p))" J0 + "(total_eval (Pring R J0) (indexed_const \ g) (\ Q))" ] + by (smt A(1) A(2) B(1) B(2) Pring_add Pring_car in_mono indexed_pset_mono order_refl + subsetD subset_iff total_eval_add) + qed + show "\P i. P \ Pring_set R I \ + total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P)) \ + i \ I \ total_eval R g (P \ i) = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ i)))" + proof- fix P i + assume A: "P \ Pring_set R I" + "total_eval R g P = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P))" + "i \ I" + have 0: "(P \ i) = P \\<^bsub>Pring R I\<^esub> (pvar R i)" + using A poly_index_mult + by (metis Pring_mult var_to_IP_def) + have 1: "\ (P \ i) = \ P \\<^bsub>Pring (Pring R J0) J1\<^esub> \ (pvar R i)" + using 0 A assms var_factor_morphism'(1)[of I J0 J1 \] + pvar_closed[of R i] ring_hom_mult ring_hom_ring.homh + by (smt Pring_car Pring_var_closed) + have 2: "\ P \ carrier (Pring (Pring R J0) J1)" + using assms A var_factor_morphism'(1)[of I J0 J1 \] + ring_hom_ring.homh ring_hom_closed Pring_car + by metis + have 3: "\ (pvar R i) \ carrier (Pring (Pring R J0) J1)" + using assms A var_factor_morphism'(1)[of I J0 J1 \] + ring_hom_ring.homh ring_hom_closed Pring_car pvar_closed[of R i I] + by (metis is_cring) + have 4: "closed_fun (Pring R J0) (indexed_const \ g)" + apply(rule cring.closed_funI) + using Pring_is_cring is_cring apply blast + using assms indexed_const_closed closed_fun_simp[of g] comp_apply[of indexed_const g] + proof - + fix x :: 'c + show "(indexed_const \ g) x \ carrier (Pring R J0)" + by (metis (no_types) \\n. closed_fun R g \ g n \ carrier R\ \\x. (indexed_const \ g) x = indexed_const (g x)\ \closed_fun R g\ indexed_const_closed) + qed + have 5: "(total_eval (Pring R J0) (indexed_const \ g) (\ (P \ i))) = + (total_eval (Pring R J0) (indexed_const \ g) (\ P)) \\<^bsub>Pring R J0\<^esub> + (total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i)))" + using 2 3 4 cring.total_eval_ring_hom[of "Pring R J0" "indexed_const \ g" J1] + by (metis "1" Pring_is_cring cring.total_eval_mult is_cring) + have 6: "(total_eval (Pring R J0) (indexed_const \ g) (\ P)) \ carrier (Pring R J0)" + using total_eval_closed "2" "4" Pring_is_cring cring.total_eval_closed is_cring + by blast + have 7: " (total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i))) \ carrier (Pring R J0)" + using "3" "4" Pring_is_cring cring.total_eval_closed is_cring + by blast + have 8: " total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ i))) = + total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ P)) \\<^bsub>R\<^esub> + total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i)))" + using 6 7 + by (metis "5" assms(5) cring.total_eval_mult is_cring) + have 9: " total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ i))) = + total_eval R g P \\<^bsub>R\<^esub> + total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i)))" + using "8" A(2) + by presburger + have 10: "total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i))) = + g i" + proof(cases "i \ J0") + case True + then have "\ (pvar R i) = ring.indexed_const (Pring R J0) (pvar R i)" + using assms + by (metis Pring_var_closed inf_sup_absorb var_factor_morphism'(3)) + then have "(total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i))) = (pvar R i)" + by (metis Pring_is_cring Pring_var_closed True cring.total_eval_const is_cring) + then show ?thesis + using total_eval_var[of g i] assms var_to_IP_def + by metis + next + case False + then have "\ (pvar R i) = (pvar (Pring R J0) i)" + by (metis A(3) UnE assms(1) assms(2) assms(3) assms(4) cring.var_factor_morphism'(2) is_cring) + then have "total_eval (Pring R J0) (indexed_const \ g) (\ (pvar R i)) = + indexed_const (g i)" + using cring.total_eval_var[of "Pring R J0" "indexed_const \ g" i] comp_apply + by (metis "4" Pring_is_cring is_cring var_to_IP_def) + then show ?thesis + using total_eval_const[of "g i" g] assms closed_fun_simp[of g i] + by metis + qed + show "total_eval R g (P \ i) = total_eval R g (total_eval (Pring R J0) (indexed_const \ g) (\ (P \ i)))" + using 9 10 + by (metis A(1) A(3) Pring_car assms(5) total_eval_indexed_pmult) + qed +qed + +lemma var_factor_closed: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "P \ carrier (Pring R I)" + shows "var_factor I J0 J1 P \ carrier (Pring (Pring R J0 ) J1)" + using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh + by (metis ring_hom_closed) + +lemma var_factor_add: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + shows "var_factor I J0 J1 (P \\<^bsub>Pring R I\<^esub> Q) = var_factor I J0 J1 P \\<^bsub>Pring (Pring R J0) J1\<^esub> + var_factor I J0 J1 Q" + using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh + by (metis (no_types, lifting) ring_hom_add) + +lemma var_factor_mult: + assumes "I = J0 \ J1" + assumes "J1 \ I" + assumes "J1 \ J0 = {}" + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + shows "var_factor I J0 J1 (P \\<^bsub>Pring R I\<^esub> Q) = var_factor I J0 J1 P \\<^bsub>Pring (Pring R J0) J1\<^esub> + var_factor I J0 J1 Q" + using assms var_factor_morphism'[of I J0 J1] ring_hom_ring.homh + by (metis (no_types, lifting) ring_hom_mult) + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Viewing a Mulitvariable Polynomial as a Univariate Polynomial over a Multivariate Polynomial Base Ring\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition multivar_poly_to_univ_poly :: + "'c set \ 'c \ ('a,'c) mvar_poly \ + (('a,'c) mvar_poly) u_poly " where +"multivar_poly_to_univ_poly I i P = ((IP_to_UP i) \ (var_factor I (I - {i}) {i})) P" + +definition univ_poly_to_multivar_poly :: + "'c set \ 'c \ (('a,'c) mvar_poly) u_poly \ + ('a,'c) mvar_poly" where +"univ_poly_to_multivar_poly I i P =((var_factor_inv I (I - {i}) {i}) \ (UP_to_IP (Pring R (I - {i})) i)) P" + +lemma multivar_poly_to_univ_poly_is_hom: + assumes "i \ I" + shows "multivar_poly_to_univ_poly I i \ ring_hom (Pring R I) (UP (Pring R (I - {i})))" + unfolding multivar_poly_to_univ_poly_def comp_apply + apply(rule ring_hom_compose[of _ "Pring (Pring R (I - {i})) {i}" _ "var_factor I (I - {i}) {i}" "IP_to_UP i"]) + apply (simp add: Pring_is_ring; fail) + apply (simp add: local.ring_axioms ring.Pring_is_ring; fail) + apply(rule UP_ring.UP_ring) unfolding UP_ring_def + apply (simp add: Pring_is_ring; fail) + using assms var_factor_morphism'[of I "I - {i}" "{i}"] unfolding ring_hom_ring_def ring_hom_ring_axioms_def + apply blast + using UP_cring.IP_to_UP_ring_hom[of "Pring R (I - {i})" i] + unfolding ring_hom_ring_def ring_hom_ring_axioms_def UP_cring_def + using Pring_is_cring is_cring apply blast + by blast + +lemma multivar_poly_to_univ_poly_inverse: + assumes "i \ I" + assumes "\0 = multivar_poly_to_univ_poly I i" + assumes "\1 = univ_poly_to_multivar_poly I i" + assumes "P \ carrier (Pring R I)" + shows "\1 (\0 P) = P" +proof- + have closed: "var_factor I (I - {i}) {i} P \ carrier (Pring (Pring R (I - {i})) {i})" + using assms var_factor_closed[of I "I - {i}" "{i}" P] + by blast + have cancel_1: "UP_to_IP (Pring R (I - {i})) i + (IP_to_UP i (var_factor I (I - {i}) {i} P)) = + var_factor I (I - {i}) {i} P" + using closed assms ring.Pring_car + UP_cring.UP_to_IP_inv[of "Pring R (I - {i})" "var_factor I (I - {i}) {i} P" i] + Pring_is_ring unfolding UP_cring_def + using is_cring local.ring_axioms ring.Pring_is_cring + by blast + have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) = + univ_poly_to_multivar_poly I i ((IP_to_UP i) (var_factor I (I - {i}) {i} P))" + by (metis comp_apply multivar_poly_to_univ_poly_def) + then have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) = + ((var_factor_inv I (I - {i}) {i}) ((UP_to_IP (Pring R (I - {i})) i) + ((IP_to_UP i) (var_factor I (I - {i}) {i} P))))" + using comp_apply univ_poly_to_multivar_poly_def + by metis + then have "univ_poly_to_multivar_poly I i (multivar_poly_to_univ_poly I i P) = + ((var_factor_inv I (I - {i}) {i}) (var_factor I (I - {i}) {i} P))" + using cancel_1 + by presburger + then show ?thesis using assms var_factor_inv_inverse[of I "I - {i}" "{i}" _ _ P] + by (metis Diff_cancel Diff_disjoint Diff_eq_empty_iff Diff_partition Un_Diff_cancel + Un_Diff_cancel2 Un_insert_right empty_Diff empty_subsetI insert_Diff_if insert_absorb ) +qed + +lemma multivar_poly_to_univ_poly_total_eval: + assumes "i \ I" + assumes "\ = multivar_poly_to_univ_poly I i" + assumes "P \ carrier (Pring R I)" + assumes "closed_fun R g" + shows "total_eval R g P = total_eval R g (to_function (Pring R (I - {i})) (\ P) (indexed_const (g i)))" +proof- + have 0: "var_factor I (I - {i}) {i} P \ Pring_set (Pring R (I - {i})) {i}" + proof- + have 00: " ring (Pring R (I - {i}))" + using Pring_is_ring + by blast + then show ?thesis + using assms(1) assms(3) var_factor_closed[of I "I - {i}" "{i}" P] ring.Pring_car[of "Pring R (I - {i})" "{i}"] + by blast + qed + have 1: "closed_fun (Pring R (I - {i})) (indexed_const \ g)" + using assms comp_apply indexed_const_closed cring.closed_funI[of "Pring R (I - {i})"] + by (metis Pring_is_cring closed_fun_simp is_cring) + have 2: "cring (Pring R (I - {i}))" + using Pring_is_cring is_cring by blast + have "(to_function (Pring R (I - {i})) (\ P) (indexed_const (g i))) + = total_eval (Pring R (I - {i})) (indexed_const \ g) ((var_factor I (I - {i}) {i}) P) " + using assms 0 1 2 comp_apply UP_cring.IP_to_UP_poly_eval[of "Pring R (I - {i})" + "(var_factor I (I - {i}) {i}) P" i "indexed_const \ g"] + unfolding UP_cring_def multivar_poly_to_univ_poly_def + by metis + then have 3: "total_eval R g ((to_function (Pring R (I - {i})) (\ P) (indexed_const (g i))) ) + = total_eval R g ( total_eval (Pring R (I - {i})) (indexed_const \ g) ((var_factor I (I - {i}) {i}) P)) " + by presburger + then have "total_eval R g P = total_eval R g (total_eval (Pring R (I - {i})) (indexed_const \ g) ((var_factor I (I - {i}) {i}) P)) " + using assms var_factor_total_eval[of I "I - {i}" "{i}" "var_factor I (I - {i}) {i}" g P] + by blast + then show ?thesis + using 3 + by presburger +qed + +text\ + Induction for polynomial rings. Basically just \texttt{indexed\_pset.induct} with some + boilerplate translations removed +\ + +lemma(in ring) Pring_car_induct'': + assumes "Q \ carrier (Pring R I)" + assumes "\c. c \ carrier R \ P (indexed_const c)" + assumes "\p q. p \ carrier (Pring R I) \ q \ carrier (Pring R I) \ P p \ P q \ P (p \\<^bsub>Pring R I\<^esub> q)" + assumes "\p i. p \ carrier (Pring R I) \ i \ I \ P p \ P (p \\<^bsub>Pring R I\<^esub> pvar R i)" + shows "P Q" + apply(rule indexed_pset.induct[of Q I "carrier R"]) + using Pring_car assms(1) apply blast + using assms(2) apply blast + apply (metis (full_types) Pring_add Pring_car assms(3)) +proof- + fix Pa i assume A: "Pa \ Pring_set R I" "P Pa" "i \ I" + then have 0: "Pa \ carrier (Pring R I)" + using assms A Pring_car + by blast + have "Pa \ i = Pa \\<^bsub>Pring R I\<^esub> pvar R i" + using 0 poly_index_mult assms A + by (metis Pring_mult var_to_IP_def) + then show "P (Pa \ i)" + by (simp add: "0" A(2) A(3) assms) +qed + +subsubsection\Application: A Polynomial Ring over a Domain is a Domain\ + +text \ + In this section, we use the fact the UP \R\ is a domain when \R\ is a domain to show the analogous + result for indexed polynomial rings. We first prove it inductively for rings with a finite + variable set, and then generalize to infinite variable sets using the fact that any two + multivariable polynomials over an indexed polynomial ring must also lie in a finitely indexed + polynomial subring. +\ + +lemma indexed_const_mult: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "indexed_const a \\<^bsub> Pring R I\<^esub> indexed_const b = indexed_const (a \ b)" + by (metis Pring_mult assms(1) assms(2) indexed_const_P_mult_eq) + +lemma(in domain) Pring_fin_vars_is_domain: + assumes "finite (I ::'c set)" + shows "domain (Pring R I)" +proof(rule finite_induct, rule assms) + show "domain (Pring R ({}::'c set))" + proof(rule domainI) + show " cring (Pring R {})" + by (simp add: Pring_is_cring is_cring) + show "\\<^bsub>Pring R {}\<^esub> \ \\<^bsub>Pring R {}\<^esub>" + proof- + have "\\<^bsub>Pring R {}\<^esub> {#} \ \\<^bsub>Pring R {}\<^esub> {#}" + by (metis Pring_one Pring_zero local.ring_axioms ring.indexed_const_def zero_not_one) + thus ?thesis + by metis + qed + show "\a b. a \\<^bsub>Pring R ({}::'c set)\<^esub> b = \\<^bsub>Pring R {}\<^esub> \ + a \ carrier (Pring R {}) \ b \ carrier (Pring R {}) \ a = \\<^bsub>Pring R {}\<^esub> \ b = \\<^bsub>Pring R {}\<^esub>" + proof- + fix a b assume A: "a \\<^bsub>Pring R ({}::'c set)\<^esub> b = \\<^bsub>Pring R {}\<^esub>" + "a \ carrier (Pring R {})" " b \ carrier (Pring R {})" + have 0: "monomials_of R a \ {{#}}" + using A Pring_set_zero unfolding monomials_of_def Pring_car + by blast + have 1: "monomials_of R b \ {{#}}" + using A Pring_set_zero unfolding monomials_of_def Pring_car + by blast + obtain A where A_def: "A = a {#}" + by blast + obtain B where B_def: "B = b {#}" + by blast + have 2: "a = indexed_const A" + unfolding A_def + apply(rule ext) + by (metis 0 complement_of_monomials_of empty_iff in_mono indexed_const_def insert_iff) + have 3: "b = indexed_const B" + unfolding B_def + apply(rule ext) + by (metis 1 complement_of_monomials_of empty_iff in_mono indexed_const_def insert_iff) + have 4: "a \\<^bsub>Pring R {}\<^esub> b = indexed_const (A \ B)" + using A unfolding 2 3 + by (metis "2" "3" A_def B_def Pring_carrier_coeff' indexed_const_mult) + have 5: "A \ B = \" + using A unfolding 4 by (metis Pring_zero indexed_const_def) + have 6: "A = \ \ B = \" + using 5 A_def B_def A + by (simp add: domain.integral_iff domain_axioms) + show "a = \\<^bsub>Pring R {}\<^esub> \ b = \\<^bsub>Pring R {}\<^esub>" + unfolding 2 3 using 6 + by (metis Pring_zero) + qed + qed + show "\x F. finite (F:: 'c set) \ x \ F \ domain (Pring R F) \ domain (Pring R (insert x F))" + proof- fix S:: "'c set" fix s assume A: "finite S" "s \ S" "domain (Pring R S)" + show "domain (Pring R (insert s S))" + proof- + have ring_hom: "multivar_poly_to_univ_poly (insert s S) s \ ring_hom (Pring R (insert s S)) (UP (Pring R S))" + using multivar_poly_to_univ_poly_is_hom + by (metis A(2) Diff_insert_absorb insertI1) + have domain: "domain (UP (Pring R S))" + apply(rule UP_domain.UP_domain) + unfolding UP_domain_def by(rule A) + show "domain (Pring R (insert s S))" + apply(rule ring_hom_ring.inj_on_domain[of _ "UP (Pring R S)" "multivar_poly_to_univ_poly (insert s S) s"]) + apply(rule ring_hom_ring.intro) + apply (simp add: Pring_is_ring; fail) + apply(rule UP_ring.UP_ring) + unfolding UP_ring_def + apply (simp add: Pring_is_ring; fail) + unfolding ring_hom_ring_axioms_def apply(rule ring_hom) + proof(rule inj_onI) + fix x y assume A: "x \ carrier (Pring R (insert s S))" + "y \ carrier (Pring R (insert s S))" + "multivar_poly_to_univ_poly (insert s S) s x = multivar_poly_to_univ_poly (insert s S) s y" + then have 0: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s x) + = univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s y)" + by auto + have 1: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s x) = x" + using A by (meson cring.multivar_poly_to_univ_poly_inverse insertI1 is_cring) + have 2: "univ_poly_to_multivar_poly (insert s S) s (multivar_poly_to_univ_poly (insert s S) s y) = y" + using A by (meson insertI1 multivar_poly_to_univ_poly_inverse) + show "x = y" + using 0 unfolding 1 2 by auto + next + show "domain (UP (Pring R S))" + apply(rule UP_domain.UP_domain) + unfolding UP_domain_def by(rule A) + qed + qed + qed +qed + +lemma locally_finite: + assumes "a \ carrier (Pring R I)" + shows "\J. J \ I \ finite J \ a \ carrier (Pring R J)" +proof(rule Pring_car_induct[of _ I], rule assms) + have 0: "\\<^bsub>Pring R I\<^esub> = \\<^bsub>Pring R {}\<^esub>" + by (simp add: Pring_zero_eq) + show "\J\I. finite J \ \\<^bsub>Pring R I\<^esub> \ carrier (Pring R J)" + unfolding 0 + by (meson Pring_zero_closed empty_subsetI finite.emptyI) +next + fix m k assume A: "set_mset m \ I \ k \ carrier R" + obtain J where J_def: "J = set_mset m" + by blast + have 0: "k \\<^bsub>Pring R I\<^esub> mset_to_IP R m = k \\<^bsub>Pring R J\<^esub> mset_to_IP R m" + unfolding J_def by (simp add: Pring_smult) + have 1: "k \\<^bsub>Pring R I\<^esub> mset_to_IP R m \ carrier (Pring R J)" + unfolding 0 J_def using A + by (simp add: Pring_car Pring_smult mset_to_IP_closed poly_scalar_mult_closed) + show "\J\I. finite J \ k \\<^bsub>Pring R I\<^esub> mset_to_IP R m \ carrier (Pring R J)" + using J_def 1 A by blast +next + fix Q m k + assume A: "Q \ carrier (Pring R I)\(\J\I. finite J \ Q \ carrier (Pring R J)) \ set_mset m \ I \ k \ carrier R" + then obtain J where J_def: "J\I \ finite J \ Q \ carrier (Pring R J)" + by blast + obtain J' where J'_def: "J' = J \ set_mset m" + by blast + have 0: "finite J'" + unfolding J'_def using J_def by blast + have 1: "k \\<^bsub>Pring R I\<^esub> mset_to_IP R m = k \\<^bsub>Pring R J'\<^esub> mset_to_IP R m" + unfolding J'_def using A by (simp add: Pring_smult) + have 2: "Q \ carrier (Pring R J')" + using J_def unfolding J'_def + by (metis Pring_car Pring_carrier_subset Un_upper1 subsetD) + have 3: "Q \ k \\<^bsub>Pring R I\<^esub> mset_to_IP R m \ carrier (Pring R J')" + using 0 1 2 J'_def A + by (metis Pring_car Pring_smult_closed indexed_pset.indexed_padd mset_to_IP_closed sup.cobounded2) + show "\J'\I. finite J' \ Q \ k \\<^bsub>Pring R I\<^esub> mset_to_IP R m \ carrier (Pring R J')" + using 3 0 + by (metis A J'_def J_def Un_subset_iff) +qed + +lemma(in domain) Pring_is_domain: + "domain (Pring R I)" +proof(rule domainI, simp add: Pring_is_cring is_cring) + have 0: "\\<^bsub>Pring R I\<^esub> {#} = \" + by (simp add: Pring_one indexed_const_def) + have 1: "\\<^bsub>Pring R I\<^esub> {#} = \" + by (simp add: Pring_zero indexed_const_def) + have 2: "\ \ \" + by simp + show "\\<^bsub>Pring R I\<^esub> \ \\<^bsub>Pring R I\<^esub>" + using 0 1 2 by auto +next + fix a b assume A: "a \\<^bsub>Pring R I\<^esub> b = \\<^bsub>Pring R I\<^esub>" + "a \ carrier (Pring R I)" + "b \ carrier (Pring R I)" + obtain J0 where J0_def: "J0 \ I \ finite J0 \ a \ carrier (Pring R J0)" + using A locally_finite by blast + obtain J1 where J1_def: "J1 \ I \ finite J1 \ b \ carrier (Pring R J1)" + using A locally_finite by blast + obtain J where J_def: "J = J0 \ J1" + by blast + have J_finite: "finite J" + using J_def J0_def J1_def by blast + have 0: "a \ carrier (Pring R J)" + using J0_def unfolding J_def + by (metis (no_types, lifting) Pring_car Pring_carrier_subset Un_upper1 subset_eq) + have 1: "b \ carrier (Pring R J)" + using J1_def unfolding J_def + by (metis Pring_car Pring_carrier_subset in_mono sup.cobounded2) + have 2: "a \\<^bsub>Pring R J\<^esub> b = \\<^bsub>Pring R J\<^esub>" + using A J_def 0 1 by (metis Pring_mult_eq Pring_zero_eq) + have 3: "domain (Pring R J)" + using J_finite Pring_fin_vars_is_domain[of J] by blast + have 4: "a = \\<^bsub>Pring R J\<^esub> \ b = \\<^bsub>Pring R J\<^esub>" + using 3 2 0 1 by (simp add: domain.integral) + thus "a = \\<^bsub>Pring R I\<^esub> \ b = \\<^bsub>Pring R I\<^esub>" + using Pring_zero_eq by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsubsection\Relabelling of Variables for Indexed Polynomial Rings\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition relabel_vars :: "'d set \ 'e set \ ('d \ 'e) \ + ('a, 'd) mvar_poly \ ('a, 'e) mvar_poly" where +"relabel_vars I J g = indexed_poly_induced_morphism I (Pring R J) indexed_const (\i. pvar R (g i))" + +lemma relabel_vars_is_morphism: + assumes "g \ I \ J" + shows "ring_hom_ring (Pring R I) (Pring R J) (relabel_vars I J g)" + "\i. i \ I \ relabel_vars I J g (pvar R i) = pvar R (g i)" + "\c. c \ carrier R \ relabel_vars I J g (indexed_const c) = indexed_const c" + using Pring_universal_prop(1)[of "Pring R J" "\i. pvar R (g i)" I indexed_const] + assms unfolding relabel_vars_def + apply (meson Pi_iff Pring_is_cring Pring_var_closed indexed_const_ring_hom is_cring) +proof- + have 0: "cring (Pring R J)" " (\i. mset_to_IP R {#g i#}) \ I \ carrier (Pring R J)" + "ring_hom_ring R (Pring R J) indexed_const" + using assms Pring_is_cring is_cring apply blast + apply (smt Pi_iff Pring_var_closed assms var_to_IP_def) + by (simp add: indexed_const_ring_hom) + show "\i. i \ I \ indexed_poly_induced_morphism I (Pring R J) + indexed_const (\i. pvar R (g i)) (pvar R i) = pvar R (g i)" + using 0 assms Pring_universal_prop(2)[of "Pring R J" "\i. pvar R (g i)" I indexed_const] + unfolding relabel_vars_def var_to_IP_def + by blast + show "\c. c \ carrier R \ indexed_poly_induced_morphism I (Pring R J) + indexed_const (\i. pvar R (g i)) (indexed_const c) = indexed_const c" + using 0 assms Pring_universal_prop(3)[of "Pring R J" "\i. pvar R (g i)" I indexed_const] + unfolding var_to_IP_def + by blast +qed + +lemma relabel_vars_add: + assumes "g \ I \ J" + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + shows "relabel_vars I J g (P \\<^bsub>Pring R I\<^esub> Q) = relabel_vars I J g P \\<^bsub>Pring R J\<^esub> relabel_vars I J g Q" + using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh ring_hom_add + by (metis (no_types, lifting)) + +lemma relabel_vars_mult: + assumes "g \ I \ J" + assumes "P \ carrier (Pring R I)" + assumes "Q \ carrier (Pring R I)" + shows "relabel_vars I J g (P \\<^bsub>Pring R I\<^esub> Q) = relabel_vars I J g P \\<^bsub>Pring R J\<^esub> relabel_vars I J g Q" + using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh ring_hom_mult + by (metis (no_types, lifting)) + +lemma relabel_vars_closed: + assumes "g \ I \ J" + assumes "P \ carrier (Pring R I)" + shows "relabel_vars I J g P \ carrier (Pring R J)" + using assms relabel_vars_is_morphism(1)[of g I J] ring_hom_ring.homh + by (metis ring_hom_closed) + +lemma relabel_vars_smult: + assumes "g \ I \ J" + assumes "P \ carrier (Pring R I)" + assumes "a \ carrier R" + shows "relabel_vars I J g (a \\<^bsub>Pring R I\<^esub>P) = a \\<^bsub>Pring R J\<^esub>relabel_vars I J g P" +proof- + have 0: "a \\<^bsub>Pring R I\<^esub>P = indexed_const a \\<^bsub>Pring R I\<^esub> P" + by (metis Pring_car Pring_mult Pring_smult assms(2) assms(3) poly_scalar_mult_eq) + have 1: "a \\<^bsub>Pring R J\<^esub>relabel_vars I J g P = indexed_const a \\<^bsub>Pring R J\<^esub> relabel_vars I J g P" + by (metis Pring_car Pring_mult Pring_smult assms(1) assms(2) assms(3) poly_scalar_mult_eq relabel_vars_closed) + show ?thesis using 0 1 assms relabel_vars_mult relabel_vars_is_morphism(3)[of g I J a] + by (metis indexed_const_closed) +qed + +lemma relabel_vars_inverse: + assumes "g \ I \ J" + assumes "h \ J \ I" + assumes "\i. i \ I \ h (g i) = i" + assumes "P \ carrier (Pring R I)" + shows "relabel_vars J I h (relabel_vars I J g P) = P" + apply(rule Pring_car_induct''[of _ I]) + using assms(4) apply blast + using assms + apply (metis relabel_vars_is_morphism(3)) +proof- + show "\p q. p \ carrier (Pring R I) \ + q \ carrier (Pring R I) \ + relabel_vars J I h (relabel_vars I J g p) = p \ + relabel_vars J I h (relabel_vars I J g q) = q \ + relabel_vars J I h (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q)) = p \\<^bsub>Pring R I\<^esub> q" + proof- fix p q assume A: " p \ carrier (Pring R I)" "q \ carrier (Pring R I)" + "relabel_vars J I h (relabel_vars I J g p) = p" + "relabel_vars J I h (relabel_vars I J g q) = q" + show "relabel_vars J I h (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q)) = p \\<^bsub>Pring R I\<^esub> q" + proof- + have 0: "relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q) = relabel_vars I J g p \\<^bsub>Pring R J\<^esub> relabel_vars I J g q" + using A(1) A(2) assms(1) relabel_vars_add by blast + then show ?thesis + using assms A relabel_vars_is_morphism(3)[of g I J] relabel_vars_is_morphism(3)[of h J I] + relabel_vars_closed[of g I J p] relabel_vars_closed[of g I J q] + relabel_vars_add[of h J I "relabel_vars I J g p" "relabel_vars I J g q"] + by presburger + qed + qed + show "\p i. p \ carrier (Pring R I) \ + i \ I \ + relabel_vars J I h (relabel_vars I J g p) = p \ + relabel_vars J I h (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> pvar R i)) = p \\<^bsub>Pring R I\<^esub> pvar R i" + proof- fix p i assume A: " p \ carrier (Pring R I)" + "relabel_vars J I h (relabel_vars I J g p) = p" + "i \ I" + show " relabel_vars J I h (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> pvar R i)) = p \\<^bsub>Pring R I\<^esub> pvar R i" + proof- + have "relabel_vars J I h (relabel_vars I J g (pvar R i)) = pvar R i" + using assms relabel_vars_is_morphism[of g I J] relabel_vars_is_morphism[of h J I] + by (metis A(3) funcset_mem ) + then show ?thesis + using assms relabel_vars_is_morphism[of g I J] relabel_vars_is_morphism[of h J I] + relabel_vars_closed[of g I J p] relabel_vars_closed[of g I J "pvar R i"] + relabel_vars_mult[of g I J "p" "pvar R i"] + relabel_vars_mult[of h J I "relabel_vars I J g p" "relabel_vars I J g (pvar R i)"] + by (metis A(1) A(2) A(3) Pring_var_closed) + qed + qed +qed + +lemma relabel_vars_total_eval: + assumes "g \ I \ J" + assumes "P \ carrier (Pring R I)" + assumes "closed_fun R f" + shows "total_eval R (f \ g) P = total_eval R f (relabel_vars I J g P)" +proof(rule Pring_car_induct''[of P I]) + show "P \ carrier (Pring R I)" + using assms(2) by blast + show "\c. c \ carrier R \ total_eval R (f \ g) (indexed_const c) = total_eval R f (relabel_vars I J g (indexed_const c))" + by (metis assms(1) relabel_vars_is_morphism(3) total_eval_const) + show " \p q. p \ carrier (Pring R I) \ + q \ carrier (Pring R I) \ + total_eval R (f \ g) p = total_eval R f (relabel_vars I J g p) \ + total_eval R (f \ g) q = total_eval R f (relabel_vars I J g q) \ + total_eval R (f \ g) (p \\<^bsub>Pring R I\<^esub> q) = total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q))" + + proof- fix p q assume A: "p \ carrier (Pring R I)" "q \ carrier (Pring R I)" + "total_eval R (f \ g) p = total_eval R f (relabel_vars I J g p)" + "total_eval R (f \ g) q = total_eval R f (relabel_vars I J g q)" + have 0: "closed_fun R (f \ g)" + apply(rule closed_funI) + using comp_apply[of f g] closed_fun_simp[of f] assms(3) by presburger + have 1: " (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q)) = + (relabel_vars I J g p) \\<^bsub>Pring R J\<^esub> (relabel_vars I J g q)" + using A(1) A(2) assms(1) relabel_vars_add by blast + have 2: "total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q)) = + (total_eval R f (relabel_vars I J g p)) \\<^bsub>R\<^esub> (total_eval R f (relabel_vars I J g q))" + using total_eval_add[of _ J _ f] + by (metis "1" A(1) A(2) A(3) A(4) assms(1) assms(3) relabel_vars_closed) + show "total_eval R (f \ g) (p \\<^bsub>Pring R I\<^esub> q) = total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> q))" + using A 0 1 2 + by (metis total_eval_add) + qed + show "\p i. p \ carrier (Pring R I) \ + i \ I \ + total_eval R (f \ g) p = total_eval R f (relabel_vars I J g p) \ + total_eval R (f \ g) (p \\<^bsub>Pring R I\<^esub> pvar R i) = total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> pvar R i))" + proof- fix p i assume A: "p \ carrier (Pring R I)" " i \ I " + "total_eval R (f \ g) p = total_eval R f (relabel_vars I J g p)" + have 0: "closed_fun R (f \ g)" + apply(rule closed_funI) + using comp_apply[of f g] closed_fun_simp[of f] assms(3) by presburger + have 1: " (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> (pvar R i))) = + (relabel_vars I J g p) \\<^bsub>Pring R J\<^esub> (relabel_vars I J g (pvar R i))" + by (meson A(1) A(2) assms(1) local.ring_axioms relabel_vars_mult ring.Pring_var_closed) + have 2: "total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> (pvar R i))) = + (total_eval R f (relabel_vars I J g p)) \\<^bsub>R\<^esub> (total_eval R f (relabel_vars I J g (pvar R i)))" + using total_eval_mult[of "relabel_vars I J g p" J "relabel_vars I J g (pvar R i)"] + by (metis "1" A(1) A(2) A(3) assms(1) assms(3) is_cring pvar_closed relabel_vars_closed) + have 3: " total_eval R (f \ g) (p \\<^bsub>Pring R I\<^esub> pvar R i) = + total_eval R (f \ g) p \\<^bsub>R\<^esub> total_eval R (f \ g)( pvar R i)" + by (meson "0" A(1) A(2) Pring_var_closed total_eval_mult) + have 4: "total_eval R (f \ g)( pvar R i) = (total_eval R f (relabel_vars I J g (pvar R i)))" + proof- + have 40: "total_eval R (f \ g)( pvar R i) = (f \ g) i" + using total_eval_var[of "f \g" i] + by (metis "0" var_to_IP_def) + have 41: "relabel_vars I J g (pvar R i) = pvar R (g i)" + by (simp add: A(2) assms(1) relabel_vars_is_morphism(2)) + have 42: "total_eval R f (relabel_vars I J g (pvar R i)) = f (g i)" + using total_eval_var + by (metis "41" assms(3) var_to_IP_def) + show ?thesis + by (metis "40" "42" comp_apply) + qed + show " total_eval R (f \ g) (p \\<^bsub>Pring R I\<^esub> pvar R i) = + total_eval R f (relabel_vars I J g (p \\<^bsub>Pring R I\<^esub> pvar R i))" + using A 0 1 2 3 4 + by presburger + qed +qed + +end + + +end + diff --git a/thys/Padic_Field/Fraction_Field.thy b/thys/Padic_Field/Fraction_Field.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Fraction_Field.thy @@ -0,0 +1,1175 @@ +theory Fraction_Field + imports "HOL-Algebra.UnivPoly" + Localization_Ring.Localization + "HOL-Algebra.Subrings" + Padic_Ints.Supplementary_Ring_Facts +begin + +section\The Field of Fractions of a Ring\ + +text\ + This theory defines the fraction field of a domain and establishes its basic properties. + The fraction field is defined in the standard way as the localization of a domain at its nonzero + elements. This is done by importing the AFP session \texttt{Localization\_Ring}. Choice functions + for numerator and denominators of fractions are introduced, and the inclusion of a domain into + its ring of fractions is defined. +\ + +subsection\The Monoid of Nonzero Elements in a Domain\ +locale domain_frac = domain + +lemma zero_not_in_nonzero: "\\<^bsub>R\<^esub> \ nonzero R" + unfolding nonzero_def by blast + +lemma(in domain) nonzero_is_submonoid: "submonoid R (nonzero R)" +proof + show " nonzero R \ carrier R" + using nonzero_def by fastforce + show "\x y. x \ nonzero R \ y \ nonzero R \ x \ y \ nonzero R" + by (metis (mono_tags, lifting) local.integral m_closed mem_Collect_eq nonzero_def) + show "\ \ nonzero R" + by (simp add: nonzero_def) +qed + +lemma(in domain) nonzero_closed: + assumes "a \ nonzero R" + shows "a \ carrier R" + using assms + by (simp add: nonzero_def) + +lemma(in domain) nonzero_mult_closed: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "a \ b \ carrier R" + using assms + by (simp add: nonzero_def) + +lemma(in domain) nonzero_one_closed: +"\ \ nonzero R" + by (simp add: nonzero_def) + +lemma(in domain) nonzero_memI: + assumes "a \ carrier R" + assumes "a \ \" + shows "a \ nonzero R" + using assms by(simp add: nonzero_def) + +lemma(in domain) nonzero_memE: + assumes "a \ nonzero R" + shows "a \ carrier R" "a \\" + using assms by(auto simp: nonzero_def) + +lemma(in domain) not_nonzero_memE: + assumes "a \ nonzero R" + assumes "a \ carrier R" + shows "a = \" + using assms + by (simp add: nonzero_def) + +lemma(in domain) not_nonzero_memI: + assumes "a = \" + shows "a \ nonzero R" + using assms nonzero_memE(2) by auto + +lemma(in domain) one_nonzero: +"\ \ nonzero R" + by (simp add: nonzero_one_closed) + +lemma(in domain_frac) eq_obj_rng_of_frac_nonzero: + "eq_obj_rng_of_frac R (nonzero R)" + using nonzero_is_submonoid + by (simp add: eq_obj_rng_of_frac.intro + is_cring local.ring_axioms mult_submonoid_of_crng_def mult_submonoid_of_rng_def) + +subsection\Numerator and Denominator Choice Functions\ + +definition(in eq_obj_rng_of_frac) denom where +"denom a = (if (a = \\<^bsub>rec_rng_of_frac\<^esub>) then \ else ( snd (SOME x. x \ a)))" + +text\The choice function for numerators must be compatible with denom:\ + +definition(in eq_obj_rng_of_frac) numer where +"numer a = (if (a = \\<^bsub>rec_rng_of_frac\<^esub>) then \ else (fst (SOME x. x \ a \ (snd x = denom a))))" + +text\Basic properties of numer and denom:\ +lemma(in eq_obj_rng_of_frac) numer_denom_facts0: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier rec_rng_of_frac" + assumes "a \ \\<^bsub>rec_rng_of_frac\<^esub>" + shows "a = ((numer a) |\<^bsub>rel\<^esub> (denom a))" + "(numer a) \ carrier R" + "(denom a) \ S" + "numer a = \ \ a = \\<^bsub>rec_rng_of_frac\<^esub>" + "a \\<^bsub>rec_rng_of_frac\<^esub> (rng_to_rng_of_frac(denom a)) = rng_to_rng_of_frac (numer a)" + "(rng_to_rng_of_frac(denom a)) \\<^bsub>rec_rng_of_frac\<^esub> a = rng_to_rng_of_frac (numer a)" +proof- + have 0: "carrier rel \ {}" + by (metis (no_types, lifting) SigmaI empty_iff one_closed partial_object.select_convs(1) rel_def zero_closed) + have 1: "(numer a, denom a) \ a" + proof- + have "\ r s. (r \ carrier R \ s \ S \ (a = (r |\<^bsub>rel\<^esub> s)))" + using rel_def assms(3) assms(1) set_eq_class_of_rng_of_frac_def rec_rng_of_frac_def + by (smt mem_Collect_eq mem_Sigma_iff partial_object.select_convs(1)) + then obtain r s where "r \ carrier R \ s \ S \ (a = (r |\<^bsub>rel\<^esub> s))" + by blast + hence "a = class_of\<^bsub>rel\<^esub> (r, s)" + by (simp add: class_of_to_rel) + hence "(r,s) \ a" using eq_class_of_def rel_def + using \r \ carrier R \ s \ S \ a = (r |\<^bsub>rel\<^esub> s)\ equiv_obj_rng_of_frac equivalence.refl by fastforce + hence "(\ x. x \ a)" + by blast + hence "(SOME x. x \ a) \ a" + by (meson some_eq_ex) + hence "(\ x. x \ a \ (snd x = denom a))" + using denom_def assms by metis + hence "\x. x \ a \ (snd x = denom a) \ (fst x = numer a)" + using numer_def assms by (metis (mono_tags, lifting) exE_some) + thus ?thesis + by simp + qed + have "a \ {r |\<^bsub>rel\<^esub> s |r s. (r, s) \ carrier rel}" + using assms(3) rec_rng_of_frac_def set_eq_class_of_rng_of_frac_def by auto + hence "\ x y. a = (x |\<^bsub>rel\<^esub> y) \ (x,y) \ carrier rel" + using rec_rng_of_frac_def eq_class_of_rng_of_frac_def set_eq_class_of_rng_of_frac_def + by blast + then obtain x y where "a = (x |\<^bsub>rel\<^esub> y)" and P0:"(x,y) \ carrier rel" + by blast + hence P1: "(numer a, denom a) \(x |\<^bsub>rel\<^esub> y)" + using "1" by blast + thus 2:"a = ((numer a) |\<^bsub>rel\<^esub> (denom a))" + proof- + have P2:"(numer a, denom a) \ carrier rel \ (x, y).=\<^bsub>rel\<^esub> (numer a, denom a) " + using eq_class_of_rng_of_frac_def P1 by fastforce + hence "(x, y).=\<^bsub>rel\<^esub> (numer a, denom a)" + by blast + hence "(numer a, denom a).=\<^bsub>rel\<^esub>(x, y)" + using equiv_obj_rng_of_frac by (simp add: equivalence.sym P0 P2) + thus ?thesis + by (metis P0 P2 \a = (x |\<^bsub>rel\<^esub> y)\ class_of_to_rel elem_eq_class equiv_obj_rng_of_frac) + qed + have 3:"(numer a, denom a) \ carrier rel" + using P1 by (simp add: eq_class_of_rng_of_frac_def) + thus 4: "numer a \ carrier R" + by (simp add: rel_def) + show 5: "denom a \ S" + using "3" rel_def by auto + show "numer a = \ \ a = \\<^bsub>rec_rng_of_frac\<^esub>" + proof- + assume "numer a = \" + hence "a = (\ |\<^bsub>rel\<^esub> (denom a))" + using "2" by auto + thus ?thesis + using "5" class_of_zero_rng_of_frac by auto + qed + show "a \\<^bsub>rec_rng_of_frac\<^esub> rng_to_rng_of_frac (denom a) = rng_to_rng_of_frac (numer a)" + proof- + have S: "(denom a, \) \carrier rel" + using "5" rel_def subset by auto + hence "a \\<^bsub>rec_rng_of_frac\<^esub> rng_to_rng_of_frac (denom a) = (((numer a) \ (denom a)) |\<^bsub>rel\<^esub> ((denom a) \ \)) " + using 2 3 mult_rng_of_frac_fundamental_lemma rng_to_rng_of_frac_def + rec_monoid_rng_of_frac_def rec_rng_of_frac_def by fastforce + hence "a \\<^bsub>rec_rng_of_frac\<^esub> rng_to_rng_of_frac (denom a) = (((denom a)\ (numer a)) |\<^bsub>rel\<^esub> ((denom a) \ \))" + using "4" "5" m_comm subset by auto + hence "a \\<^bsub>rec_rng_of_frac\<^esub> rng_to_rng_of_frac (denom a) = ((denom a) |\<^bsub>rel\<^esub> (denom a)) \\<^bsub>rec_rng_of_frac\<^esub>( (numer a) |\<^bsub>rel\<^esub> \)" + using mult_rng_of_frac_fundamental_lemma "4" "5" S + rec_monoid_rng_of_frac_def rec_rng_of_frac_def rel_def by auto + thus ?thesis + using "4" "5" \a \\<^bsub>rec_rng_of_frac\<^esub> rng_to_rng_of_frac + (denom a) = (denom a \ numer a |\<^bsub>rel\<^esub> denom a \ \)\ rel_def + rng_to_rng_of_frac_def simp_in_frac by auto + qed + thus "(rng_to_rng_of_frac(denom a)) \\<^bsub>rec_rng_of_frac\<^esub> a = rng_to_rng_of_frac (numer a)" + by (smt "5" assms(3) cring.cring_simprules(14) crng_rng_of_frac ring_hom_closed rng_to_rng_of_frac_is_ring_hom subset subsetD) +qed + +lemma(in eq_obj_rng_of_frac) frac_zero: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier R" + assumes "b \ S" + assumes "(a |\<^bsub>rel\<^esub> b) = \\<^bsub>rec_rng_of_frac\<^esub>" + shows "a = \" +proof- + have 0: "(a |\<^bsub>rel\<^esub> b) = (\ |\<^bsub>rel\<^esub> \)" + by (simp add: assms(5) class_of_zero_rng_of_frac) + have 1: "(a,b) \ carrier rel" + by (simp add: assms(3) assms(4) rel_def) + have 2: "(\, \) \ carrier rel" + by (simp add: rel_def) + have 3: "(b, \) \ carrier rel" + using assms(4) rel_def subset by auto + have "(a,b) .=\<^bsub>rel\<^esub> (\, \)" + by (metis (no_types, lifting) "0" "1" "2" eq_class_to_rel partial_object.select_convs(1) rel_def) + have "(a |\<^bsub>rel\<^esub> b) \\<^bsub>rec_rng_of_frac\<^esub> (b |\<^bsub>rel\<^esub>\) = (\ |\<^bsub>rel\<^esub> b)" + by (metis (no_types, opaque_lifting) assms(4) assms(5) + basic_trans_rules(31) class_of_zero_rng_of_frac cring.axioms(1) + crng_rng_of_frac ring.ring_simprules(24) ring_hom_closed + rng_to_rng_of_frac_def rng_to_rng_of_frac_is_ring_hom subset) + hence 4: "((a \ b) |\<^bsub>rel\<^esub> b) = (\ |\<^bsub>rel\<^esub> b)" + using "1" "3" assms(4) mult_rng_of_frac_fundamental_lemma + rec_monoid_rng_of_frac_def rec_rng_of_frac_def subset by auto + have S: "((a \ b) , b) \ carrier rel" + using assms(3) assms(4) rel_def subset by auto + have T: "(\, b) \carrier rel" + by (simp add: assms(4) rel_def) + hence " ((a \ b) , b).=\<^bsub>rel\<^esub> (\ , b)" + using 4 S eq_class_to_rel rel_def by auto + hence "eq rel ((a \ b) , b) (\ , b)" + by blast + hence "\t\S. t \ (b \ (a \ b) \ b \ \) = \" + using rel_def by auto + then obtain t where "t \ S" and "t \ (b \ (a \ b) \ b \ \) = \" + by blast + have "t \\" + using \t \ S\ assms(2) by blast + hence "(b \ (a \ b) \ b \ \) = \" + by (meson \t \ S\ \t \ (b \ (a \ b) \ b \ \) = \\ assms(1) assms(3) + assms(4) domain.integral_iff minus_closed semiring_simprules(3) + set_mp subset zero_closed) + hence "b \ (a \ b) = \" + using "3" S rel_def abelian_group.minus_to_eq is_abelian_group by fastforce + thus "a = \" + by (metis (no_types, lifting) assms(1) assms(2) + assms(3) assms(4) domain.integral_iff + semiring_simprules(3) subset subsetCE) +qed + +text\When S does not contain 0, and R is a domain, the localization is a domain.\ + +lemma(in eq_obj_rng_of_frac) rec_rng_of_frac_is_domain: + assumes "domain R" + assumes "\ \ S" + shows "domain rec_rng_of_frac" +proof(rule domainI) + show "cring rec_rng_of_frac" + by (simp add: crng_rng_of_frac) + show "\\<^bsub>rec_rng_of_frac\<^esub> \ \\<^bsub>rec_rng_of_frac\<^esub>" + proof- + have " \\<^bsub>R\<^esub> \ \\<^bsub>R\<^esub>" + by (simp add: assms domain.one_not_zero) + hence 0: " \\<^bsub>R\<^esub> \ (a_kernel R rec_rng_of_frac rng_to_rng_of_frac)" + using assms(1) rng_to_rng_of_frac_without_zero_div_is_inj + by (simp add: assms(2) domain_axioms_def domain_def) + hence "rng_to_rng_of_frac \ \ \\<^bsub>rec_rng_of_frac\<^esub>" + by (simp add: a_kernel_def') + thus ?thesis + using ring_hom_one rng_to_rng_of_frac_is_ring_hom by blast + qed + show "\a b. a \\<^bsub>rec_rng_of_frac\<^esub> b = \\<^bsub>rec_rng_of_frac\<^esub> \ + a \ carrier rec_rng_of_frac \ + b \ carrier rec_rng_of_frac \ + a = \\<^bsub>rec_rng_of_frac\<^esub> \ b = \\<^bsub>rec_rng_of_frac\<^esub>" + proof- + fix a b + assume A1: "a \\<^bsub>rec_rng_of_frac\<^esub> b = \\<^bsub>rec_rng_of_frac\<^esub>" + assume A2: " a \ carrier rec_rng_of_frac" + assume A3: "b \ carrier rec_rng_of_frac" + show "a = \\<^bsub>rec_rng_of_frac\<^esub> \ b = \\<^bsub>rec_rng_of_frac\<^esub>" + proof(rule disjCI) + assume "b \ \\<^bsub>rec_rng_of_frac\<^esub>" + have "\ a \ \\<^bsub>rec_rng_of_frac\<^esub> " + proof + assume "a \ \\<^bsub>rec_rng_of_frac\<^esub>" + have B1: "numer a \ \" + using A2 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) numer_denom_facts0(4) by blast + have B2: "numer b \ \" + using A3 \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) numer_denom_facts0(4) by blast + have B3: "denom a \\" + using A2 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) + eq_obj_rng_of_frac.numer_denom_facts0(1) eq_obj_rng_of_frac_axioms + using numer_denom_facts0(3) by force + have B4: "denom b \\" + using A3 \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) + eq_obj_rng_of_frac_axioms by (metis (no_types, lifting) numer_denom_facts0(3)) + have "a \\<^bsub>rec_rng_of_frac\<^esub> b = (((numer a) \ (numer b)) |\<^bsub>rel\<^esub> ((denom a) \ (denom b)))" + proof- + have S0: "a = (numer a |\<^bsub>rel\<^esub> denom a)" + by (simp add: A2 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) numer_denom_facts0(1)) + have S1: "b= (numer b |\<^bsub>rel\<^esub> denom b)" + by (simp add: A3 \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) numer_denom_facts0(1)) + have S2: "(numer a, denom a) \ carrier rel" + using A2 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) + eq_obj_rng_of_frac.numer_denom_facts0(2) eq_obj_rng_of_frac.numer_denom_facts0(3) + eq_obj_rng_of_frac_axioms rel_def by fastforce + have S3: "(numer b, denom b) \ carrier rel" + using A3 \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) + eq_obj_rng_of_frac.numer_denom_facts0(2) eq_obj_rng_of_frac_axioms + numer_denom_facts0(3) rel_def by auto + show ?thesis using S0 S1 S2 S3 mult_rng_of_frac_fundamental_lemma + using rec_monoid_rng_of_frac_def rec_rng_of_frac_def by force + qed + hence "(((numer a) \ (numer b)) |\<^bsub>rel\<^esub> ((denom a) \ (denom b))) = \\<^bsub>rec_rng_of_frac\<^esub>" + using A1 by blast + hence "(numer a) \ (numer b) = \" + by (meson A2 A3 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ + assms(1) assms(2) eq_obj_rng_of_frac.numer_denom_facts0(2) + eq_obj_rng_of_frac.numer_denom_facts0(3) eq_obj_rng_of_frac_axioms + frac_zero m_closed semiring_simprules(3)) + thus False + by (meson A2 A3 B1 B2 \a \ \\<^bsub>rec_rng_of_frac\<^esub>\ + \b \ \\<^bsub>rec_rng_of_frac\<^esub>\ assms(1) assms(2) + domain.integral_iff eq_obj_rng_of_frac.numer_denom_facts0(2) + eq_obj_rng_of_frac_axioms) + qed + thus "a = \\<^bsub>rec_rng_of_frac\<^esub>" + by auto + qed + qed +qed + +lemma(in eq_obj_rng_of_frac) numer_denom_facts: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier rec_rng_of_frac" + shows "a = (numer a |\<^bsub>rel\<^esub> denom a)" + "(numer a) \ carrier R" + "(denom a) \ S" + "a \ \\<^bsub>rec_rng_of_frac\<^esub> \ (numer a) \\" + "a \\<^bsub>rec_rng_of_frac\<^esub> (rng_to_rng_of_frac(denom a)) = rng_to_rng_of_frac (numer a)" + "(rng_to_rng_of_frac(denom a)) \\<^bsub>rec_rng_of_frac\<^esub> a = rng_to_rng_of_frac (numer a)" + using assms(1) assms(2) assms(3) class_of_zero_rng_of_frac denom_def numer_def numer_denom_facts0(1) apply force + using assms(1) assms(2) assms(3) numer_def numer_denom_facts0(2) apply force + using assms(1) assms(2) assms(3) denom_def numer_denom_facts0(3) apply force + using assms(1) assms(2) assms(3) numer_denom_facts0(4) apply blast + apply (metis (no_types, lifting) assms(1) assms(2) assms(3) class_of_zero_rng_of_frac + denom_def monoid.r_one monoid.simps(2) numer_def numer_denom_facts0(5) one_closed + rec_rng_of_frac_def ringE(2) rng_rng_of_frac rng_to_rng_of_frac_def) + by (metis (no_types, lifting) assms(1) assms(2) assms(3) class_of_zero_rng_of_frac + cring.cring_simprules(12) crng_rng_of_frac denom_def monoid.simps(2) numer_def + numer_denom_facts0(6) one_closed rec_rng_of_frac_def rng_to_rng_of_frac_def) + +lemma(in eq_obj_rng_of_frac) numer_denom_closed: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier rec_rng_of_frac" + shows "(numer a, denom a) \ carrier rel" + by (simp add: assms(1) assms(2) assms(3) numer_denom_facts(2) numer_denom_facts(3) rel_def) + +text\Fraction function which suppresses the "rel" argument.\ + +definition(in eq_obj_rng_of_frac) fraction where +"fraction x y = (x |\<^bsub>rel\<^esub> y)" + +lemma(in eq_obj_rng_of_frac) a_is_fraction: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier rec_rng_of_frac" + shows "a = fraction (numer a) (denom a)" + by (simp add: assms(1) assms(2) assms(3) fraction_def numer_denom_facts(1)) + +lemma(in eq_obj_rng_of_frac) add_fraction: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier R" + assumes "b \ S" + assumes "c \ carrier R" + assumes "d \ S" + shows "(fraction a b) \\<^bsub>rec_rng_of_frac\<^esub> (fraction c d) = (fraction ((a \ d) \ (b \ c)) (b \ d))" +proof- + have 0:"(a,b) \ carrier rel" + by (simp add: assms(3) assms(4) rel_def) + have 1:"(c,d) \ carrier rel" + by (simp add: assms(5) assms(6) rel_def) + show ?thesis + using 0 1 add_rng_of_frac_fundamental_lemma assms numer_denom_facts fraction_def + domain_def m_comm subset + by auto +qed + +lemma(in eq_obj_rng_of_frac) mult_fraction: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier R" + assumes "b \ S" + assumes "c \ carrier R" + assumes "d \ S" + shows "(fraction a b) \\<^bsub>rec_rng_of_frac\<^esub> (fraction c d) = (fraction (a \ c) (b \ d))" +proof- + have 0:"(a,b) \ carrier rel" + by (simp add: assms(3) assms(4) rel_def) + have 1:"(c,d) \ carrier rel" + by (simp add: assms(5) assms(6) rel_def) + show ?thesis using 0 1 mult_rng_of_frac_fundamental_lemma + by (simp add: fraction_def rec_monoid_rng_of_frac_def rec_rng_of_frac_def) +qed + +lemma(in eq_obj_rng_of_frac) fraction_zero: +"\\<^bsub>rec_rng_of_frac\<^esub> = fraction \ \ " + by (simp add: class_of_zero_rng_of_frac fraction_def) + +lemma(in eq_obj_rng_of_frac) fraction_zero': + assumes "a \ S" + assumes "\ \ S" + shows "\\<^bsub>rec_rng_of_frac\<^esub> = fraction \ a" + by (simp add: assms(1) class_of_zero_rng_of_frac fraction_def) + +lemma(in eq_obj_rng_of_frac) fraction_one: +"\\<^bsub>rec_rng_of_frac\<^esub> = fraction \ \" + by (simp add: fraction_def rec_rng_of_frac_def) + +lemma(in eq_obj_rng_of_frac) fraction_one': + assumes "domain R" + assumes "\ \ S" + assumes "a \ S" + shows "fraction a a = \\<^bsub>rec_rng_of_frac\<^esub>" + using assms fraction_def fraction_one one_closed simp_in_frac subset + by (smt mem_Sigma_iff partial_object.select_convs(1) r_one rel_def subsetD) + +lemma(in eq_obj_rng_of_frac) fraction_closed: + assumes "domain R" + assumes "\ \ S" + assumes "a \ carrier R" + assumes "b \ S" + shows "fraction a b \ carrier rec_rng_of_frac" +proof- + have "(a,b) \ carrier rel" + by (simp add: assms(3) assms(4) rel_def) + hence "(a |\<^bsub>rel\<^esub> b) \ set_class_of\<^bsub>rel\<^esub>" + using set_eq_class_of_rng_of_frac_def + by blast + thus ?thesis using fraction_def + by (simp add: rec_rng_of_frac_def) +qed + +subsection\Defining the Field of Fractions\ + +definition Frac where +"Frac R = eq_obj_rng_of_frac.rec_rng_of_frac R (nonzero R)" + +lemma(in domain_frac) fraction_field_is_domain: +"domain (Frac R)" + using domain_axioms eq_obj_rng_of_frac.rec_rng_of_frac_is_domain + eq_obj_rng_of_frac_nonzero Frac_def + by (metis nonzero_memE(2)) + +subsubsection\Numerator and Denominator Choice Functions for \texttt{domain\_frac}\ +definition numerator where +"numerator R = eq_obj_rng_of_frac.numer R (nonzero R)" + +abbreviation(in domain_frac)(input) numer where +"numer \ numerator R" + +definition denominator where +"denominator R = eq_obj_rng_of_frac.denom R (nonzero R)" + +abbreviation(in domain_frac)(input) denom where +"denom \ denominator R" + +definition fraction where +"fraction R = eq_obj_rng_of_frac.fraction R (nonzero R)" + +abbreviation(in domain_frac)(input) frac where +"frac \ fraction R" + +subsubsection\The inclusion of R into its fraction field\ + +definition Frac_inc where +"Frac_inc R = eq_obj_rng_of_frac.rng_to_rng_of_frac R (nonzero R)" + +abbreviation(in domain_frac)(input) inc ("\") where +"inc \ Frac_inc R" + +lemma(in domain_frac) inc_equation: + assumes "a \ carrier R" + shows "\ a = frac a \" + unfolding Frac_inc_def fraction_def + using assms + by (simp add: eq_obj_rng_of_frac.fraction_def eq_obj_rng_of_frac.rng_to_rng_of_frac_def eq_obj_rng_of_frac_nonzero) + +lemma(in domain_frac) inc_is_hom: +"inc \ ring_hom R (Frac R)" + by (simp add: eq_obj_rng_of_frac.rng_to_rng_of_frac_is_ring_hom Frac_def + eq_obj_rng_of_frac_nonzero Frac_inc_def) + +lemma(in domain_frac) inc_is_hom1: +"ring_hom_ring R (Frac R) inc" + apply(rule ring_hom_ringI2) + using cring_def domain.axioms(1) fraction_field_is_domain + by(auto simp:inc_is_hom local.ring_axioms) + +text\Inclusion map is injective:\ + +lemma(in domain_frac) inc_inj0: +"a_kernel R (Frac R) inc = {\}" +proof- + have 0: "\ \ nonzero R" + by (simp add: nonzero_def) + have 1: "eq_obj_rng_of_frac R (nonzero R)" + by (simp add: eq_obj_rng_of_frac_nonzero) + have 2: "\ a\ carrier R. \ b\carrier R. a \ b = \ \ a = \ \ b = \" + using local.integral by blast + show ?thesis using 0 1 2 + by (simp add: eq_obj_rng_of_frac.rng_to_rng_of_frac_without_zero_div_is_inj + Frac_def Frac_inc_def) +qed + +lemma(in domain_frac) inc_inj1: + assumes "a \ carrier R" + assumes "inc a = \\<^bsub>Frac R\<^esub>" + shows "a = \" +proof- + have "a \ a_kernel R (Frac R) inc" using a_kernel_def' assms(2) + by (simp add: a_kernel_def' assms(1)) + thus ?thesis + using inc_inj0 by blast +qed + +lemma(in domain_frac) inc_inj2: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "inc a = inc b" + shows "a = b" +proof- + have "inc (a \ b) = (inc a) \\<^bsub>Frac R\<^esub> (inc (\ b))" + using inc_is_hom by (simp add: ring_hom_add assms(1) assms(2) minus_eq) + hence "inc (a \ b) = \\<^bsub>Frac R\<^esub>" using assms inc_is_hom + by (smt Frac_def add.inv_closed eq_obj_rng_of_frac.rng_rng_of_frac + eq_obj_rng_of_frac_nonzero local.ring_axioms r_neg ring_hom_add ring_hom_zero) + thus ?thesis + by (meson abelian_group.minus_to_eq assms(1) assms(2) domain_frac.inc_inj1 domain_frac_axioms is_abelian_group minus_closed) +qed + +text\Image of inclusion map is a subring:\ + +lemma(in domain_frac) inc_im_is_subring: +"subring (\ ` (carrier R)) (Frac R)" + using carrier_is_subring inc_is_hom1 ring_hom_ring.img_is_subring by blast + +subsubsection\Basic Properties of numer, denom, and frac\ + +lemma(in domain_frac) numer_denom_facts: + assumes "a \ carrier (Frac R)" + shows "a = frac (numer a) (denom a)" + "(numer a) \ carrier R" + "(denom a) \ nonzero R" + "a \ \\<^bsub>Frac R\<^esub> \ numer a \ \ " + "a \\<^bsub>Frac R\<^esub> (inc (denom a)) = inc (numer a)" + unfolding fraction_def numerator_def denominator_def Frac_inc_def + apply (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.a_is_fraction eq_obj_rng_of_frac_nonzero not_nonzero_memI) + apply (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.numer_denom_facts(2) eq_obj_rng_of_frac_nonzero nonzero_memE(2)) + apply (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.numer_denom_facts(3) eq_obj_rng_of_frac_nonzero not_nonzero_memI) + apply (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.numer_denom_facts0(4) eq_obj_rng_of_frac_nonzero not_nonzero_memI) + by (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.numer_denom_facts(5) eq_obj_rng_of_frac_nonzero nonzero_memE(2)) + +lemma(in domain_frac) frac_add: + assumes "a \ carrier R" + assumes "b \ nonzero R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + shows "(frac a b) \\<^bsub>Frac R\<^esub> (frac c d) = (frac ((a \ d) \ (b \ c)) (b \ d))" + using eq_obj_rng_of_frac.add_fraction[of R "nonzero R" a b c d] + eq_obj_rng_of_frac_nonzero assms zero_not_in_nonzero[of R] + by (simp add: Frac_def domain_axioms fraction_def) + +lemma(in domain_frac) frac_mult: + assumes "a \ carrier R" + assumes "b \ nonzero R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + shows "(frac a b) \\<^bsub>Frac R\<^esub> (frac c d) = (frac (a \ c) (b \ d))" + by (simp add: Frac_def assms(1) assms(2) assms(3) assms(4) domain_axioms + eq_obj_rng_of_frac.mult_fraction eq_obj_rng_of_frac_nonzero fraction_def not_nonzero_memI) + +lemma(in domain_frac) frac_one: + assumes "a \ nonzero R" + shows "frac a a = \\<^bsub>Frac R\<^esub>" + by (metis Frac_def assms domain_axioms eq_obj_rng_of_frac.fraction_one' eq_obj_rng_of_frac_nonzero fraction_def nonzero_memE(2)) + +lemma(in domain_frac) frac_closed: + assumes "a \ carrier R" + assumes "b \ nonzero R" + shows "frac a b \ carrier (Frac R)" + by (metis Frac_def assms(1) assms(2) domain_axioms eq_obj_rng_of_frac.fraction_closed eq_obj_rng_of_frac_nonzero fraction_def nonzero_memE(2)) + +lemma(in domain_frac) nonzero_fraction: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "frac a b \ \\<^bsub>Frac R\<^esub>" +proof + assume "frac a b = \\<^bsub>Frac R\<^esub>" + hence "(frac a b) \\<^bsub>Frac R\<^esub> (frac b a) = \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> (frac b a)" + by simp + hence "(frac a b) \\<^bsub>Frac R\<^esub> (frac b a) = \\<^bsub>Frac R\<^esub>" + by (metis Localization.submonoid.subset assms(1) assms(2) cring.cring_simprules(26) + domain.axioms(1) frac_closed fraction_field_is_domain nonzero_is_submonoid subsetCE) + hence "frac (a \b) (b \ a) = \\<^bsub>Frac R\<^esub>" + by (metis (no_types, lifting) Localization.submonoid.subset + assms(1) assms(2) frac_mult nonzero_is_submonoid subsetCE) + hence "frac (a \b) (a \ b) = \\<^bsub>Frac R\<^esub>" + by (metis (no_types, lifting) Localization.submonoid.subset assms(1) assms(2) m_comm nonzero_is_submonoid subsetCE) + hence "\\<^bsub>Frac R\<^esub> = \\<^bsub>Frac R\<^esub>" + using Localization.submonoid.m_closed assms(1) assms(2) frac_one nonzero_is_submonoid by force + thus False + using domain.one_not_zero fraction_field_is_domain by blast +qed + +lemma(in comm_monoid) UnitsI: + assumes "a \ carrier G" + assumes "b \ carrier G" + assumes "a \ b = \" + shows "a \ Units G" "b \ Units G" + unfolding Units_def using comm_monoid_axioms_def assms m_comm[of a b] + by auto + +lemma(in comm_monoid) is_invI: + assumes "a \ carrier G" + assumes "b \ carrier G" + assumes "a \ b = \" + shows "inv\<^bsub>G\<^esub> b = a" "inv\<^bsub>G\<^esub> a = b" + using assms inv_char m_comm + by auto + +lemma(in ring) ring_in_Units_imp_not_zero: + assumes "\ \ \" + assumes "a \ Units R" + shows "a \ \" + using assms monoid.Units_l_cancel + by (metis l_null monoid_axioms one_closed zero_closed) + +lemma(in domain_frac) in_Units_imp_not_zero: + assumes "a \ Units R" + shows "a \ \" + using assms ring_in_Units_imp_not_zero domain_axioms + by simp + +lemma(in domain_frac) units_of_fraction_field0: + assumes "a \ carrier (Frac R)" + assumes "a \ \\<^bsub>Frac R\<^esub>" + shows "inv\<^bsub>Frac R\<^esub> a = frac (denom a) (numer a)" + "a\\<^bsub>Frac R\<^esub> (inv\<^bsub>Frac R\<^esub> a) = \\<^bsub>Frac R\<^esub>" + "(inv\<^bsub>Frac R\<^esub> a)\\<^bsub>Frac R\<^esub>a = \\<^bsub>Frac R\<^esub>" +proof- + have 0: "a \\<^bsub>Frac R\<^esub> (frac (denom a) (numer a)) = + frac ((numer a) \ (denom a)) ((numer a) \ (denom a))" + proof - + have "denom a \ nonzero R" + by (simp add: assms(1) numer_denom_facts(3)) + hence "frac (numer a) (denom a) \\<^bsub>Frac R\<^esub> frac (denom a) (numer a) + = frac (numer a \ denom a) (denom a \ numer a)" + by (simp add: assms(1) assms(2) domain_frac.numer_denom_facts(2) domain_frac_axioms frac_mult nonzero_closed nonzero_memI numer_denom_facts(4)) + thus ?thesis + using assms(1) numer_denom_facts(5) domain_frac.numer_denom_facts(2) + domain_axioms m_comm nonzero_closed numer_denom_facts(1) + by (simp add: domain_frac.numer_denom_facts(2) \denominator R a \ nonzero R\ domain_frac_axioms) + qed + have 1:"((numer a) \ (denom a)) \ nonzero R" + by (metis assms(1) assms(2) domain_frac.numer_denom_facts(2) domain_frac_axioms + local.integral m_closed nonzero_closed nonzero_memI numer_denom_facts(3) numer_denom_facts(4)) + have 2: "a \\<^bsub>Frac R\<^esub> (frac (denom a) (numer a)) = \\<^bsub>Frac R\<^esub>" + using 0 1 frac_one by blast + have 3: "(frac (denom a) (numer a)) \ carrier (Frac R)" + by (simp add: assms(1) assms(2) frac_closed nonzero_closed nonzero_memI numer_denom_facts(2) numer_denom_facts(3) numer_denom_facts(4)) + hence 4: "(frac (denom a) (numer a)) \ carrier (Frac R) \ + a \\<^bsub>Frac R\<^esub> (frac (denom a) (numer a)) = \\<^bsub>Frac R\<^esub> \ + (frac (denom a) (numer a)) \\<^bsub>Frac R\<^esub> a = \\<^bsub>Frac R\<^esub>" + by (simp add: "2" assms(1) cring.cring_simprules(14) domain.axioms(1) fraction_field_is_domain) + thus 5: "inv\<^bsub>Frac R\<^esub> a = frac (denom a) (numer a)" + using m_inv_def 2 assms(1) comm_monoid.comm_inv_char cring_def + domain_def fraction_field_is_domain by fastforce + thus 6: "a\\<^bsub>Frac R\<^esub> (inv\<^bsub>Frac R\<^esub> a) = \\<^bsub>Frac R\<^esub>" + by (simp add: "2") + thus "(inv\<^bsub>Frac R\<^esub> a)\\<^bsub>Frac R\<^esub>a = \\<^bsub>Frac R\<^esub>" + using assms + by (simp add: "4" "5") +qed + +lemma(in domain_frac) units_of_fraction_field: +"Units (Frac R) = carrier (Frac R) - {\\<^bsub>Frac R\<^esub>}" +proof + show "Units (Frac R) \ carrier (Frac R) - {\\<^bsub>Frac R\<^esub>}" + proof fix x assume A: "x \ Units (Frac R)" + have "x \ carrier (Frac R)" + using Units_def \x \ Units (Frac R)\ by force + hence "x \ \\<^bsub>Frac R\<^esub>" + using fraction_field_is_domain + by (simp add: A domain_frac.in_Units_imp_not_zero domain_frac.intro) + thus "x \ carrier (Frac R) - {\\<^bsub>Frac R\<^esub>}" + by (simp add: \x \ carrier (Frac R)\) + qed + show "carrier (Frac R) - {\\<^bsub>Frac R\<^esub>} \ Units (Frac R)" + proof fix x assume A: "x \ carrier (Frac R) - {\\<^bsub>Frac R\<^esub>}" + show "x \ Units (Frac R)" + using comm_monoid.UnitsI(1)[of "Frac R" x "inv\<^bsub>Frac R\<^esub> x"] + by (metis A Diff_iff cring.axioms(2) domain.axioms(1) domain_frac.numer_denom_facts(2) + domain_frac.numer_denom_facts(3) domain_frac.units_of_fraction_field0(1) + domain_frac.units_of_fraction_field0(2) domain_frac_axioms frac_closed + fraction_field_is_domain insert_iff nonzero_closed nonzero_memI numer_denom_facts(4)) + qed +qed + +subsection\The Fraction Field as a Field\ + +lemma(in domain_frac) fraction_field_is_field: +"field (Frac R)" +proof(rule Ring.field.intro) + show "domain (Frac R)" by(auto simp: fraction_field_is_domain) + show "field_axioms (Frac R)" + proof + show "Units (Frac R) = carrier (Frac R) - {\\<^bsub>Frac R\<^esub>}" + using units_of_fraction_field by blast + qed +qed + +lemma(in domain_frac) frac_inv: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "inv\<^bsub>Frac R\<^esub> (frac a b) = (frac b a)" + using cring_def[of "Frac R"] domain_def[of "Frac R"] fraction_field_is_domain + frac_closed[of a b] frac_closed[of b a] nonzero_closed[of a] + nonzero_closed[of b] assms comm_monoid.is_invI(2)[of "Frac R" "frac a b" "frac b a"] + by (metis frac_mult frac_one integral_iff m_comm nonzero_memE(2) nonzero_memI nonzero_mult_closed) + +lemma(in domain_frac) frac_inv_id: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + assumes "c \ nonzero R" + assumes "d \ nonzero R" + assumes "frac a b = frac c d" + shows "frac b a = frac d c" + using frac_inv assms by metis + +lemma(in domain_frac) frac_uminus: + assumes "a \ carrier R" + assumes "b \ nonzero R" + shows "\\<^bsub>Frac R\<^esub> (frac a b) = frac (\ a) b" +proof- + have "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = frac (((\ a)\b) \ (a \b)) (b\b)" + using frac_add by (smt Localization.submonoid.subset add.inv_closed + assms(1) assms(2) m_comm nonzero_is_submonoid subsetCE) + hence "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = frac (b \((\ a) \ a)) (b\b)" + by (metis (no_types, lifting) add.inv_closed assms(1) assms(2) + local.ring_axioms m_comm mem_Collect_eq nonzero_def ringE(4) ) + hence "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = (frac \ (b \b))" + using Localization.submonoid.subset assms(1) assms(2) l_neg nonzero_is_submonoid by fastforce + hence "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = (frac \ \) \\<^bsub>Frac R\<^esub> (frac \ (b \b))" + using frac_mult by (smt Localization.submonoid.m_closed Localization.submonoid.one_closed + Localization.submonoid.subset assms(2) l_one nonzero_is_submonoid r_null subsetCE zero_closed) + hence "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> (frac \ (b \b))" + using Frac_def eq_obj_rng_of_frac.fraction_zero' eq_obj_rng_of_frac_nonzero + by (simp add: Frac_def eq_obj_rng_of_frac.fraction_zero fraction_def) + hence "frac (\ a) b \\<^bsub>Frac R\<^esub> (frac a b) = \\<^bsub>Frac R\<^esub>" + using fraction_field_is_field + by (simp add: Localization.submonoid.m_closed assms(2) cring.cring_simprules(26) + domain.axioms(1) frac_closed fraction_field_is_domain nonzero_is_submonoid) + thus 0: "\\<^bsub>Frac R\<^esub> (frac a b) = frac (\ a) b" + by (metis add.inv_closed assms(1) assms(2) cring.sum_zero_eq_neg + domain.axioms(1) frac_closed fraction_field_is_domain) +qed + +lemma(in domain_frac) frac_eqI: + assumes "a \ carrier R" + assumes "b \ nonzero R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + assumes "a \ d \ b \ c = \" + shows "frac a b = frac c d" +proof- + have "frac a b \\<^bsub>Frac R\<^esub> frac c d = frac (a \ d \ b \ c) (b\d)" + by (simp add: a_minus_def assms(1) assms(2) assms(3) assms(4) frac_uminus local.frac_add nonzero_closed r_minus) + hence "frac a b \\<^bsub>Frac R\<^esub> frac c d = \\<^bsub>Frac R\<^esub>" + by (metis Frac_def assms(2) assms(4) assms(5) eq_obj_rng_of_frac.fraction_zero' + eq_obj_rng_of_frac_nonzero fraction_def local.integral nonzero_memE(1) nonzero_memE(2) + nonzero_memI nonzero_mult_closed) + thus ?thesis + by (meson assms(1) assms(2) assms(3) assms(4) field.is_ring frac_closed fraction_field_is_field ring.r_right_minus_eq) +qed + +lemma(in domain_frac) frac_eqI': + assumes "a \ carrier R" + assumes "b \ nonzero R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + assumes "a \ d = b \ c" + shows "frac a b = frac c d" + using assms frac_eqI[of a b c d] + by (simp add: nonzero_closed) + +lemma(in domain_frac) frac_cancel_l: + assumes "a \nonzero R" + assumes "b \nonzero R" + assumes "c \carrier R" + shows "frac (a \ c) (a \ b) = frac c b" +proof- + have 0: "frac (a \c) (a \b) =(frac b b) \\<^bsub>Frac R\<^esub> (frac c b)" + by (metis (no_types, lifting) assms(1) assms(2) assms(3) + frac_mult frac_one mem_Collect_eq nonzero_def) + have 1: "frac b b = \\<^bsub>Frac R\<^esub>" + by (simp add: assms(2) frac_one) + show ?thesis using 0 1 + using Frac_def assms(2) assms(3) eq_obj_rng_of_frac.rng_rng_of_frac + eq_obj_rng_of_frac_nonzero frac_closed ring.ring_simprules(12) + by (simp add: Frac_def eq_obj_rng_of_frac.rng_rng_of_frac ring.ring_simprules(12)) +qed + +lemma(in domain_frac) frac_cancel_r: + assumes "a \nonzero R" + assumes "b \nonzero R" + assumes "c \carrier R" + shows "frac (c \ a) (b \ a) = frac c b" + by (metis (no_types, lifting) Localization.submonoid.subset assms(1) + assms(2) assms(3) frac_cancel_l m_comm nonzero_is_submonoid subsetCE) + +lemma(in domain_frac) frac_cancel_lr: + assumes "a \nonzero R" + assumes "b \nonzero R" + assumes "c \carrier R" + shows "frac (a \ c) (b \ a) = frac c b" + by (metis (no_types, lifting) Localization.submonoid.subset assms(1) + assms(2) assms(3) frac_cancel_l m_comm nonzero_is_submonoid subsetCE) + +lemma(in domain_frac) frac_cancel_rl: + assumes "a \nonzero R" + assumes "b \nonzero R" + assumes "c \carrier R" + shows "frac (c \ a) (a \ b) = frac c b" + by (metis (no_types, lifting) Localization.submonoid.subset assms(1) + assms(2) assms(3) frac_cancel_l m_comm nonzero_is_submonoid subsetCE) + +lemma(in domain_frac) i_mult: + assumes "a \ carrier R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + shows "(\ a) \\<^bsub>Frac R\<^esub> (frac c d) = frac (a \ c) d" +proof- + have "(\ a) \\<^bsub>Frac R\<^esub> (frac c d) = (frac a \) \\<^bsub>Frac R\<^esub> (frac c d)" + by (simp add: assms(1) inc_equation) + thus ?thesis + by (metis (mono_tags, lifting) assms(1) assms(2) assms(3) cring_simprules(12) + cring_simprules(6) frac_mult local.one_not_zero mem_Collect_eq nonzero_def) +qed + +lemma(in domain_frac) frac_eqE: + assumes "a \ carrier R" + assumes "b \ nonzero R" + assumes "c \ carrier R" + assumes "d \ nonzero R" + assumes "frac a b = frac c d" + shows "a \ d = b \ c" +proof- + have "(\ b) \\<^bsub>Frac R\<^esub> (frac a b) = (\ b) \\<^bsub>Frac R\<^esub> (frac c d)" + by (simp add: assms(5)) + hence "(frac (a \ b) b) = (frac (c \ b) d)" + using i_mult + by (metis (no_types, lifting) Localization.submonoid.subset assms(1) + assms(2) assms(3) assms(4) m_comm nonzero_is_submonoid subsetCE) + hence "(frac a \) = (frac (c \ b) d)" + by (smt assms(1) assms(2) frac_cancel_r l_one mem_Collect_eq nonzero_def one_closed zero_not_one) + hence "(\ d) \\<^bsub>Frac R\<^esub>(frac a \) =(\ d) \\<^bsub>Frac R\<^esub> (frac (c \ b) d)" + by auto + hence "(frac (a \ d) \) =(frac ((c \ b)\ d) d)" + using i_mult + by (smt Localization.submonoid.m_closed Localization.submonoid.subset assms(1) assms(2) assms(3) + assms(4) cring_simprules(27) cring_simprules(6) local.one_not_zero m_comm + mem_Collect_eq nonzero_def nonzero_is_submonoid) + hence "(frac (a \ d) \) =(frac (c \ b) \)" + by (smt Localization.submonoid.subset assms(2) assms(3) assms(4) cring_simprules(5) + cring_simprules(6) frac_one i_mult inc_equation inc_is_hom nonzero_is_submonoid + r_one ring_hom_mult ring_hom_one subsetCE) + thus ?thesis using assms + unfolding fraction_def + by (simp add: fraction_def inc_equation inc_inj2 m_comm nonzero_closed) +qed + +lemma(in domain_frac) frac_add_common_denom: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c \ nonzero R" + shows "(frac a c) \\<^bsub>Frac R\<^esub> (frac b c) = frac (a \ b) c" +proof- + have "(frac a c) \\<^bsub>Frac R\<^esub> (frac b c) = frac ((a \ c) \ (b \ c)) (c \ c)" + using assms(1) assms(2) assms(3) domain_frac.frac_add domain_axioms frac_eqE local.frac_add + by auto + hence "(frac a c) \\<^bsub>Frac R\<^esub> (frac b c) = frac ((a \ b) \ c) (c \ c)" + by (metis Localization.submonoid.subset assms(1) assms(2) assms(3) l_distr nonzero_is_submonoid subsetCE) + thus ?thesis + by (simp add: assms(1) assms(2) assms(3) frac_cancel_r) +qed + +lemma(in domain_frac) common_denominator: + assumes "x \ carrier (Frac R)" + assumes "y \ carrier (Frac R)" + obtains a b c where + "a \ carrier R" + "b \ carrier R" + "c \ nonzero R" + "x = frac a c" + "y = frac b c" +proof- + obtain x1 x2 where X1: "x1 \ carrier R" and X2: "x2 \ nonzero R" and Fx: "x = frac x1 x2" + by (meson assms(1) numer_denom_facts(1) numer_denom_facts(2) numer_denom_facts(3)) + obtain y1 y2 where Y1: "y1 \ carrier R" and Y2: "y2 \ nonzero R" and Fy: "y = frac y1 y2" + by (meson assms(2) numer_denom_facts(1) numer_denom_facts(2) numer_denom_facts(3)) + let ?a = "x1 \ y2" + let ?b = "y1 \ x2" + let ?c = "x2 \ y2" + have 0: "?a \ carrier R" + using X1 Y2 by (simp add: nonzero_def) + have 1: "?b \ carrier R" using Y1 X2 + by (simp add: nonzero_def) + have 2: "?c \ nonzero R" using X2 Y2 + by (simp add: Localization.submonoid.m_closed nonzero_is_submonoid) + have 3: "x = frac ?a ?c" + using Fx X1 X2 Y2 frac_cancel_r by auto + have 4: "y = frac ?b ?c" + using Fy X2 Y1 Y2 frac_cancel_rl by auto + thus ?thesis using 0 1 2 3 4 + using that by blast +qed + +sublocale domain_frac < F: field "Frac R" + by (simp add: fraction_field_is_field) + +text\Inclusions of natural numbers into \texttt{(Frac R)}:\ + +lemma(in domain_frac) nat_0: +"[(0::nat)]\\ = \" + by simp + +lemma(in domain_frac) nat_suc: +"[Suc n]\\ = \ \ [n]\\" + using add.nat_pow_Suc2 by auto + +lemma(in domain_frac) nat_inc_rep: + fixes n::nat + shows "[n]\\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = frac ([n]\\) \" +proof(induction n) + case 0 + have "[(0::nat)] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = \\<^bsub>Frac R\<^esub>" + using fraction_field_is_domain + by (simp add: domain_frac.intro domain_frac.nat_0) + thus ?case + by (simp add: Frac_def eq_obj_rng_of_frac.fraction_zero eq_obj_rng_of_frac_nonzero fraction_def) +next + case (Suc n) + assume A: "[n] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = frac ([n] \ \) \" + have "[Suc n] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> [n] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub>" + using F.add.nat_pow_Suc2 by auto + hence "[Suc n] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = (frac \ \) \\<^bsub>Frac R\<^esub> (frac ([n] \ \) \)" + by (simp add: Suc.IH frac_one nonzero_def) + hence "[Suc n] \\<^bsub>Frac R\<^esub> \\<^bsub>Frac R\<^esub> = (frac (\\[n] \ \) \)" + by (simp add: frac_add_common_denom nonzero_def) + thus ?case + using nat_suc by auto +qed + +lemma(in domain_frac) pow_0: + assumes "a \ nonzero R" + shows "a[^](0::nat) = \" + by simp + +lemma(in domain_frac) pow_suc: + assumes "a \ carrier R" + shows "a[^](Suc n) = a \(a[^]n)" + using assms nat_pow_Suc2 by auto + +lemma(in domain_frac) nat_inc_add: +"[((n::nat) + (m::nat))]\\ = [n]\\ \ [m]\\" + using domain_def add_pow_def + by (simp add: add.nat_pow_mult) + +lemma(in domain_frac) int_inc_add: +"[((n::int) + (m::int))]\\ = [n]\\ \ [m]\\" + using domain_def add_pow_def + by (simp add: add.int_pow_mult) + +lemma(in domain_frac) nat_inc_mult: +"[((n::nat) * (m::nat))]\\ = [n]\\ \ [m]\\" + using domain_def add_pow_def + by (simp add: Groups.mult_ac(2) add.nat_pow_pow add_pow_ldistr) + +lemma(in domain_frac) int_inc_mult: +"[((n::int) * (m::int))]\\ = [n]\\ \ [m]\\" + using domain_def add_pow_def + by (simp add: Groups.mult_ac(2) add.int_pow_pow add_pow_ldistr_int) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Facts About Ring Units\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma(in ring) Units_nonzero: + assumes "u \ Units R" + assumes "\\<^bsub>R\<^esub> \ \\<^bsub>R\<^esub>" + shows "u \ nonzero R" +proof- + have "u \carrier R" + using Units_closed assms by auto + have "u \\" + using Units_r_inv_ex assms(1) assms(2) + by force + thus ?thesis + by (simp add: \u \ carrier R\ nonzero_def) +qed + +lemma(in ring) Units_inverse: + assumes "u \ Units R" + shows "inv u \ Units R" + by (simp add: assms) + +lemma(in cring) invI: + assumes "x \ carrier R" + assumes "y \ carrier R" + assumes "x \\<^bsub>R\<^esub> y = \\<^bsub>R\<^esub>" + shows "y = inv \<^bsub>R\<^esub> x" + "x = inv \<^bsub>R\<^esub> y" + using assms(1) assms(2) assms(3) is_invI + by auto + +lemma(in cring) inv_cancelR: + assumes "x \ Units R" + assumes "y \ carrier R" + assumes "z \ carrier R" + assumes "y = x \\<^bsub>R\<^esub> z" + shows "inv\<^bsub>R\<^esub> x \\<^bsub>R\<^esub> y = z" + "y \\<^bsub>R\<^esub> (inv\<^bsub>R\<^esub> x) = z" + apply (metis Units_closed assms(1) assms(3) assms(4) cring.cring_simprules(12) + is_cring m_assoc monoid.Units_inv_closed monoid.Units_l_inv monoid_axioms) + by (metis Units_closed assms(1) assms(3) assms(4) m_assoc m_comm monoid.Units_inv_closed + monoid.Units_r_inv monoid.r_one monoid_axioms) + +lemma(in cring) inv_cancelL: + assumes "x \ Units R" + assumes "y \ carrier R" + assumes "z \ carrier R" + assumes "y = z \\<^bsub>R\<^esub> x" + shows "inv\<^bsub>R\<^esub> x \\<^bsub>R\<^esub> y = z" + "y \\<^bsub>R\<^esub> (inv\<^bsub>R\<^esub> x) = z" + apply (simp add: Units_closed assms(1) assms(3) assms(4) m_lcomm) + by (simp add: Units_closed assms(1) assms(3) assms(4) m_assoc) + +lemma(in domain_frac) nat_pow_nonzero: + assumes "x \nonzero R" + shows "x[^](n::nat) \ nonzero R" + unfolding nonzero_def + apply(induction n) + using assms integral_iff nonzero_closed zero_not_in_nonzero by auto + +lemma(in monoid) Units_int_pow_closed: + assumes "x \ Units G" + shows "x[^](n::int) \ Units G" + by (metis Units_pow_closed assms int_pow_def2 monoid.Units_inv_Units monoid_axioms) + +subsection\Facts About Fraction Field Units\ + +lemma(in domain_frac) frac_nonzero_Units: +"nonzero (Frac R) = Units (Frac R)" + unfolding nonzero_def using F.field_Units + by blast + +lemma(in domain_frac) frac_nonzero_inv_Unit: + assumes "b \ nonzero (Frac R)" + shows "inv\<^bsub>Frac R\<^esub> b \ Units (Frac R)" + using assms frac_nonzero_Units + by simp + +lemma(in domain_frac) frac_nonzero_inv_closed: + assumes "b \ nonzero (Frac R)" + shows "inv\<^bsub>Frac R\<^esub> b \ carrier (Frac R)" + using frac_nonzero_inv_Unit + by (simp add: Units_def assms) + +lemma(in domain_frac) frac_nonzero_inv: + assumes "b \ nonzero (Frac R)" + shows "b \\<^bsub>Frac R\<^esub> inv \<^bsub>Frac R\<^esub> b = \\<^bsub>Frac R\<^esub>" + "inv \<^bsub>Frac R\<^esub> b \\<^bsub>Frac R\<^esub> b = \\<^bsub>Frac R\<^esub>" + using frac_nonzero_Units assms by auto + +lemma(in domain_frac) fract_cancel_right[simp]: + assumes "a \ carrier (Frac R)" + assumes "b \ nonzero (Frac R)" + shows "b \\<^bsub>Frac R\<^esub> (a \\<^bsub>Frac R\<^esub> inv\<^bsub>Frac R\<^esub> b) = a" + by (metis F.Units_inv_inv F.inv_cancelL(1) F.m_closed assms(1) assms(2) frac_nonzero_Units frac_nonzero_inv_Unit frac_nonzero_inv_closed) + +lemma(in domain_frac) fract_cancel_left[simp]: + assumes "a \ carrier (Frac R)" + assumes "b \ nonzero (Frac R)" + shows "(a \\<^bsub>Frac R\<^esub> inv\<^bsub>Frac R\<^esub> b) \\<^bsub>Frac R\<^esub> b = a" + by (metis Diff_iff assms(1) assms(2) cring.cring_simprules(14) cring.cring_simprules(5) + domain.axioms(1) frac_nonzero_Units frac_nonzero_inv_closed fract_cancel_right + fraction_field_is_domain units_of_fraction_field) + +lemma(in domain_frac) fract_mult_inv: + assumes "b \ nonzero (Frac R)" + assumes "d \ nonzero (Frac R)" + shows "(inv\<^bsub>Frac R\<^esub> b) \\<^bsub>Frac R\<^esub> (inv\<^bsub>Frac R\<^esub> d) = (inv\<^bsub>Frac R\<^esub> (b \\<^bsub>Frac R\<^esub>d))" + by (metis F.Units_inv_closed F.Units_m_closed F.inv_cancelR(2) F.nonzero_closed assms(1) assms(2) frac_nonzero_Units) + +lemma(in domain_frac) fract_mult: + assumes "a \ carrier (Frac R)" + assumes "b \ nonzero (Frac R)" + assumes "c \ carrier (Frac R)" + assumes "d \ nonzero (Frac R)" + shows "(a \\<^bsub>Frac R\<^esub> inv\<^bsub>Frac R\<^esub> b) \\<^bsub>Frac R\<^esub> (c \\<^bsub>Frac R\<^esub> inv\<^bsub>Frac R\<^esub> d) = ((a \\<^bsub>Frac R\<^esub> c)\\<^bsub>Frac R\<^esub> inv\<^bsub>Frac R\<^esub> (b \\<^bsub>Frac R\<^esub>d))" + using F.m_assoc F.m_lcomm assms(1) assms(2) assms(3) assms(4) frac_nonzero_Units fract_mult_inv by auto + +lemma(in domain_frac) Frac_nat_pow_nonzero: + assumes "x \ nonzero (Frac R)" + shows "x[^]\<^bsub>Frac R\<^esub>(n::nat) \ nonzero (Frac R)" + by (simp add: assms domain_frac.intro domain_frac.nat_pow_nonzero fraction_field_is_domain) + +lemma(in domain_frac) Frac_nonzero_nat_pow: + assumes "x \ carrier (Frac R)" + assumes "n > 0" + assumes "x[^]\<^bsub>Frac R\<^esub>(n::nat) \ nonzero (Frac R)" + shows "x \ nonzero (Frac R)" +proof(rule ccontr) + assume "x \ nonzero (Frac R)" + hence 0: "x = \\<^bsub>Frac R\<^esub>" + by (simp add: assms(1) nonzero_def) + have "x[^]\<^bsub>Frac R\<^esub>(n::nat) = \\<^bsub>Frac R\<^esub>" + proof- + obtain k where "n = Suc k" + using assms(2) lessE by blast + hence 00: "x[^]\<^bsub>Frac R\<^esub>(n::nat) = x[^]\<^bsub>Frac R\<^esub>k \\<^bsub>Frac R\<^esub> x" + by simp + have "x[^]\<^bsub>Frac R\<^esub>k \ carrier (Frac R)" + using assms monoid.nat_pow_closed[of "Frac R" x k] + by (meson field.is_ring fraction_field_is_field ring_def) + thus ?thesis using 0 assms + using "00" cring.cring_simprules(27) domain.axioms(1) fraction_field_is_domain by fastforce + qed + thus False + using "0" \x \ nonzero (Frac R)\ assms(3) by auto +qed + +lemma(in domain_frac) Frac_int_pow_nonzero: + assumes "x \ nonzero (Frac R)" + shows "x[^]\<^bsub>Frac R\<^esub>(n::int) \ nonzero (Frac R)" + using assms field.is_ring frac_nonzero_Units fraction_field_is_field monoid.Units_pow_closed[of "Frac R" x] + by (simp add: field.is_ring monoid.Units_int_pow_closed ring.is_monoid) + +lemma(in domain_frac) Frac_nonzero_int_pow: + assumes "x \ carrier (Frac R)" + assumes "n > 0" + assumes "x[^]\<^bsub>Frac R\<^esub>(n::int) \ nonzero (Frac R)" + shows "x \ nonzero (Frac R)" + by (metis (mono_tags, opaque_lifting) Frac_nonzero_nat_pow assms int_pow_int pos_int_cases) + +lemma(in domain_frac) numer_denom_frac[simp]: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "frac (numer (frac a b)) (denom (frac a b)) = frac a b" + using assms(1) assms(2) domain_frac.numer_denom_facts(1) +domain_axioms frac_closed nonzero_closed numer_denom_facts(1) by auto + +lemma(in domain_frac) numer_denom_swap: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "a \ (denom (frac a b)) = b \ (numer (frac a b))" + using numer_denom_frac[of a b] assms + by (simp add: frac_closed frac_eqE nonzero_closed numer_denom_facts(2) numer_denom_facts(3)) + +lemma(in domain_frac) numer_nonzero: + assumes "a \ nonzero (Frac R)" + shows "numer a \ nonzero R" + using assms numer_denom_facts(4)[of a] zero_not_in_nonzero[of R] + by (simp add: frac_nonzero_Units nonzero_memI numer_denom_facts(2) units_of_fraction_field) + +lemma(in domain_frac) fraction_zero[simp]: + assumes "b \ nonzero R" + shows "frac \ b = \\<^bsub>Frac R\<^esub>" + by (metis Frac_def assms eq_obj_rng_of_frac.fraction_zero' eq_obj_rng_of_frac_nonzero fraction_def nonzero_memE(2)) + +end diff --git a/thys/Padic_Field/Generated_Boolean_Algebra.thy b/thys/Padic_Field/Generated_Boolean_Algebra.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Generated_Boolean_Algebra.thy @@ -0,0 +1,1388 @@ +theory Generated_Boolean_Algebra + imports Main +begin + +(**************************************************************************************************) +(**************************************************************************************************) +section\Generated Boolean Algebras of Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Definitions and Basic Lemmas\ +(**************************************************************************************************) +(**************************************************************************************************) +lemma equalityI': + assumes "\x. x \ A \ x \ B" + assumes "\x. x \ B \ x \ A" + shows "A = B" + using assms by blast + +lemma equalityI'': + assumes "\x. A x \ B x" + assumes "\x. B x \ A x" + shows "{x. A x} = {x. B x}" + using assms by blast + +lemma SomeE: + assumes "a = (SOME x. P x)" + assumes "P c" + shows "P a" + using assms by (meson verit_sko_ex) + +lemma SomeE': + assumes "a = (SOME x. P x)" + assumes "\ x. P x" + shows "P a" + using assms by (meson verit_sko_ex) + +section\Basic notions about boolean algebras over a set \S\, generated by a set of generators \B\\ + +text\Note that the generators \B\ need not be subsets of the set \S\\ + +inductive_set gen_boolean_algebra + for S and B where + universe: "S \ gen_boolean_algebra S B" + | generator: "A \ B \ A \ S \ gen_boolean_algebra S B" + | union: "\ A \ gen_boolean_algebra S B; C \ gen_boolean_algebra S B\ \ A \ C \ gen_boolean_algebra S B" + | complement: "A \ gen_boolean_algebra S B \ S - A \ gen_boolean_algebra S B" + +lemma gen_boolean_algebra_subset: + shows "A \ gen_boolean_algebra S B \ A \ S" + apply(induction A rule: gen_boolean_algebra.induct) + apply blast + apply blast + apply blast + by blast + +lemma gen_boolean_algebra_intersect: + assumes "A \ gen_boolean_algebra S B" + assumes "C \ gen_boolean_algebra S B" + shows "A \ C \ gen_boolean_algebra S B" +proof- + have 0: "S - A \ gen_boolean_algebra S B" + using assms(1) gen_boolean_algebra.complement by blast + have 1: "S - C \ gen_boolean_algebra S B" + using assms(2) gen_boolean_algebra.complement by blast + have 2: "(S - A) \ (S - C) \ gen_boolean_algebra S B" + using "0" "1" gen_boolean_algebra.union by blast + have "S - (A \ C) \ gen_boolean_algebra S B" + by (simp add: 2 Diff_Int) + then have 3: "S - (S - (A \ C)) \ gen_boolean_algebra S B" + using gen_boolean_algebra.complement + by blast + have "A \ C \ S" + using assms(1) gen_boolean_algebra_subset + by blast + then show ?thesis + using 3 + by (metis "0" Diff_partition Un_subset_iff assms(1) double_diff gen_boolean_algebra_subset) +qed + +lemma gen_boolean_algebra_diff: + assumes "A \ gen_boolean_algebra S B" + assumes "C \ gen_boolean_algebra S B" + shows "A - C \ gen_boolean_algebra S B" +proof- + have "A - C = A \ (S - C)" + by (metis Int_Diff assms(1) gen_boolean_algebra_subset inf_absorb1) + then show ?thesis + by (metis assms(1) assms(2) gen_boolean_algebra.complement gen_boolean_algebra_intersect) +qed + +lemma gen_boolean_algebra_diff_eq: + assumes "A \ gen_boolean_algebra S B" + assumes "C \ gen_boolean_algebra S B" + shows "A - C = A \ (S - C)" + by (metis Int_Diff assms(1) gen_boolean_algebra_subset inf_absorb1) + +lemma gen_boolean_algebra_finite_union: + assumes "\a. a \ A \ a \ gen_boolean_algebra S B" + assumes "finite A" + shows "\A \ gen_boolean_algebra S B" +proof- + have "(\a \ A. a \ gen_boolean_algebra S B) \ \A \ gen_boolean_algebra S B" + apply(rule finite.induct[of A]) + apply (simp add: assms(2); fail) + apply (metis DiffE Union_empty ex_in_conv gen_boolean_algebra.simps) + by (metis Union_insert gen_boolean_algebra.simps insert_iff) + then show ?thesis using assms by blast +qed + +lemma gen_boolean_algebra_finite_intersection: + assumes "\a. a \ A \ a \ gen_boolean_algebra S B" + assumes "finite A" + assumes "A \ {}" + shows "\A \ gen_boolean_algebra S B" +proof- + have "(\a \ A. a \ gen_boolean_algebra S B) \ A \ {} \ \A \ gen_boolean_algebra S B" + apply(rule finite.induct[of A]) + apply (simp add: assms(2)) + apply force + using gen_boolean_algebra_intersect by auto + then show ?thesis using assms by blast +qed + +lemma gen_boolean_algebra_generators: + assumes "\b. b \ B \ b \ S" + assumes "b \ B" + shows "b \ gen_boolean_algebra S B" + unfolding gen_boolean_algebra.simps[of b] using assms(1)[of b] assms(2) by blast + +lemma gen_boolean_algebra_generator_subset: + assumes "A \ gen_boolean_algebra S As" + assumes "As \ Bs" + shows "A \ gen_boolean_algebra S Bs" + apply(rule gen_boolean_algebra.induct[of A S As]) + using assms(1) apply blast + apply (simp add: gen_boolean_algebra.intros(1); fail) + apply (meson Set.basic_monos(7) assms(2) gen_boolean_algebra.intros(2)) + using gen_boolean_algebra.intros(3) apply blast + using gen_boolean_algebra.intros(4) by blast + +lemma gen_boolean_algebra_generators_union: + assumes "A \ gen_boolean_algebra S As" + assumes "C \ gen_boolean_algebra S Cs" + shows "A \ C \ gen_boolean_algebra S (As \ Cs)" + apply(rule gen_boolean_algebra.induct[of C S Cs]) + using assms apply blast +apply(rule gen_boolean_algebra.union) + apply(rule gen_boolean_algebra_generator_subset[of _ _ As], rule assms, blast) + apply(rule gen_boolean_algebra.universe) + apply(rule gen_boolean_algebra.union) + apply(rule gen_boolean_algebra_generator_subset[of _ _ As], rule assms, blast) + apply(rule gen_boolean_algebra.generator, blast) + apply(rule gen_boolean_algebra.union) + apply(rule gen_boolean_algebra_generator_subset[of _ _ As], rule assms, blast) + apply(rule gen_boolean_algebra.union) + apply(rule gen_boolean_algebra_generator_subset[of _ _ Cs], blast, blast) + apply(rule gen_boolean_algebra_generator_subset[of _ _ Cs], blast, blast) + apply(rule gen_boolean_algebra.union) + apply(rule gen_boolean_algebra_generator_subset[of _ _ As], rule assms, blast) + apply(rule gen_boolean_algebra_generator_subset[of _ _ Cs]) + apply(rule gen_boolean_algebra_diff) + apply(rule gen_boolean_algebra.universe) + apply blast +by blast + +lemma gen_boolean_algebra_finite_gen_wits: + assumes "A \ gen_boolean_algebra S B" + shows "\ Bs. finite Bs \ Bs \ B \ A \ gen_boolean_algebra S Bs" +proof(rule gen_boolean_algebra.induct[of A S B]) + show " A \ gen_boolean_algebra S B" + using assms by blast + show "\Bs. finite Bs \ Bs \ B \ S \ gen_boolean_algebra S Bs" + using gen_boolean_algebra.universe[of S "{}"] + by blast + show "\A. A \ B \ \Bs. finite Bs \ Bs \ B \ A \ S \ gen_boolean_algebra S Bs" + proof- fix A assume A: "A \ B" + have 0: "{A} \ B" + using A by blast + show "\Bs. finite Bs \ Bs \ B \ A \ S \ gen_boolean_algebra S Bs" + using gen_boolean_algebra.generator[of A "{A}" S] 0 + by (meson finite.emptyI finite.insertI singletonI) + qed + show "\A C. A \ gen_boolean_algebra S B \ + \Bs. finite Bs \ Bs \ B \ A \ gen_boolean_algebra S Bs \ + C \ gen_boolean_algebra S B \ + \Bs. finite Bs \ Bs \ B \ C \ gen_boolean_algebra S Bs \ \Bs. finite Bs \ Bs \ B \ A \ C \ gen_boolean_algebra S Bs" + proof- fix A C + assume A: "A \ gen_boolean_algebra S B" + "\Bs. finite Bs \ Bs \ B \ A \ gen_boolean_algebra S Bs" + "C \ gen_boolean_algebra S B" + "\Bs. finite Bs \ Bs \ B \ C \ gen_boolean_algebra S Bs" + obtain As where As_def: "finite As \ As \ B \ A \ gen_boolean_algebra S As" + using A by blast + obtain Cs where Cs_def: "finite Cs \ Cs \ B \ C \ gen_boolean_algebra S Cs" + using A by blast + obtain Bs where Bs_def: "Bs = As \ Cs" + by blast + have Bs_sub: "Bs \ B" + unfolding Bs_def using As_def Cs_def by blast + have 0: " A \ C \ gen_boolean_algebra S Bs" + unfolding Bs_def + apply(rule gen_boolean_algebra_generators_union) + using As_def apply blast + using Cs_def by blast + have 1: "finite Bs" + unfolding Bs_def using As_def Cs_def by blast + show " \Bs. finite Bs \ Bs \ B \ A \ C \ gen_boolean_algebra S Bs" + using Bs_sub 0 1 by blast + qed + show "\A. A \ gen_boolean_algebra S B \ + \Bs. finite Bs \ Bs \ B \ A \ gen_boolean_algebra S Bs \ \Bs. finite Bs \ Bs \ B \ S - A \ gen_boolean_algebra S Bs" + using gen_boolean_algebra.complement by blast +qed + +lemma gen_boolean_algebra_univ_mono: + assumes "A \ gen_boolean_algebra S B" + shows "gen_boolean_algebra A B \ gen_boolean_algebra S B " +proof(rule subsetI) fix x assume A: "x \ gen_boolean_algebra A B" + obtain a where a_def: "a = A" + by blast + have 0: "a \ gen_boolean_algebra S B" + unfolding a_def using assms by blast + have 1: "a = A \ S" + using assms gen_boolean_algebra_subset unfolding a_def by blast + show "x \ gen_boolean_algebra S B " + apply(rule gen_boolean_algebra.induct[of x a B]) + using A a_def apply blast apply(rule 0) + apply (metis 1 Int_left_commute assms gen_boolean_algebra.intros(2) gen_boolean_algebra_intersect) + apply(rule gen_boolean_algebra.union, blast, blast) + apply(rule gen_boolean_algebra_diff) + apply(rule 0) + by blast +qed + +text\ + The boolean algebra generated by a collection of elements in another algebra is contained + in the original algebra: +\ +lemma gen_boolean_algebra_subalgebra: + assumes "Xs \ gen_boolean_algebra S B" + shows "gen_boolean_algebra S Xs \ gen_boolean_algebra S B" +proof fix x assume A: "x \ gen_boolean_algebra S Xs" + show "x \ gen_boolean_algebra S B " + apply(rule gen_boolean_algebra.induct[of x S Xs]) + apply (simp add: A; fail) + apply (simp add: gen_boolean_algebra.universe; fail) + using assms gen_boolean_algebra.universe gen_boolean_algebra_intersect apply blast + apply (simp add: gen_boolean_algebra.union; fail) + by (simp add: gen_boolean_algebra.complement) +qed + +lemma gen_boolean_algebra_idempotent: + assumes "S = \ Xs" + shows "gen_boolean_algebra S (gen_boolean_algebra S Xs) = (gen_boolean_algebra S Xs)" + apply(rule equalityI) + apply(rule subsetI) + apply (meson equalityD2 gen_boolean_algebra_subalgebra in_mono) + apply(rule subsetI) + by (metis gen_boolean_algebra.simps gen_boolean_algebra_subset inf.absorb1) + +text\We can always replace the set of generators \Xs\ with their intersections with the universe + set \S\, and obtain the same algebra.\ + +lemma gen_boolean_algebra_restrict_generators: +"gen_boolean_algebra S Xs =gen_boolean_algebra S ((\) S ` Xs)" +proof(rule equalityI') + fix x assume A: "x \ gen_boolean_algebra S Xs" + show "x \ gen_boolean_algebra S ((\) S ` Xs)" + apply(rule gen_boolean_algebra.induct[of x S Xs], rule A, rule gen_boolean_algebra.universe) + apply (metis gen_boolean_algebra.generator image_eqI inf.right_idem inf_commute) + apply(rule gen_boolean_algebra.union, blast, blast) + by(rule gen_boolean_algebra_diff, rule gen_boolean_algebra.universe, blast) +next + fix x assume A: "x \ gen_boolean_algebra S ((\) S ` Xs)" + show "x \ gen_boolean_algebra S Xs" + apply(rule gen_boolean_algebra.induct[of x S "(\) S ` Xs"], rule A, rule gen_boolean_algebra.universe, + rule gen_boolean_algebra_intersect ) + using gen_boolean_algebra.generator[of _ Xs S] + apply (metis (no_types, lifting) Int_commute image_iff) + apply(rule gen_boolean_algebra.universe) + apply(rule gen_boolean_algebra.union, blast, blast) + by(rule gen_boolean_algebra_diff, rule gen_boolean_algebra.universe, blast) +qed + +text\Adding a generator to a generated boolean algebra is redundant if the generator already + lies in the algebra.\ + +lemma add_generators: + assumes "A \ gen_boolean_algebra S Xs" + shows "gen_boolean_algebra S Xs = gen_boolean_algebra S (insert A Xs)" +proof(rule equalityI') + fix x assume A: "x \ gen_boolean_algebra S Xs" + show "x \ gen_boolean_algebra S (insert A Xs)" + apply(rule gen_boolean_algebra.induct[of x S Xs], rule A, rule gen_boolean_algebra.universe) + apply(rule gen_boolean_algebra.generator, blast) + apply(rule gen_boolean_algebra.union, blast,blast) + by(rule gen_boolean_algebra_diff, rule gen_boolean_algebra.universe, blast) +next + fix x assume A: "x \ gen_boolean_algebra S (insert A Xs)" + show "x \ gen_boolean_algebra S Xs" + apply(rule gen_boolean_algebra.induct[of x S "insert A Xs"], rule A, rule gen_boolean_algebra.universe) + using assms gen_boolean_algebra.generator[of _ Xs S] + using gen_boolean_algebra.universe gen_boolean_algebra_intersect apply blast + apply(rule gen_boolean_algebra.union, blast, blast) + by(rule gen_boolean_algebra_diff, rule gen_boolean_algebra.universe, blast) +qed +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Turning a Family of Sets into a Family of Disjoint Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + This section outlines the standard construction where sets $A_0, \dots, A_n$ are replaced by sets + $A_0, A_1 - A_0, A_2 - (A_0 \cup A_1), ..., A_n - (\bigcup \limits_{i = 0}^{n-1} A_i)$ to obtain + a disjoint family of the same cardinality. +\ +fun rec_disjointify where +"rec_disjointify 0 f = {}"| +"rec_disjointify (Suc m) f = insert (f m - \ (rec_disjointify m f)) (rec_disjointify m f)" + +lemma card_of_rec_disjointify: +"card (rec_disjointify m f) \ m" + apply(induction m) unfolding rec_disjointify.simps + apply simp + by (metis Suc_le_mono card.infinite card_insert_disjoint finite_insert insert_absorb le_SucI) + +lemma rec_disjointify_finite: +"finite (rec_disjointify m f)" + apply(induction m) + unfolding rec_disjointify.simps by auto + +lemma rec_disjointify_in_gen_boolean_algebra: + assumes "f ` {.. gen_boolean_algebra S B" + shows "rec_disjointify m f \ gen_boolean_algebra S B" +proof- + have "\k. k \ m \ rec_disjointify k f \ gen_boolean_algebra S B" + proof- fix k show "k \ m \ rec_disjointify k f \ gen_boolean_algebra S B" + apply(induction k) unfolding rec_disjointify.simps(1) using assms apply blast + proof fix k + assume IH: " k \ m \ rec_disjointify k f \ gen_boolean_algebra S B" + "Suc k \ m" + then have 0: "rec_disjointify k f \ gen_boolean_algebra S B" + by (simp add: IH(2)) + have 1: "finite (rec_disjointify k f )" + using rec_disjointify_finite by blast + have 2: "f k \ gen_boolean_algebra S B" + using IH(2) assms + by (simp add: image_subset_iff) + show "rec_disjointify (Suc k) f \ gen_boolean_algebra S B" + using 0 1 2 unfolding rec_disjointify.simps + by (simp add: gen_boolean_algebra_diff gen_boolean_algebra_finite_union subset_iff) + qed + qed + thus ?thesis by blast +qed + +lemma rec_disjointify_union: +"\ (rec_disjointify m f) = (\ i \ {.. (rec_disjointify m f)" + +lemma rec_disjointify_as_enum_rec_disjointify_image: +"rec_disjointify m f = enum_rec_disjointify f ` {.. f m" + unfolding enum_rec_disjointify_def + by auto + +lemma enum_rec_disjointify_disjoint: + assumes "k < m" + shows "enum_rec_disjointify f m \ enum_rec_disjointify f k = {}" +proof- + have "enum_rec_disjointify f k \ \ (rec_disjointify m f)" + unfolding rec_disjointify_union + using assms enum_rec_disjointify_subset by fastforce + thus ?thesis + unfolding enum_rec_disjointify_def + by auto +qed + +lemma enum_rec_disjointify_disjoint': + assumes "k \ m" + shows "enum_rec_disjointify f m \ enum_rec_disjointify f k = {}" + apply(cases "k < m") using enum_rec_disjointify_disjoint[of k m f] + apply simp + using assms enum_rec_disjointify_disjoint[of m k f] by auto + +lemma rec_disjointify_is_disjoint: + assumes "A \ rec_disjointify m f" + assumes "B \ rec_disjointify m f" + assumes "A \ B" + shows "A \ B = {}" + using rec_disjointify_as_enum_rec_disjointify_image enum_rec_disjointify_disjoint' assms + by (smt image_iff) + +definition enumerates where +"enumerates A f \ finite A \ A = f ` {..< (card A)} \ inj_on f {..< (card A)}" + +lemma finite_imp_exists_enumeration: + assumes "finite A" + shows "\f. enumerates A f" + unfolding enumerates_def + using assms finite_imp_nat_seg_image_inj_on[of A] + by (metis card_Collect_less_nat card_image lessThan_def) + +lemma enumeratesE: + assumes "enumerates A f" + shows "finite A" "A = f ` {..< card A}" "inj_on f {..< card A}" + using assms unfolding enumerates_def apply blast + using assms unfolding enumerates_def apply blast + using assms unfolding enumerates_def by blast + +lemma rec_disjointify_finite_set: + assumes "enumerates A f" + shows "\ (rec_disjointify (card A) f) = \ A" + unfolding rec_disjointify_union[of "card A" f] + using enumeratesE[of A f] assms by auto + +definition enumerate where +"enumerate A = (SOME f. enumerates A f)" + +lemma enumerate_enumerates: + assumes "finite A" + shows "enumerates A (enumerate A)" + unfolding enumerate_def using finite_imp_exists_enumeration assms + by (simp add: finite_imp_exists_enumeration some_eq_ex) + +lemma enumerateE: + assumes "finite A" + assumes "a \ A" + shows "\ i < card A. a = (enumerate A) i" + using enumerate_enumerates[of A] enumeratesE[of A] assms by blast + +definition disjointify where +"disjointify As = rec_disjointify (card As) (enumerate As)" + +lemma disjointify_is_disjoint: + assumes "finite As" + assumes "A \ disjointify As" + assumes "B \ disjointify As" + assumes "A \ B" + shows "A \ B = {}" + using assms rec_disjointify_is_disjoint[of A _ _ B] unfolding disjointify_def + by simp + +lemma disjointify_union: + assumes "finite As" + shows "\ (disjointify As) = \ As" + using assms + by (simp add: disjointify_def enumerate_enumerates rec_disjointify_finite_set) + +lemma disjointify_gen_boolean_algebra: + assumes "finite As" + assumes "As \ gen_boolean_algebra S B" + shows " disjointify As \ gen_boolean_algebra S B" + using assms unfolding disjointify_def + by (metis enumerate_enumerates enumeratesE(2) rec_disjointify_in_gen_boolean_algebra) + +lemma disjointify_finite: + assumes "finite As" + shows "finite (disjointify As)" + using assms unfolding disjointify_def + by (simp add: rec_disjointify_finite) + +lemma disjointify_card: + assumes "finite As" + shows"card (disjointify As) \ card As" + by (simp add: card_of_rec_disjointify disjointify_def) + +lemma disjointify_subset: + assumes "finite As" + assumes "A \ disjointify As" + shows "\B \ As. A \ B" + using assms enum_rec_disjointify_subset enumerate_enumerates enumeratesE + unfolding disjointify_def + by (smt image_iff rec_disjointify_as_enum_rec_disjointify_image) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\The Atoms Generated by Collections of Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + We can also turn a family of sets into a disjoint family by taking the atoms of the boolean + algebra generated by these sets. This will still yield a finite family if the initial family is + finite, but in general will be much larger in size. +\ + +(**********************************************************************) +(**********************************************************************) +subsubsection\Defining the Atoms of a Family of Sets\ +(**********************************************************************) +(**********************************************************************) +text\ + Here we intend that \As\ is a subset of the collection of sets \Xs\. This function associate to each + subset \As \ Xs\ a set which is contained in each element of \As\, and is disjoint from + each element of \Xs - As\. Note that in general this may yield the empty set, but we will + ultimately be interested in the cases where the result is nonempty.\ + +definition subset_to_atom where +"subset_to_atom Xs As = \ As - \ (Xs - As)" + +lemma subset_to_atom_memI: + assumes "\A. A \ As \ x \ A" + assumes "\A. A \ Xs \ A \ As \ x \ A" + shows "x \ subset_to_atom Xs As" + using assms unfolding subset_to_atom_def + by blast + +lemma subset_to_atom_memE: + assumes "x \ subset_to_atom Xs As" + shows "\A. A \ As \ x \ A" + "\A. A \ Xs \ A \ As \ x \ A" + using assms unfolding subset_to_atom_def by auto + +lemma subset_to_atom_closed: + assumes "As \ {}" + assumes "As \ Xs" + shows "subset_to_atom Xs As \ \ Xs" +proof- + have 0: "\ As \ \ As " + apply(rule subsetI) + using assms(1) by blast + show ?thesis + apply(rule subsetI) + using assms 0 unfolding subset_to_atom_def + by (meson DiffD1 Union_mono subsetD) +qed + +lemma subset_to_atom_as_intersection: + assumes "As \ {}" + assumes "As \ Xs" + assumes "S = \ Xs" + shows "subset_to_atom Xs As = \ As \ (\ X \ Xs - As. S - X)" + unfolding assms subset_to_atom_def + apply(rule equalityI') + apply(rule IntI, blast) + apply(rule InterI) + using INT_I assms(1) assms(2) apply auto[1] + apply(rule DiffI, blast) + by blast + +definition atoms_of where +"atoms_of Xs = (subset_to_atom Xs ` ((Pow Xs) - {{}})) - {{}}" + +lemma atoms_nonempty: + assumes "A \ atoms_of Xs" + shows "A \ {}" + using assms unfolding atoms_of_def by blast + +lemma atoms_of_disjoint: + assumes "A \ atoms_of Xs" + assumes "B \ atoms_of Xs" + assumes "A \ B" + shows "A \ B = {}" +proof- + obtain a where a_def: "a \ Xs \ A = subset_to_atom Xs a" + using assms unfolding atoms_of_def by blast + obtain b where b_def: "b \ Xs \ B = subset_to_atom Xs b" + using assms unfolding atoms_of_def by blast + have a_neq_b: "a \ b" + using assms a_def b_def by blast + have "A \ B \ {}" + proof fix x assume A: "x \ A \ B" + show "x \ {}" + proof(cases "a \ b") + case True + then obtain c where c_def: "c \ b - a" + using a_neq_b by blast + have c_in_Xs: "c \ Xs" + using c_def b_def by blast + have x_in_c: "x \ c" + using A b_def c_def subset_to_atom_memE[of x Xs b c] by blast + have x_notin_c: "x \ c" + using A a_def c_in_Xs c_def subset_to_atom_memE[of x Xs a c] by blast + then show ?thesis using x_in_c by blast + next + case False + then obtain c where c_def: "c \ a - b" + using a_neq_b by blast + have c_in_Xs: "c \ Xs" + using c_def a_def by blast + have x_in_c: "x \ c" + using A a_def c_def subset_to_atom_memE[of x Xs a c] by blast + have x_notin_c: "x \ c" + using A b_def c_in_Xs c_def subset_to_atom_memE[of x Xs b c] by blast + then show ?thesis using x_in_c by blast + qed + qed + thus "A \ B = {}" + by blast +qed + +text \ + The atoms of a family of sets \Xs\ are minimal in the sense that they are either contained in or + disjoint from each element of \Xs\. +\ +lemma atoms_are_minimal: + assumes "A \ atoms_of Xs" + assumes "X \ Xs" + shows "X \ A = {} \ A \ X" +proof(cases "X \ A = {}") + case True + then show ?thesis by blast +next + case False + obtain As where As_def: "As \ Pow Xs - {{}} \ A = subset_to_atom Xs As" + using assms unfolding atoms_of_def by blast + have A_simp: "A = subset_to_atom Xs As" + using As_def by blast + then show ?thesis using assms unfolding atoms_of_def subset_to_atom_def A_simp + using DiffD1 subset_eq by auto +qed + +(**********************************************************************) +(**********************************************************************) +subsubsection\Atoms Induced by Types of Points\ +(**********************************************************************) +(**********************************************************************) +text\ + The set of sets in \Xs\ which contain some point \x\. In the case where \Xs\ is some collection of + first order formulas, this is just the type of \x\ over these formulas.\ +definition point_to_type where +"point_to_type Xs x = {X \ Xs. x \ X}" + +text \The type of a point \x\ induces the unique atom of \Xs\ which contains \x\.\ +lemma point_in_atom_of_type: + assumes "x \ \ Xs" + shows "x \ subset_to_atom Xs (point_to_type Xs x)" + using assms unfolding subset_to_atom_def point_to_type_def + by blast + +lemma point_to_type_nonempty: + assumes "x \ \ Xs" + shows "point_to_type Xs x \{}" + using assms unfolding point_to_type_def + by blast + +lemma point_to_type_closed: + "point_to_type Xs x \ Pow (\ Xs)" + unfolding point_to_type_def + by blast + +lemma atoms_of_covers: + assumes "X = \ Xs" + shows "\ (atoms_of Xs) = X" +proof + show " \ (atoms_of Xs) \ X" + proof fix x assume A: "x \ \ (atoms_of Xs)" + then obtain As where As_def: "As \ Pow Xs - {{}} \ x \ subset_to_atom Xs As" + unfolding atoms_of_def by blast + have "subset_to_atom Xs As \ \ Xs" + using subset_to_atom_closed[of As Xs] As_def by blast + then show "x \ X" unfolding assms + using As_def by blast + qed + show "X \ \ (atoms_of Xs)" apply(rule subsetI) + using point_to_type_nonempty point_in_atom_of_type point_to_type_closed + unfolding assms point_to_type_def atoms_of_def + by fastforce +qed + +lemma atoms_of_covers': + shows "\ (atoms_of Xs) = \ Xs" + using atoms_of_covers[of "\ Xs"] by blast + +text \Every atom of a collection \Xs\ of sets is realized as the atom generated by the type of + an element in that atom.\ +lemma nonemtpy_atom_from_point_to_type: + assumes "A \ atoms_of Xs" + assumes "a \ A" + shows "A = subset_to_atom Xs (point_to_type Xs a)" +proof- + obtain As where As_def: "As \ (Pow Xs) - {} \ A = subset_to_atom Xs As" + using assms unfolding atoms_of_def by blast + have A_simp: "A = subset_to_atom Xs As" + using As_def by blast + have 0: "As = point_to_type Xs a" + apply(rule equalityI) + apply(rule subsetI) + apply (smt As_def Diff_empty UnionI Union_Pow_eq assms point_in_atom_of_type subset_to_atom_memE(1) subset_to_atom_memE(2)) + apply(rule subsetI) + using As_def assms subset_to_atom_memE(2) + by (metis (no_types, lifting) mem_Collect_eq point_to_type_def) + show ?thesis + using point_in_atom_of_type 0 + atoms_of_covers'[of Xs] assms unfolding A_simp + by auto +qed + +text \ + In light of the previous theorem, a point a and a collection of sets \Xs\ is enough to recover + the the unique atom of \Xs\ which contains \a\. +\ +definition point_to_atom where +"point_to_atom Xs a = subset_to_atom Xs (point_to_type Xs a)" + +lemma point_to_atom_closed: + assumes "x \ \ Xs" + shows "point_to_atom Xs x \ atoms_of Xs" + using assms unfolding atoms_of_def point_to_atom_def + by (metis (full_types) Union_iff atoms_of_covers atoms_of_def nonemtpy_atom_from_point_to_type) + +text \All atoms of \Xs\ are the atom induced by some point in the union of \Xs\.\ +lemma atoms_induced_by_points: +"atoms_of Xs = point_to_atom Xs ` (\ Xs)" + apply(rule equalityI) + apply(rule subsetI) + using nonemtpy_atom_from_point_to_type atoms_nonempty atoms_of_covers' + unfolding point_to_atom_def + apply (smt DiffE Pow_empty Pow_iff atoms_of_def image_iff subsetD subsetI subset_to_atom_closed) + apply(rule subsetI) + by (metis (no_types, lifting) imageE point_to_atom_closed point_to_atom_def) + +(**********************************************************************) +(**********************************************************************) +subsubsection\Atoms of Generated Boolean Algebras\ +(**********************************************************************) +(**********************************************************************) + +lemma atoms_of_gen_boolean_algebra: + assumes "Xs \ gen_boolean_algebra S B" + assumes "finite Xs" + shows "atoms_of Xs \ gen_boolean_algebra S B" +proof + fix x assume A: "x \ atoms_of Xs" + then obtain As where As_def: "As \ ((Pow Xs) - {{}}) \ x = subset_to_atom Xs As" + unfolding atoms_of_def by blast + have x_simp: "x = subset_to_atom Xs As" + using As_def by blast + have 0: "finite As" + using As_def assms finite_subset by auto + have 1: "As \ gen_boolean_algebra S B" + using As_def assms by blast + have 2: "\ As \ gen_boolean_algebra S B" + using 0 1 assms + by (metis As_def DiffE gen_boolean_algebra_finite_intersection singletonI subset_eq) + show "x \ gen_boolean_algebra S B" + using A 2 unfolding atoms_of_def subset_to_atom_def x_simp + by (metis (no_types, lifting) As_def DiffD1 Diff_partition Pow_iff Un_subset_iff assms(1) assms(2) finite_subset gen_boolean_algebra_diff gen_boolean_algebra_finite_union order_refl subsetD) +qed + + +text \If the generators of a boolean algebra are contained in the universe, the atoms induced by + the generators alone are minimal elements of the entire algebra.\ +lemma finite_algebra_atoms_are_minimal: + assumes "finite Xs" + assumes "\ Xs \ S" + assumes "A \ atoms_of Xs" + assumes "X \ gen_boolean_algebra S Xs" + shows "X \ A = {} \ A \ X" + apply(rule gen_boolean_algebra.induct[of X S Xs]) + apply (simp add: assms(4); fail) + apply (metis Union_upper assms(2) assms(3) atoms_of_covers dual_order.trans) + using assms(2) assms(3) atoms_are_minimal apply fastforce + apply blast + using assms + by (metis Diff_Int_distrib2 Diff_empty Diff_eq_empty_iff Sup_upper atoms_of_covers' equalityE inf.absorb_iff2 order_trans) + +lemma finite_set_imp_finite_atoms: + assumes "finite Xs" + shows "finite (atoms_of Xs)" + using assms unfolding atoms_of_def + by blast + +text \ + Every element in the boolean algebra generated by \Xs\ over \S\ is a (disjoint) union + of atoms of generators: +\ + +lemma gen_boolean_algebra_elem_uni_of_atoms: + assumes "finite Xs" + assumes "S = \ Xs" + assumes "X \ gen_boolean_algebra S Xs" + shows "X = \ {a \ atoms_of Xs. a \ X}" +proof + show "X \ \ {a \ atoms_of Xs. a \ X}" + proof fix x assume A: "x \ X" + then have "point_to_atom Xs x \ atoms_of Xs" + using assms by (meson gen_boolean_algebra_subset point_to_atom_closed subsetD) + then show "x \ \ {a \ atoms_of Xs. a \ X}" + by (smt A IntI Union_iff assms(1) assms(2) assms(3) empty_iff finite_algebra_atoms_are_minimal gen_boolean_algebra.universe gen_boolean_algebra_subset mem_Collect_eq point_in_atom_of_type point_to_atom_def subsetD) + qed + show "\ {a \ atoms_of Xs. a \ X} \ X" + by blast +qed + +text\In fact, every generated boolean algebra is the power set of the atoms of its generators:\ +lemma gen_boolean_algebra_generated_by_atoms: + assumes "finite Xs" + assumes "S = \ Xs" + shows "gen_boolean_algebra S Xs = \ ` (Pow (atoms_of Xs))" +proof + show "gen_boolean_algebra S Xs \ \ ` Pow (atoms_of Xs)" + apply(rule subsetI) + using gen_boolean_algebra_elem_uni_of_atoms[of Xs S] assms + by fastforce + show "\ ` Pow (atoms_of Xs) \ gen_boolean_algebra S Xs" + apply(rule subsetI) + using atoms_of_gen_boolean_algebra[of Xs S Xs] + finite_subset[of _ "atoms_of Xs"] assms + finite_set_imp_finite_atoms[of Xs] + gen_boolean_algebra_finite_union[of _ S Xs] + by (smt Pow_iff Union_upper gen_boolean_algebra.intros(2) image_iff inf.absorb1 subsetD subsetI) +qed + +text\Finitely generated boolean algebras are finite\ +lemma fin_gens_imp_fin_algebra: + assumes "finite Xs" + assumes "S = \ Xs" + shows "finite (gen_boolean_algebra S Xs)" + using finite_set_imp_finite_atoms[of Xs] assms gen_boolean_algebra_generated_by_atoms[of Xs S] + by simp + + +lemma point_to_atom_equal: + assumes "finite Xs" + assumes "S = \ Xs" + assumes "x \ S" + shows "point_to_atom Xs x = point_to_atom (gen_boolean_algebra S Xs) x" +proof + show P0: "point_to_atom Xs x \ point_to_atom (gen_boolean_algebra S Xs) x" + proof- + have 0: "point_to_atom Xs x \ point_to_atom (gen_boolean_algebra S Xs) x \ {}" + using assms + by (metis IntI UnionI empty_iff gen_boolean_algebra.universe point_in_atom_of_type point_to_atom_def) + have 1: "point_to_atom (gen_boolean_algebra S Xs) x \ gen_boolean_algebra S Xs" + using assms fin_gens_imp_fin_algebra[of Xs S] + by (meson UnionI atoms_of_gen_boolean_algebra gen_boolean_algebra.simps point_to_atom_closed subset_eq subset_refl) + then show ?thesis + using 0 finite_algebra_atoms_are_minimal[of Xs S "point_to_atom Xs x" "point_to_atom (gen_boolean_algebra S Xs) x"] + assms(1) assms(2) assms(3) atoms_induced_by_points by auto + qed + show "point_to_atom (gen_boolean_algebra S Xs) x \ point_to_atom Xs x" + proof- + have 0: "point_to_atom (gen_boolean_algebra S Xs) x \ point_to_atom Xs x \{}" + using assms P0 point_in_atom_of_type point_to_atom_def by fastforce + have 1: "point_to_atom (gen_boolean_algebra S Xs) x \ (gen_boolean_algebra S Xs)" + using assms gen_boolean_algebra_idempotent[of S Xs] atoms_of_gen_boolean_algebra + by (metis UnionI fin_gens_imp_fin_algebra gen_boolean_algebra.universe point_to_atom_closed subset_eq) + have 2: "\ (gen_boolean_algebra S Xs) \ S" + using assms + by (simp add: Sup_le_iff gen_boolean_algebra_subset) + hence 3: "\ (gen_boolean_algebra S Xs) = S" + by (simp add: Union_upper gen_boolean_algebra.universe subset_antisym) + have 4: "gen_boolean_algebra S (gen_boolean_algebra S Xs) = gen_boolean_algebra S Xs" + using assms gen_boolean_algebra_idempotent[of S Xs] by blast + have 5: "point_to_atom Xs x \ gen_boolean_algebra S (gen_boolean_algebra S Xs)" + unfolding 4 using assms + by (metis (no_types, opaque_lifting) Int_absorb1 Int_commute Union_upper atoms_of_gen_boolean_algebra gen_boolean_algebra.generator point_to_atom_closed subsetD subsetI) + show ?thesis + using 2 5 finite_algebra_atoms_are_minimal[of "gen_boolean_algebra S Xs" S "point_to_atom (gen_boolean_algebra S Xs) x" "point_to_atom Xs x"] 0 1 2 + unfolding 4 + by (metis "3" Int_commute assms(1) assms(2) assms(3) fin_gens_imp_fin_algebra point_to_atom_closed) + qed +qed + +text \ + When the set \Xs\ of generators covers the universe set \S\, the atoms of \Xs\ in the above + sense are the same as the atoms of the boolean algebra they generate over \S\. +\ + +lemma atoms_of_sets_eq_atoms_of_algebra: + assumes "finite Xs" + assumes "S = \ Xs" + shows "atoms_of Xs = atoms_of (gen_boolean_algebra S Xs)" +proof + show "atoms_of Xs \ atoms_of (gen_boolean_algebra S Xs)" + proof fix A assume A: "A \ atoms_of Xs" + then obtain x where x_def: "x \ S \ A = point_to_atom Xs x" + using assms + by (metis atoms_induced_by_points image_iff) + have 0: "A = point_to_atom (gen_boolean_algebra S Xs) x" + using assms point_to_atom_equal x_def by fastforce + show "A \ atoms_of (gen_boolean_algebra S Xs)" + unfolding 0 using assms A + by (metis (full_types) "0" UnionI gen_boolean_algebra.universe point_to_atom_closed x_def) + qed + show "atoms_of (gen_boolean_algebra S Xs) \ atoms_of Xs" + proof fix A assume A: "A \ atoms_of (gen_boolean_algebra S Xs)" + then obtain x where x_def: "x \ S \ A = point_to_atom (gen_boolean_algebra S Xs) x" + by (metis atoms_induced_by_points cSup_eq_maximum gen_boolean_algebra.universe gen_boolean_algebra_subset image_iff) + then show "A \ atoms_of Xs" + using assms(1) assms(2) point_to_atom_closed point_to_atom_equal by fastforce + qed +qed + +lemma atoms_closed: + assumes "finite Xs" + assumes "A \ atoms_of (gen_boolean_algebra S Xs)" + assumes "S = \ Xs" + shows "A \ (gen_boolean_algebra S Xs)" +proof- + have 1: "A = \ {A}" + by blast + have 2: "A \ atoms_of Xs" + using assms atoms_of_sets_eq_atoms_of_algebra + by blast + show ?thesis + using gen_boolean_algebra_generated_by_atoms[of Xs S] + assms 1 2 unfolding Pow_def by blast +qed + +lemma atoms_finite: + assumes "finite Xs" + shows "finite ((atoms_of (gen_boolean_algebra S Xs)))" +proof- + have 0: "gen_boolean_algebra S Xs =gen_boolean_algebra S ((\) S ` Xs)" + using gen_boolean_algebra_restrict_generators by blast + have 1: "gen_boolean_algebra S Xs = gen_boolean_algebra S (insert S ((\) S ` Xs))" + unfolding 0 by(rule add_generators, rule gen_boolean_algebra.universe) + obtain Ys where Ys_def: "Ys = (insert S ((\) S ` Xs))" + by blast + have Ys_finite: "finite Ys" + unfolding Ys_def using assms by blast + have 2: "\ Ys = S" + unfolding Ys_def + by blast + have 3: "atoms_of Ys = atoms_of (gen_boolean_algebra S Xs) " + unfolding Ys_def 1 + apply(rule atoms_of_sets_eq_atoms_of_algebra) + using Ys_finite unfolding Ys_def apply blast + by blast + have 4: "finite (atoms_of Ys)" + by(rule finite_set_imp_finite_atoms, rule Ys_finite) + show ?thesis using 4 unfolding 3 by blast +qed + + +text \ + We can distinguish atoms of a set of generators \Cs\ by finding some element of \Cs\ which + includes one and excludes the other. +\ + +lemma distinct_atoms: + assumes "Cs \ {}" + assumes "a \ atoms_of Cs" + assumes "b \ atoms_of Cs" + assumes "a \ b" + shows "(\B \ Cs. b \ B \ a \ B = {}) \ (\A \ Cs. a \ A \ b \ A = {})" +proof- + obtain x where x_def: "x \ \ Cs \ a = point_to_atom Cs x" + by (metis assms(2) atoms_induced_by_points imageE) + obtain y where y_def: "y \ \ Cs \ b = point_to_atom Cs y" + by (metis assms(3) atoms_induced_by_points imageE) + have 0: "point_to_atom Cs x \ point_to_atom Cs y" + using x_def y_def assms by simp + hence 1: "point_to_type Cs x \ point_to_type Cs y" + unfolding point_to_atom_def subset_to_atom_def by blast + then obtain B where B_def: "B \ Cs \ (B \ point_to_type Cs x - point_to_type Cs y \ B \ point_to_type Cs y - point_to_type Cs x)" + unfolding point_to_type_def by blast + have 2: "B \ point_to_type Cs x - point_to_type Cs y \ a \ B" + using x_def point_to_atom_def subset_to_atom_memE(1) by fastforce + have 3: "B \ point_to_type Cs y - point_to_type Cs y \ b \ B" + using y_def by blast + show ?thesis using B_def 2 3 + by (smt Diff_iff disjoint_iff_not_equal point_to_atom_def subset_eq subset_to_atom_memE(1) subset_to_atom_memE(2) x_def y_def) +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Partitions of a Set\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition disjoint :: "'a set set \ bool" where +"disjoint Ss = (\ A \ Ss. \B \ Ss. A \B \ A \ B = {})" + +lemma disjointE: + assumes "disjoint Ss" + assumes "A \ Ss" + assumes "B \ Ss" + assumes "A \B" + shows "A \ B = {}" + by (meson assms(1) assms(2) assms(3) assms(4) disjoint_def) + +lemma disjointI: + assumes "\A B. A \ Ss \ B \ Ss \ A \ B \ A \ B = {}" + shows "disjoint Ss" + by (meson assms disjoint_def) + +definition is_partition :: "'a set set \ 'a set \ bool" (infixl "partitions" 75) where +"S partitions A = (disjoint S \ \ S = A)" + +lemma is_partitionE: + assumes "S partitions A" + shows "disjoint S" + "\ S = A" + using assms is_partition_def apply blast + using assms + by (simp add: is_partition_def) + +lemma is_partitionI: + assumes "disjoint S" + assumes "\ S = A" + shows "S partitions A" + using assms is_partition_def by blast + +text \ + If we start with a finite partition of a set \A\, and each element in that partition has a + finite partition with some property \P\, then \A\ itself has a finite partition where each + element has property \P\.\ + +lemma iter_partition: + assumes "As partitions A" + assumes "finite As" + assumes "\a. a \ As \ \Bs. finite Bs \ Bs partitions a \ (\b \ Bs. P b)" + shows "\Bs. finite Bs \ Bs partitions A \ (\b \ Bs. P b)" +proof- + obtain F where F_def: "F = (\a. (SOME Bs. finite Bs \ Bs partitions a \ (\b \ Bs. P b)))" + by blast + have FE: "\a. a \ As \ finite (F a) \ (F a) partitions a \ (\b \ (F a). P b)" + proof- fix a assume A: "a \ As" + show "finite (F a) \ (F a) partitions a \ (\b \ (F a). P b)" + apply(rule SomeE'[of _ "\Bs. finite Bs \ Bs partitions a \ (\b \ Bs. P b)"]) + unfolding F_def apply blast + using assms by (simp add: A) + qed + obtain Bs where Bs_def: "Bs = (\ a \ As. F a)" + by blast + have 0: "finite Bs" + unfolding Bs_def using FE assms by blast + have 1: "disjoint Bs" + proof(rule disjointI) + fix a b assume A: "a \ Bs" "b \ Bs" "a \ b" + obtain c where c_def: "c \ As \ a \ F c" + using Bs_def A by blast + obtain d where d_def: "d \ As \ b \ F d" + using Bs_def A by blast + have 0: "a \ c" + using c_def FE[of c] is_partitionE(2)[of "F c" c] by blast + have 1: "b \ d" + using d_def FE[of d] is_partitionE(2)[of "F d" d] by blast + show "a \ b = {}" + proof(cases "c = d") + case True + show ?thesis apply(rule disjointE[of "F c"]) + unfolding True using FE is_partitionE d_def apply blast + using c_def unfolding True apply blast + using d_def apply blast + by(rule A) + next + case False + have "c \ d = {}" + apply(rule disjointE[of As]) + using assms is_partitionE apply blast + using c_def apply blast + using d_def apply blast + using False by blast + then show ?thesis using 0 1 by blast + qed + qed + have 2: "(\b \ Bs. P b)" + apply(rule ) + unfolding Bs_def using FE + by blast + have FE': "\a. a \ As \ (\ (F a)) = a " + apply(rule is_partitionE) + using FE by blast + have 3: "Bs partitions A" + apply(rule is_partitionI, rule 1) +apply(rule equalityI') + unfolding Bs_def using assms is_partitionE(2)[of As A] + FE' is_partitionE(2) apply blast + proof- + fix x assume A: "x \ A" + then obtain a where a_def: "a \ As \ x \ a" + using assms is_partitionE by blast + then have "x \ (\ (F a))" + using a_def FE' by blast + thus " x \ \ (\ (F ` As))" + using a_def A by blast + qed + show "\Bs. finite Bs \ Bs partitions A \ (\b\Bs. P b)" + using 0 2 3 by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Intersections of Families of Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition pairwise_intersect where +"pairwise_intersect As Bs = {c. \a \ As. \b \ Bs. c = a \ b}" + +lemma partition_intersection: + assumes "As partitions A" + assumes "Bs partitions B" + shows "(pairwise_intersect As Bs) partitions (A \ B)" +proof(rule is_partitionI, rule disjointI) + fix a b assume a0: "a \ pairwise_intersect As Bs" "b \ pairwise_intersect As Bs" "a \ b" + obtain a1 b1 where def1: "a1 \ As \ b1 \ Bs \ a = a1 \ b1" + using a0 unfolding pairwise_intersect_def by blast + obtain a2 b2 where def2: "a2 \ As \ b2 \ Bs \ b = a2 \ b2" + using a0 unfolding pairwise_intersect_def by blast + have 0: "a \ b = (a1 \ a2) \ (b1 \ b2)" + using def1 def2 by blast + show " a \ b= {}" + proof(cases "a1 \ a2") + case True + have T0: "a1 \ a2 = {}" + apply(rule disjointE[of As a1 a2] ) + using def1 def2 assms(1) True is_partitionE(1)[of As A] apply blast + using def1 apply blast using def2 apply blast by(rule True) + thus ?thesis unfolding 0 by blast + next + case False + then have F0: "b1 \ b2" + using a0 def1 def2 by blast + have F1: "b1 \ b2 = {}" + apply(rule disjointE[of Bs b1 b2]) + using def1 def2 assms(2) F0 is_partitionE(1)[of Bs B] apply blast + using def1 apply blast using def2 apply blast by(rule F0) + thus ?thesis unfolding 0 by blast + qed +next + show "\ (pairwise_intersect As Bs) = A \ B" + proof(rule equalityI') + fix x assume A: "x \ \ (pairwise_intersect As Bs)" + then obtain a b where def1: "a \ As \ b \ Bs \ x \ a \ b" + unfolding pairwise_intersect_def by blast + have 0: "a \ A" + using def1 assms is_partitionE by blast + have 1: "b \ B" + using def1 assms is_partitionE by blast + show " x \ A \ B" + using 0 1 def1 by blast + next + fix x assume A: "x \ A \ B" + obtain a where a_def: "a \ As \ x \ a" + using A assms is_partitionE by blast + obtain b where b_def: "b \ Bs \ x \ b" + using A assms is_partitionE by blast + have 0: "x \ a \ b" + using a_def b_def by blast + show "x \ \ (pairwise_intersect As Bs)" + using a_def b_def 0 unfolding pairwise_intersect_def + by blast + qed +qed + +lemma pairwise_intersect_finite: + assumes "finite As" + assumes "finite Bs" + shows "finite (pairwise_intersect As Bs)" +proof- + have 0: "(pairwise_intersect As Bs) = (\ a \ As. (\) a ` Bs)" + unfolding pairwise_intersect_def + apply(rule equalityI') + unfolding mem_Collect_eq apply blast + by blast + have 1: "\a. a \ As \ finite ((\) a ` Bs)" + using assms by blast + show ?thesis unfolding 0 using assms(1) 1 by blast +qed + +definition family_intersect where +"family_intersect parts = atoms_of (\ parts)" + +lemma family_intersect_partitions: + assumes "\Ps. Ps \ parts \ Ps partitions A" + assumes "\Ps. Ps \ parts \ finite Ps" + assumes "finite parts" + assumes "parts \ {}" + shows "family_intersect parts partitions A" +proof(rule is_partitionI) + show "disjoint (family_intersect parts)" + apply(rule disjointI) + unfolding family_intersect_def apply(rule atoms_of_disjoint) + apply blast + apply blast + by blast + show " \ (family_intersect parts) = A" + proof- + have 0: "\ (family_intersect parts) = \ (\ parts)" + unfolding family_intersect_def + apply(rule atoms_of_covers) + by blast + have 1: "\Ps. Ps \ parts \ \Ps = A" + by(rule is_partitionE, rule assms, blast) + show ?thesis unfolding 0 + using 1 assms by blast + qed +qed + +lemma family_intersect_memE: + assumes "\Ps. Ps \ parts \ Ps partitions A" + assumes "\Ps. Ps \ parts \ finite Ps" + assumes "finite parts" + assumes "parts \ {}" + shows "\Ps a. a \ family_intersect parts \ Ps \ parts \ \P \ Ps. a \ P" +proof- + fix Ps a assume A: "a \ family_intersect parts" "Ps \ parts" + have 0: "\ Ps = A" + apply(rule is_partitionE) + using A assms by blast + have 1: "\ (family_intersect parts) = A" + apply(rule is_partitionE) + using family_intersect_partitions assms by blast + have 2: "a \ {}" + using A unfolding family_intersect_def atoms_of_def by blast + obtain P where P_def: "P \ Ps \ a \ P \ {}" + using 0 1 A 2 by blast + have P_in: "P \ (\ parts)" + using P_def A by blast + have a_sub: "a \ P" + using atoms_are_minimal P_def A P_in unfolding family_intersect_def by blast + show "\P \ Ps. a \ P" + using a_sub P_def by blast +qed + +lemma family_intersect_mem_inter: + assumes "\Ps. Ps \ (parts:: 'a set set set) \ Ps partitions A" + assumes "\Ps. Ps \ parts \ finite Ps" + assumes "finite parts" + assumes "parts \ {}" + assumes "a \ family_intersect parts" + shows "\f. \ Ps \ parts. f Ps \ Ps \ a = (\ Ps \ parts. f Ps)" +proof- + obtain f where f_def: "f = (\Ps:: 'a set set. (SOME P. P \ Ps \ a \ P))" + by blast + have f_eval: "\Ps. Ps \ parts \ f Ps \ Ps \ a \ (f Ps)" + proof- + fix Ps assume A: "Ps \ parts" + obtain P where P_def: "P \ Ps \ a \ P" + using assms family_intersect_memE A by blast + show " f Ps \ Ps \ a \ f Ps" + apply(rule SomeE[of "f Ps" _ P]) + unfolding f_def using A apply simp + by(rule P_def) + qed + have 0: "a \ {}" + using assms unfolding family_intersect_def + using atoms_nonempty by blast + have 1: "a = (\ Ps \ parts. f Ps)" + proof(rule equalityI) + show 10: "a \ \ (f ` parts)" + using f_eval by blast + show "\ (f ` parts) \ a" + proof + fix x assume A: "x \ \ (f ` parts)" + obtain b where b_def: "b = point_to_atom (\ parts) x" + by blast + have b_atom: "b \ atoms_of (\ parts)" + unfolding b_def apply(rule point_to_atom_closed) + using A f_eval assms by blast + show x_in_a: "x \ a" + proof(rule ccontr) + assume "x \ a" + then have "\ b \ a" + using b_def unfolding point_to_atom_def point_to_type_def subset_to_atom_def by blast + hence p0: "a \ b" + by blast + have p1: "b \ a = {}" + apply(rule atoms_of_disjoint[of _ "(\ parts)"] ) + apply(rule b_atom) + using assms unfolding family_intersect_def apply blast + using p0 by blast + have p2: " (\B\\ parts. b \ B \ a \ B = {}) \ (\A\\ parts. a \ A \ b \ A = {})" + using distinct_atoms[of "\ parts" a b] assms + by (metis Sup_bot_conv(1) b_atom equalityI' f_eval family_intersect_def mem_simps(2) p0) + show False + proof(cases "(\B\\ parts. b \ B \ a \ B = {})") + case True + then obtain B where B_def: "B\\ parts \ b \ B \ a \ B = {}" + by blast + obtain Ps where Ps_def: "B \ Ps \ Ps \ parts" + using B_def by blast + have B_neq: "B \ f Ps" + using Ps_def B_def 10 0 by blast + have B_cap: "B \ f Ps = {}" + apply(rule disjointE[of Ps]) + apply(rule is_partitionE[of Ps A]) + using Ps_def assms apply blast + using Ps_def apply blast + using f_eval Ps_def apply blast + by(rule B_neq) + have b_cap: "b \ f Ps = {}" + using B_cap B_def by blast + have x_in_b: "x \ b" + using b_def unfolding point_to_atom_def point_to_type_def subset_to_atom_def + by blast + show False using x_in_b b_cap Ps_def A by blast + next + case False + then obtain B where B_def: "B\\ parts \ a \ B \ b \ B = {}" + using p2 by blast + obtain Ps where Ps_def: "B \ Ps \ Ps \ parts" + using B_def by blast + have F0: "B = f Ps" + proof(rule ccontr) + assume not: "B \ f Ps" + have F0: "B \ f Ps = {}" + apply(rule disjointE[of Ps]) + apply(rule is_partitionE[of Ps A]) + using Ps_def assms apply blast + using Ps_def apply blast + using f_eval Ps_def apply blast + by(rule not) + have a_sub: "a \ f Ps" + using 10 Ps_def by blast + show False using F0 B_def a_sub 0 by blast + qed + have x_in_B: "x \ B" + unfolding F0 using A Ps_def by blast + have x_in_b: "x \ b" + using b_def unfolding point_to_atom_def point_to_type_def subset_to_atom_def + by blast + show False using x_in_b x_in_B B_def by blast + qed + qed + qed + qed + show ?thesis using f_eval 1 by blast +qed + +text \ + If we take a finite family of partitions in a particular generated boolean algebra, where each + partition itself is finite, then their induced partition is also in the algebra.\ +lemma family_intersect_in_gen_boolean_algebra: + assumes "A \ gen_boolean_algebra S B" + assumes "\Ps. Ps \ parts \ Ps partitions A" + assumes "\Ps. Ps \ parts \ finite Ps" + assumes "\Ps P. Ps \ parts \ P \ Ps \ P \ gen_boolean_algebra S B" + assumes "finite parts" + assumes "parts \ {}" + shows "\P. P \ family_intersect parts \ P \ gen_boolean_algebra S B" +proof- + fix P assume A: "P \ family_intersect parts" + have 0: "P \ atoms_of (\ parts)" + using A unfolding family_intersect_def by blast + have 1: "finite (\ parts)" + using assms by blast + have 2: "\ parts \ gen_boolean_algebra S B" + using assms by blast + obtain Ps where Ps_def: "Ps \ parts" + using assms by blast + have 3: "\ (\ parts) = A" + apply(rule equalityI') + using assms is_partitionE(2)[of _ A] apply blast + using assms is_partitionE(2)[of Ps A] Ps_def by blast + have 4: "atoms_of (\ parts) = atoms_of (gen_boolean_algebra A (\ parts))" + apply(rule atoms_of_sets_eq_atoms_of_algebra[of "\ parts" A]) + apply(rule 1) + unfolding 3 by blast + have 5: "atoms_of (\ parts) \ (gen_boolean_algebra A (\ parts))" + apply(rule atoms_of_gen_boolean_algebra) + using 3 gen_boolean_algebra.generator[of _ "\ parts" A] + apply (meson Sup_upper gen_boolean_algebra_generators subsetI) + by(rule 1) + have 6: "A \ S" + using assms gen_boolean_algebra_subset by blast + have 7: "(gen_boolean_algebra A (\ parts)) \ gen_boolean_algebra (S) (\ parts)" + apply(rule gen_boolean_algebra_univ_mono) + using 3 gen_boolean_algebra_finite_union[of "\ parts" "S" "\ parts"] + gen_boolean_algebra.generator[of _ "\ parts" "S" ] 6 1 + by (meson Sup_le_iff gen_boolean_algebra_generators) + have 8: "gen_boolean_algebra (S) (\ parts) \ gen_boolean_algebra S B" + apply(rule gen_boolean_algebra_subalgebra) + using 2 by blast + show "P \ gen_boolean_algebra S B" + using 0 5 6 7 8 by blast +qed + + + +end diff --git a/thys/Padic_Field/Indices.thy b/thys/Padic_Field/Indices.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Indices.thy @@ -0,0 +1,571 @@ +theory Indices +imports Main +begin + +(**************************************************************************************************) +(**************************************************************************************************) +section\Basic Lemmas for Manipulating Indices and Lists\ +(**************************************************************************************************) +(**************************************************************************************************) +fun index_list where +"index_list 0 = []"| +"index_list (Suc n) = index_list n @ [n]" + +lemma index_list_length: +"length (index_list n) = n" + by(induction n, simp, auto ) + +lemma index_list_indices: +"k < n \ (index_list n)!k = k" + apply(induction n) + apply (simp; fail) + by (simp add: index_list_length nth_append) + +lemma index_list_set: +"set (index_list n) = {.. 'b list) => 'a list => 'b list" where + "flat_map f [] = []" + |"flat_map f (h#t) = (f h)@(flat_map f t)" + +abbreviation(input) project_at_indices ("\\<^bsub>_\<^esub>") where +"project_at_indices S as \ nths as S" + +fun insert_at_index :: " 'a list \'a \ nat \ 'a list" where +"insert_at_index as a n= (take n as) @ (a#(drop n as))" + +lemma insert_at_index_length: + shows "length (insert_at_index as a n) = length as + 1" + by(induction n, auto) + +lemma insert_at_index_eq[simp]: + assumes "n \ length as" + shows "(insert_at_index as a n)!n = a" + by (metis assms insert_at_index.elims length_take min.absorb2 nth_append_length) + +lemma insert_at_index_eq'[simp]: + assumes "n \ length as" + assumes "k < n" + shows "(insert_at_index as a n)!k = as ! k" + using assms + by (simp add: nth_append) + +lemma insert_at_index_eq''[simp]: + assumes "n < length as" + assumes "k \ n" + shows "(insert_at_index as a k)!(Suc n) = as ! n" + using assms insert_at_index.simps[of as a k] + by (smt Suc_diff_Suc append_take_drop_id diff_Suc_Suc dual_order.order_iff_strict + le_imp_less_Suc length_take less_trans min.absorb2 not_le nth_Cons_Suc nth_append) + +text\Correctness of project\_at\_indices\ + +definition indices_of :: "'a list \ nat set" where +"indices_of as = {..<(length as)}" + +lemma proj_at_index_list_length[simp]: + assumes "S \ indices_of as" + shows "length (project_at_indices S as) = card S" +proof- + have "S = {i. i < length as \ i \ S}" + using assms unfolding indices_of_def + by blast + thus ?thesis + using length_nths[of as S] by auto +qed + +text\A function which enumerates finite sets\ + +abbreviation(input) set_to_list :: "nat set \ nat list" where +"set_to_list S \ sorted_list_of_set S" + +lemma set_to_list_set: + assumes "finite S" + shows "set (set_to_list S) = S" + by (simp add: assms) + +lemma set_to_list_length: + assumes "finite S" + shows "length (set_to_list S) = card S" + by (metis assms length_remdups_card_conv length_sort set_sorted_list_of_set sorted_list_of_set_sort_remdups) + +lemma set_to_list_empty: + assumes "card S = 0" + shows "set_to_list S = []" + by (metis assms length_0_conv length_sorted_list_of_set) + +lemma set_to_list_first: + assumes "card S > 0" + shows "Min S = set_to_list S ! 0 " +proof- + have 0: "set (set_to_list S) = S" + using assms card_ge_0_finite set_sorted_list_of_set by blast + have 1: "sorted (set_to_list S)" + by simp + show ?thesis apply(rule Min_eqI) + using assms card_ge_0_finite apply blast + apply (metis "0" "1" in_set_conv_nth less_Suc0 less_or_eq_imp_le not_less_eq sorted_iff_nth_mono_less) + by (metis "0" Max_in assms card_0_eq card_ge_0_finite gr_zeroI in_set_conv_nth not_less0) +qed + +lemma set_to_list_last: + assumes "card S > 0" + shows "Max S = last (set_to_list S)" +proof- + have 0: "set (set_to_list S) = S" + using assms card_ge_0_finite set_sorted_list_of_set by blast + have 1: "sorted (set_to_list S)" + by simp + show ?thesis apply(rule Max_eqI) + using assms card_ge_0_finite apply blast + apply (smt "0" "1" Suc_diff_1 in_set_conv_nth last_conv_nth le_simps(2) length_greater_0_conv + less_or_eq_imp_le nat_neq_iff neq0_conv not_less_eq sorted_iff_nth_mono_less) + by (metis "0" assms card.empty empty_set last_in_set less_numeral_extra(3)) +qed + +lemma set_to_list_insert_Max: + assumes "finite S" + assumes "\s. s \ S \ a > s" + shows "set_to_list (insert a S) = set_to_list S @[a]" + by (metis assms(1) assms(2) card_0_eq card_insert_if finite.insertI infinite_growing + insert_not_empty less_imp_le_nat sorted_insort_is_snoc sorted_list_of_set(1) sorted_list_of_set(2) + sorted_list_of_set_insert) + +lemma set_to_list_insert_Min: + assumes "finite S" + assumes "\s. s \ S \ a < s" + shows "set_to_list (insert a S) = a#set_to_list S" + by (metis assms(1) assms(2) insort_is_Cons nat_less_le sorted_list_of_set(1) sorted_list_of_set_insert) + +fun nth_elem where +"nth_elem S n = set_to_list S ! n" + +lemma nth_elem_closed: + assumes "i < card S" + shows "nth_elem S i \ S" + by (metis assms card.infinite not_less0 nth_elem.elims nth_mem set_to_list_length sorted_list_of_set(1)) + +lemma nth_elem_Min: + assumes "card S > 0" + shows "nth_elem S 0 = Min S" + by (simp add: assms set_to_list_first) + +lemma nth_elem_Max: + assumes "card S > 0" + shows "nth_elem S (card S - 1) = Max S" +proof- + have "last (set_to_list S) = set_to_list S ! (card S - 1)" + by (metis assms card_0_eq card_ge_0_finite last_conv_nth neq0_conv set_to_list_length sorted_list_of_set_eq_Nil_iff) + thus ?thesis + using assms set_to_list_last set_to_list_length + by simp +qed + +lemma nth_elem_Suc: + assumes "card S > Suc n" + shows "nth_elem S (Suc n) > nth_elem S n" + using assms sorted_sorted_list_of_set[of S] set_to_list_length[of S] + by (metis Suc_lessD card.infinite distinct_sorted_list_of_set lessI nat_less_le not_less0 nth_elem.elims nth_eq_iff_index_eq sorted_iff_nth_mono_less) + +lemma nth_elem_insert_Min: + assumes "card S > 0" + assumes "a < Min S" + shows "nth_elem (insert a S) (Suc i) = nth_elem S i" + using assms + by (metis Min_gr_iff card_0_eq card_ge_0_finite neq0_conv nth_Cons_Suc nth_elem.elims set_to_list_insert_Min) + +lemma set_to_list_Suc_map: + assumes "finite S" + shows "set_to_list (Suc ` S) = map Suc (set_to_list S)" +proof- + obtain n where n_def: "n = card S" + by blast + have "\S. card S = n \ set_to_list (Suc ` S) = map Suc (set_to_list S)" + proof(induction n) + case 0 + then show ?case + by (metis card_eq_0_iff finite_imageD image_is_empty inj_Suc list.simps(8) set_to_list_empty) + next + case (Suc n) + have 0: "S = insert (Min S) (S - {Min S})" + by (metis Min_in Suc.prems card_gt_0_iff insert_Diff zero_less_Suc) + have 1: "sorted_list_of_set (Suc ` (S - {Min S})) = map Suc (sorted_list_of_set (S - {Min S}))" + by (metis "0" Suc.IH Suc.prems card_Diff_singleton card.infinite diff_Suc_1 insertI1 nat.simps(3)) + have 2: "set_to_list S = (Min S)#(set_to_list (S - {Min S}))" + by (metis "0" DiffD1 Min_le Suc.prems card_Diff_singleton card.infinite card_insert_if + diff_Suc_1 finite_Diff n_not_Suc_n nat.simps(3) nat_less_le set_to_list_insert_Min) + have 3: "sorted_list_of_set (Suc ` S) = (Min (Suc ` S))#(set_to_list ((Suc ` S) - {Min (Suc ` S)}))" + by (metis DiffD1 Diff_idemp Min_in Min_le Suc.prems card_Diff1_less card_eq_0_iff finite_Diff + finite_imageI image_is_empty insert_Diff nat.simps(3) nat_less_le set_to_list_insert_Min) + have 4: "(Min (Suc ` S)) = Suc (Min S)" + by (metis Min.hom_commute Suc.prems Suc_le_mono card_eq_0_iff min_def nat.simps(3)) + have 5: "sorted_list_of_set (Suc ` S) = Suc (Min S)#(set_to_list ((Suc ` S) - {Suc (Min S)}))" + using 3 4 by auto + have 6: "sorted_list_of_set (Suc ` S) = Suc (Min S)#(set_to_list (Suc ` (S - {Min S})))" + by (metis (no_types, lifting) "0" "5" Diff_insert_absorb image_insert inj_Suc inj_on_insert) + show ?case + using 6 + by (simp add: "1" "2") + qed + thus ?thesis + using n_def by blast +qed + +lemma nth_elem_Suc_im: + assumes "i < card S" + shows "nth_elem (Suc ` S) i = Suc (nth_elem S i) " + using set_to_list_Suc_map + by (metis assms card_ge_0_finite dual_order.strict_trans not_gr0 nth_elem.elims nth_map set_to_list_length) + +lemma set_to_list_upto: +"set_to_list {..Characterizing the entries of project\_at\_indices \ + +lemma project_at_indices_append: +"project_at_indices S (as@bs) = project_at_indices S as @ project_at_indices {j. j + length as \ S} bs" + using nths_append[of as bs S] by auto + +lemma project_at_indices_nth: + assumes "S \ indices_of as" + assumes "card S > i" + shows "project_at_indices S as ! i = as ! (nth_elem S i)" +proof- + have "\ S i. S \ indices_of as \ card S > i \ project_at_indices S as ! i = as ! (nth_elem S i)" + proof(induction as) + case Nil + then show ?case + by (metis list.size(3) not_less0 nths_nil proj_at_index_list_length) + next + case (Cons a as) + assume A: "S \ indices_of (a # as) \ i < card S" + have 0: "nths (a # as) S = (if 0 \ S then [a] else []) @ nths as {j. Suc j \ S}" + using nths_Cons[of a as S] by simp + show "nths (a # as) S ! i = (a # as) ! nth_elem S i" + proof(cases "0 \ S") + case True + show ?thesis + proof(cases "S = {0}") + case True + then show ?thesis + using "0" Cons.prems by auto + next + case False + have T0: "nths (a # as) S = a#nths as {j. Suc j \ S}" + using 0 + by (simp add: True) + have T1: "{j. Suc j \ S} \ indices_of as" + proof fix x assume A: "x \ {j. Suc j \ S}" + then have "Suc x < length (a#as)" + using Cons.prems indices_of_def by blast + then show "x \ indices_of as" + by (simp add: indices_of_def) + qed + have T2: "\i. i < card {j. Suc j \ S} \ nths as {j. Suc j \ S} ! i = as ! nth_elem {j. Suc j \ S} i" + using Cons.IH T1 by blast + have T3: "\i. i < card {j. Suc j \ S} \ nth_elem {j. j > 0 \ j\ S} i = nth_elem S (Suc i)" + proof- + have 0: " 0 < card {j. Suc j \ S}" + by (smt Cons.prems Diff_iff Diff_subset False T0 T1 True add_diff_cancel_left' + card.insert card_0_eq card.infinite finite_subset gr_zeroI insert_Diff + length_Cons n_not_Suc_n plus_1_eq_Suc proj_at_index_list_length singletonI) + have 1: "(insert 0 {j. 0 < j \ j \ S}) = S" + apply(rule set_eqI) using True gr0I by blast + have 2: "0 < Min {j. 0 < j \ j \ S}" using False + by (metis (mono_tags, lifting) "1" Cons.prems Min_in finite_insert finite_lessThan + finite_subset indices_of_def less_Suc_eq less_Suc_eq_0_disj mem_Collect_eq singleton_conv) + show "\i. i < card {j. Suc j \ S} \ nth_elem {j. 0 < j \ j \ S} i = nth_elem S (Suc i)" + using 0 1 2 nth_elem_insert_Min[of "{j. 0 < j \ j \ S}" 0] True False + by (metis (no_types, lifting) Cons.prems T0 T1 card_gt_0_iff finite_insert length_Cons less_SucI proj_at_index_list_length) + qed + show "nths (a # as) S ! i = (a # as) ! nth_elem S i" + apply(cases "i = 0") + apply (metis Cons.prems Min_le T0 True card_ge_0_finite le_zero_eq nth_Cons' nth_elem_Min) + proof- + assume "i \ 0" + then have "i = Suc (i - 1)" + using Suc_pred' by blast + hence "nths (a # as) S ! i = nths as {j. Suc j \ S} ! (i-1)" + using A by (simp add: T0) + thus "nths (a # as) S ! i = (a # as) ! nth_elem S i" + proof- + have "i - 1 < card {j. Suc j \ S}" + by (metis Cons.prems Suc_less_SucD T0 T1 \i = Suc (i - 1)\ length_Cons proj_at_index_list_length) + hence 0: "nth_elem {j. 0 < j \ j \ S} (i - 1) = nth_elem S i" + using T3[of "i-1"] \i = Suc (i - 1)\ by auto + + have 1: "nths as {j. Suc j \ S} ! (i-1) = as ! nth_elem {j. Suc j \ S} (i-1)" + using T2 \i - 1 < card {j. Suc j \ S}\ by blast + have 2: "(a # as) ! nth_elem S i = as! ((nth_elem S i) - 1)" + by (metis Cons.prems \i = Suc (i - 1)\ not_less0 nth_Cons' nth_elem_Suc) + have 3: "(nth_elem S i) - 1 = nth_elem {j. Suc j \ S} (i-1)" + proof- + have "Suc ` {j. Suc j \ S} = {j. 0 < j \ j \ S}" + proof + show "Suc ` {j. Suc j \ S} \ {j. 0 < j \ j \ S}" + by blast + show "{j. 0 < j \ j \ S} \ Suc ` {j. Suc j \ S}" + using Suc_pred gr0_conv_Suc by auto + qed + thus ?thesis + using "0" \i - 1 < card {j. Suc j \ S}\ nth_elem_Suc_im by fastforce + qed + show "nths (a # as) S ! i = (a # as) ! nth_elem S i" + using "1" "2" "3" \nths (a # as) S ! i = nths as {j. Suc j \ S} ! (i - 1)\ by auto + qed + qed + qed + next + case False + have F0: "nths (a # as) S = nths as {j. Suc j \ S}" + by (simp add: "0" False) + have F1: "Suc `{j. Suc j \ S} = S" + proof show "Suc ` {j. Suc j \ S} \ S" by auto + show "S \ Suc ` {j. Suc j \ S}" using False Suc_pred + by (smt image_iff mem_Collect_eq neq0_conv subsetI) + qed + have F2: "{j. Suc j \ S} \ indices_of as \ i < card {j. Suc j \ S}" + using F1 + by (metis (mono_tags, lifting) A F0 Suc_less_SucD indices_of_def length_Cons lessThan_iff + mem_Collect_eq proj_at_index_list_length subset_iff) + have F3: "project_at_indices {j. Suc j \ S} as ! i = as ! (nth_elem {j. Suc j \ S} i)" + using F2 Cons(1)[of "{j. Suc j \ S}"] Cons(2) + by blast + then show ?thesis + using F0 F1 F2 nth_elem_Suc_im by fastforce + qed + qed + then show ?thesis + using assms(1) assms(2) by blast +qed + +text\An inverse for nth\_elem\ + +definition set_rank where +"set_rank S x = (THE i. i < card S \ x = nth_elem S i)" + +lemma set_rank_exist: + assumes "finite S" + assumes "x \ S" + shows "\i. i < card S \ x = nth_elem S i" + using assms nth_elem.simps[of S] + by (metis in_set_conv_nth set_to_list_length sorted_list_of_set(1)) + +lemma set_rank_unique: + assumes "finite S" + assumes "x \ S" + assumes "i < card S \ x = nth_elem S i" + assumes "j < card S \ x = nth_elem S j" + shows "i = j" + using assms nth_elem.simps[of S] + by (simp add: \i < card S \ x = nth_elem S i\ \j < card S \ x = nth_elem S j\ + nth_eq_iff_index_eq set_to_list_length) + +lemma nth_elem_set_rank_inv: + assumes "finite S" + assumes "x \ S" + shows "nth_elem S (set_rank S x) = x" + using the_equality set_rank_unique set_rank_exist assms + unfolding set_rank_def + by smt + +lemma set_rank_nth_elem_inv: + assumes "finite S" + assumes "i < card S" + shows "set_rank S (nth_elem S i) = i" + using the_equality set_rank_unique set_rank_exist assms + unfolding set_rank_def +proof - + show "(THE n. n < card S \ nth_elem S i = nth_elem S n) = i" + using assms(1) assms(2) nth_elem_closed set_rank_unique by blast +qed + +lemma set_rank_range: + assumes "finite S" + assumes "x \ S" + shows "set_rank S x < card S" + using assms(1) assms(2) set_rank_exist set_rank_nth_elem_inv by fastforce + +lemma project_at_indices_nth': + assumes "S \ indices_of as" + assumes "i \ S" + shows "as ! i = project_at_indices S as ! (set_rank S i) " + by (metis assms(1) assms(2) finite_lessThan finite_subset indices_of_def nth_elem_set_rank_inv + project_at_indices_nth set_rank_range) + +fun proj_away_from_index :: "nat \ 'a list \ 'a list" ("\\<^bsub>\_\<^esub>")where +"proj_away_from_index n as = (take n as)@(drop (Suc n) as)" + +text\proj\_away\_from\_index is an inverse to insert\_at\_index\ + +lemma insert_at_index_project_away[simp]: + assumes "k < length as" + assumes "bs = (insert_at_index as a k)" + shows "\\<^bsub>\ k\<^esub> bs = as" + using assms insert_at_index.simps[of as a k] proj_away_from_index.simps[of k bs] + by (simp add: \k < length as\ less_imp_le_nat min.absorb2) + +definition fibred_cell :: "'a list set \ ('a list \ 'a \ bool) \ 'a list set" where +"fibred_cell C P = {as . \x t. as = (t#x) \ x \ C \ (P x t)}" + +definition fibred_cell_at_ind :: "nat \ 'a list set \ ('a list \ 'a \ bool) \ 'a list set" where +"fibred_cell_at_ind n C P = {as . \x t. as = (insert_at_index x t n) \ x \ C \ (P x t)}" + +lemma fibred_cell_lengths: + assumes "\k. k \ C \ length k = n" + shows "k \ (fibred_cell C P) \ length k = Suc n" +proof- + assume "k \ (fibred_cell C P)" + obtain x t where "k = (t#x) \ x \ C \ P x t" + proof - + assume a1: "\t x. k = t # x \ x \ C \ P x t \ thesis" + have "\as a. k = a # as \ as \ C \ P as a" + using \k \ fibred_cell C P\ fibred_cell_def by blast + then show ?thesis + using a1 by blast + qed + then show ?thesis + by (simp add: assms) +qed + +lemma fibred_cell_at_ind_lengths: + assumes "\k. k \ C \ length k = n" + assumes "k \ n" + shows "c \ (fibred_cell_at_ind k C P) \ length c = Suc n" +proof- + assume "c \ (fibred_cell_at_ind k C P)" + then obtain x t where "c = (insert_at_index x t k) \ x \ C \ (P x t)" + using assms + unfolding fibred_cell_at_ind_def + by blast + then show ?thesis + by (simp add: assms(1)) +qed + +lemma project_fibred_cell: + assumes "\k. k \ C \ length k = n" + assumes "k < n" + assumes "\x \ C. \t. P x t" + shows "\\<^bsub>\ k\<^esub> ` (fibred_cell_at_ind k C P) = C" +proof + show "\\<^bsub>\k\<^esub> ` fibred_cell_at_ind k C P \ C" + proof + fix x + assume x_def: "x \ \\<^bsub>\k\<^esub> ` fibred_cell_at_ind k C P" + then obtain c where c_def: "x = \\<^bsub>\k\<^esub> c \ c \ fibred_cell_at_ind k C P" + by blast + then obtain y t where yt_def: "c = (insert_at_index y t k) \ y \ C \ (P y t)" + using assms + unfolding fibred_cell_at_ind_def + by blast + have 0: "x =\\<^bsub>\k\<^esub> c" + by (simp add: c_def) + have 1: "y =\\<^bsub>\k\<^esub> c" + using yt_def assms(1) assms(2) + by (metis insert_at_index_project_away) + have 2: "x = y" using 0 1 by auto + then show "x \ C" + by (simp add: yt_def) + qed + show "C \ \\<^bsub>\k\<^esub> ` fibred_cell_at_ind k C P" + proof fix x + assume A: "x \ C" + obtain t where t_def: "P x t" + using assms A by auto + then show "x \ \\<^bsub>\k\<^esub> ` fibred_cell_at_ind k C P" + proof - + have f1: "\a n A as. take n as @ (a::'a) # drop n as \ A \ as \ \\<^bsub>\n\<^esub> ` A \ \ n < length as" + by (metis insert_at_index.simps insert_at_index_project_away rev_image_eqI) + have "\n. \as a. take n x @ t # drop n x = insert_at_index as a n \ as \ C \ P as a" + using A t_def by auto + then have "\n. take n x @ t # drop n x \ {insert_at_index as a n |as a. as \ C \ P as a}" + by blast + then have "x \ \\<^bsub>\k\<^esub> ` {insert_at_index as a k |as a. as \ C \ P as a}" + using f1 by (metis (lifting) A assms(1) assms(2)) + then show ?thesis + by (simp add: fibred_cell_at_ind_def) + qed + qed +qed + +definition list_segment where +"list_segment i j as = map (nth as) [i.. j" + assumes "j \ length as" + shows "length (list_segment i j as) = j - i" + using assms + unfolding list_segment_def + by (metis length_map length_upt) + +lemma list_segment_drop: + assumes "i < length as" + shows "(list_segment i (length as) as) = drop i as" + by (metis One_nat_def Suc_diff_Suc add_diff_inverse_nat drop0 drop_map drop_upt + less_Suc_eq list_segment_def map_nth neq0_conv not_less0 plus_1_eq_Suc) + +lemma list_segment_concat: + assumes "j \ k" + assumes "i \ j" + shows "(list_segment i j as) @ (list_segment j k as) = (list_segment i k as)" + using assms unfolding list_segment_def + using le_Suc_ex upt_add_eq_append + by fastforce + +lemma list_segment_subset: + assumes "j \ k" + shows "set (list_segment i j as) \ set (list_segment i k as)" + apply(cases "i > j") + unfolding list_segment_def + apply (metis in_set_conv_nth length_map list.size(3) order.asym subsetI upt_rec zero_order(3)) +proof- + assume "\ j < i" + then have "i \j" + using not_le + by blast + then have "list_segment i j as @ list_segment j k as = list_segment i k as" + using assms list_segment_concat[of j k i as] by auto + then show "set (map ((!) as) [i.. set (map ((!) as) [i.. length as" + shows "set (list_segment i j as) \ set as" + apply(cases "i \ j") + apply (simp add: list_segment_def) +proof- + assume A: "\ j \ i" + then have B: "i < j" + by auto + have 0: "list_segment i (length as) as = drop i as" + using B assms list_segment_drop[of i as] less_le_trans + by blast + have 1: "set (list_segment i j as) \ set (list_segment i (length as) as)" + using B assms list_segment_subset[of j "length as" i as] + by blast + then show ?thesis + using assms 0 dual_order.trans set_drop_subset[of i as] + by metis +qed + +definition fun_inv where +"fun_inv = inv" + + + +end diff --git a/thys/Padic_Field/Padic_Field_Polynomials.thy b/thys/Padic_Field/Padic_Field_Polynomials.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Padic_Field_Polynomials.thy @@ -0,0 +1,1811 @@ +theory Padic_Field_Polynomials + imports Padic_Fields + +begin + +(**************************************************************************************************) +(**************************************************************************************************) +section\$p$-adic Univariate Polynomials and Hensel's Lemma\ +(**************************************************************************************************) +(**************************************************************************************************) +type_synonym padic_field_poly = "nat \ padic_number" + +type_synonym padic_field_fun = "padic_number \ padic_number" + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Gauss Norms of Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +text \ + The Gauss norm of a polynomial is defined to be the minimum valuation of a coefficient of that + polynomial. This induces a valuation on the ring of polynomials, and in particular it satisfies + the ultrametric inequality. In addition, the Gauss norm of a polynomial $f(x)$ gives a lower + bound for the value $\text{val } (f(a))$ in terms of $\text{val }(a)$, for a point + $a \in \mathbb{Q}_p$. We introduce Gauss norms here as a useful tool for stating and proving + Hensel's Lemma for the field $\mathbb{Q}_p$. We are abusing terminology slightly in calling + this the Gauss norm, rather than the Gauss valuation, but this is just to conform with our + decision to work exclusively with the $p$-adic valuation and not discuss the equivalent + real-valued $p$-adic norm. For a detailed treatment of Gauss norms one can see, for example + \cite{engler2005valued}. +\ +context padic_fields +begin + +no_notation Zp.to_fun (infixl\\\ 70) + +abbreviation(input) Q\<^sub>p_x where +"Q\<^sub>p_x \ UP Q\<^sub>p" + +definition gauss_norm where +"gauss_norm g = Min (val ` g ` {..degree g}) " + +lemma gauss_normE: + assumes "g \ carrier Q\<^sub>p_x" + shows "gauss_norm g \ val (g k)" + apply(cases "k \ degree g") + unfolding gauss_norm_def + using assms apply auto[1] +proof- + assume "\ k \ degree g" + then have "g k = \\<^bsub>Q\<^sub>p\<^esub> " + by (simp add: UPQ.deg_leE assms) + then show "Min (val ` g ` {..deg Q\<^sub>p g}) \ val (g k)" + by (simp add: local.val_zero) +qed + +lemma gauss_norm_geqI: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\n. val (g n) \ \" + shows "gauss_norm g \ \" + unfolding gauss_norm_def using assms + by simp + +lemma gauss_norm_eqI: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\n. val (g n) \ \" + assumes "val (g i) = \" + shows "gauss_norm g = \" +proof- + have 0: "gauss_norm g \ \" + using assms gauss_normE gauss_norm_def by fastforce + have 1: "gauss_norm g \ \" + using assms gauss_norm_geqI by auto + show ?thesis using 0 1 by auto +qed + +lemma nonzero_poly_nonzero_coeff: + assumes "g \ carrier Q\<^sub>p_x" + assumes "g \ \\<^bsub>Q\<^sub>p_x\<^esub>" + shows "\k. k \degree g \ g k \\\<^bsub>Q\<^sub>p\<^esub>" +proof(rule ccontr) + assume "\ (\k\degree g. g k \ \\<^bsub>Q\<^sub>p\<^esub>)" + then have 0: "\k. g k = \\<^bsub>Q\<^sub>p\<^esub>" + by (meson UPQ.deg_leE assms(1) not_le_imp_less) + then show False + using assms UPQ.cfs_zero by blast +qed + +lemma gauss_norm_prop: + assumes "g \ carrier Q\<^sub>p_x" + assumes "g \ \\<^bsub>Q\<^sub>p_x\<^esub>" + shows "gauss_norm g \ \" +proof- + obtain k where k_def: "k \degree g \ g k \\\<^bsub>Q\<^sub>p\<^esub>" + using assms nonzero_poly_nonzero_coeff + by blast + then have 0: "gauss_norm g \ val (g k)" + using assms(1) gauss_normE by blast + have "g k \ carrier Q\<^sub>p" + using UPQ.cfs_closed assms(1) by blast + hence "val (g k) < \" + using k_def assms + by (metis eint_ord_code(3) eint_ord_simps(4) val_ineq) + then show ?thesis + using 0 not_le by fastforce +qed + +lemma gauss_norm_coeff_norm: + "\n \ degree g. (gauss_norm g) = val (g n)" +proof- + have "finite (val ` g ` {..deg Q\<^sub>p g})" + by blast + hence "\x \ (val ` g ` {..deg Q\<^sub>p g}). gauss_norm g = x" + unfolding gauss_norm_def + by auto + thus ?thesis unfolding gauss_norm_def + by blast +qed + +lemma gauss_norm_smult_cfs: + assumes "g \ carrier Q\<^sub>p_x" + assumes "a \ carrier Q\<^sub>p" + assumes "gauss_norm g = val (g k)" + shows "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) = val a + val (g k)" +proof- + obtain l where l_def: "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) = val ((a \\<^bsub>Q\<^sub>p_x\<^esub> g) l)" + using gauss_norm_coeff_norm + by blast + then have "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) = val (a \\<^bsub>Q\<^sub>p\<^esub> (g l))" + using assms + by simp + then have "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) = val a + val (g l)" + by (simp add: UPQ.cfs_closed assms(1) assms(2) val_mult) + then have 0: "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) \ val a +val (g k)" + using assms gauss_normE[of g l] + by (metis UPQ.UP_smult_closed UPQ.cfs_closed UPQ.cfs_smult gauss_normE val_mult) + have "val a + val (g k) = val ((a \\<^bsub>Q\<^sub>p_x\<^esub> g) k)" + by (simp add: UPQ.cfs_closed assms(1) assms(2) val_mult) + then have "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) \ val a + val (g k)" + by (metis \gauss_norm (a \\<^bsub>UP Q\<^sub>p\<^esub> g) = val a + val (g l)\ add_left_mono assms(1) assms(3) gauss_normE) + then show ?thesis + using 0 by auto +qed + +lemma gauss_norm_smult: + assumes "g \ carrier Q\<^sub>p_x" + assumes "a \ carrier Q\<^sub>p" + shows "gauss_norm (a \\<^bsub>Q\<^sub>p_x\<^esub> g) = val a + gauss_norm g" + using gauss_norm_smult_cfs[of g a] gauss_norm_coeff_norm[of g] assms + by metis + +lemma gauss_norm_ultrametric: + assumes "g \ carrier Q\<^sub>p_x" + assumes "h \ carrier Q\<^sub>p_x" + shows "gauss_norm (g \\<^bsub>Q\<^sub>p_x\<^esub> h) \ min (gauss_norm g) (gauss_norm h)" +proof- + obtain k where "gauss_norm (g \\<^bsub>Q\<^sub>p_x\<^esub> h) = val ((g \\<^bsub>Q\<^sub>p_x\<^esub> h) k)" + using gauss_norm_coeff_norm + by blast + then have 0: "gauss_norm (g \\<^bsub>Q\<^sub>p_x\<^esub> h) = val (g k \\<^bsub>Q\<^sub>p\<^esub> h k)" + by (simp add: assms(1) assms(2)) + have "min (val (g k)) (val (h k))\ min (gauss_norm g) (gauss_norm h)" + using gauss_normE[of g k] gauss_normE[of h k] assms(1) assms(2) min.mono + by blast + then show ?thesis + using 0 val_ultrametric[of "g k" "h k"] assms(1) assms(2) dual_order.trans + by (metis (no_types, lifting) UPQ.cfs_closed) +qed + +lemma gauss_norm_a_inv: + assumes "f \ carrier (UP Q\<^sub>p)" + shows "gauss_norm (\\<^bsub>UP Q\<^sub>p\<^esub>f) = gauss_norm f" +proof- + have 0: "\n. ((\\<^bsub>UP Q\<^sub>p\<^esub>f) n) = \ (f n)" + using assms by simp + have 1: "\n. val ((\\<^bsub>UP Q\<^sub>p\<^esub>f) n) = val (f n)" + using 0 assms UPQ.UP_car_memE(1) val_minus by presburger + obtain i where i_def: "gauss_norm f = val (f i)" + using assms gauss_norm_coeff_norm by blast + have 2: "\k. val ((\\<^bsub>UP Q\<^sub>p\<^esub>f) k) \ val (f i)" + unfolding 1 + using i_def assms gauss_normE by fastforce + show ?thesis + apply(rule gauss_norm_eqI[of _ _ i]) + apply (simp add: assms; fail) + unfolding 1 using assms gauss_normE apply blast + unfolding i_def by blast +qed + +lemma gauss_norm_ultrametric': + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "g \ carrier (UP Q\<^sub>p)" + shows "gauss_norm (f \\<^bsub>UP Q\<^sub>p\<^esub> g) \ min (gauss_norm f) (gauss_norm g)" + unfolding a_minus_def + using assms gauss_norm_a_inv[of g] gauss_norm_ultrametric + by (metis UPQ.UP_a_inv_closed) + +lemma gauss_norm_finsum: + assumes "f \ A \ carrier Q\<^sub>p_x" + assumes "finite A" + assumes "A \ {}" + shows " gauss_norm (\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) \ Min (gauss_norm ` (f`A))" +proof- + obtain k where k_def: "val ((\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) k) = gauss_norm (\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i)" + by (metis gauss_norm_coeff_norm) + then have 0: "val (\\<^bsub>Q\<^sub>p\<^esub>i\A. f i k) \ Min (val ` (\ i. f i k) ` A)" + using finsum_val_ultrametric[of "\ i. f i k" A] assms + by (simp add: \\(\i. f i k) \ A \ carrier Q\<^sub>p; finite A; A \ {}\ \ Min (val ` (\i. f i k) ` A) \ val (\i\A. f i k)\ Pi_iff UPQ.cfs_closed) + have "(\a. a \ A \ (val \ (\i. f i k)) a \ gauss_norm (f a))" + using gauss_normE assms + by (metis (no_types, lifting) Pi_split_insert_domain Set.set_insert comp_apply) + then have "Min (val ` (\ i. f i k) ` A) \ Min ((\ i. gauss_norm (f i)) ` A)" + using Min_mono'[of A] + by (simp add: assms(2) image_comp) + then have 1: "Min (val ` (\ i. f i k) ` A) \ Min (gauss_norm ` f ` A)" + by (metis image_image) + have "f \ A \ carrier (UP Q\<^sub>p) \ ((\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) \ carrier Q\<^sub>p_x \ ((\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) k) = (\\<^bsub>Q\<^sub>p\<^esub>i\A. f i k)) " + apply(rule finite.induct[of A]) + apply (simp add: assms(2); fail) + apply (metis (no_types, lifting) Pi_I Qp.add.finprod_one_eqI UPQ.P.finsum_closed UPQ.P.finsum_empty UPQ.cfs_zero empty_iff) + proof- + fix a A assume A: "finite A" "f \ A \ carrier (UP Q\<^sub>p) \ ( finsum (UP Q\<^sub>p) f A \ carrier (UP Q\<^sub>p) \ finsum (UP Q\<^sub>p) f A k = (\i\A. f i k)) " + show " f \ insert a A \ carrier (UP Q\<^sub>p) \ finsum (UP Q\<^sub>p) f (insert a A) \ carrier (UP Q\<^sub>p) \ finsum (UP Q\<^sub>p) f (insert a A) k = (\i\insert a A. f i k)" + apply(cases "a \ A") + using A + apply (simp add: insert_absorb; fail) + proof assume B: "a \ A" " f \ insert a A \ carrier (UP Q\<^sub>p)" + then have f_a: "f a \ carrier (UP Q\<^sub>p)" + by blast + have f_A: "f \ A \ carrier (UP Q\<^sub>p)" + using B by blast + have "finsum (UP Q\<^sub>p) f (insert a A) = f a \\<^bsub>UP Q\<^sub>p\<^esub>finsum (UP Q\<^sub>p) f A" + using assms A B f_a f_A finsum_insert by simp + then have 0: "finsum (UP Q\<^sub>p) f (insert a A) k = f a k \\<^bsub>Q\<^sub>p\<^esub> (finsum (UP Q\<^sub>p) f A) k" + using f_a f_A A B + by simp + have " ( \ a. f a k) \ A \ carrier Q\<^sub>p" + proof fix a assume "a \ A" + then have "f a \ carrier (UP Q\<^sub>p)" + using f_A by blast + then show "f a k \ carrier Q\<^sub>p" + using A cfs_closed by blast + qed + then have 0: "finsum (UP Q\<^sub>p) f (insert a A) k = (\i\insert a A. f i k)" + using A B Qp.finsum_insert[of A a "\ a. f a k"] + by (simp add: UPQ.cfs_closed) + thus " finsum (UP Q\<^sub>p) f (insert a A) \ carrier (UP Q\<^sub>p) \ finsum (UP Q\<^sub>p) f (insert a A) k = (\i\insert a A. f i k)" + using B(2) UPQ.P.finsum_closed by blast + qed + qed + then have "(\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) \ carrier Q\<^sub>p_x \ ((\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) k) = (\\<^bsub>Q\<^sub>p\<^esub>i\A. f i k)" + using assms by blast + hence 3: "gauss_norm (\\<^bsub>Q\<^sub>p_x\<^esub>i\A. f i) \ Min (val ` (\ i. f i k) ` A)" + using 0 k_def by auto + thus ?thesis + using 1 le_trans by auto +qed + +lemma gauss_norm_monom: + assumes "a \ carrier Q\<^sub>p" + shows "gauss_norm (monom Q\<^sub>p_x a n) = val a" +proof- + have "val ((monom Q\<^sub>p_x a n) n) \ gauss_norm (monom Q\<^sub>p_x a n)" + using assms gauss_normE[of "monom Q\<^sub>p_x a n" n] UPQ.monom_closed + by blast + then show ?thesis + using gauss_norm_coeff_norm[of "monom Q\<^sub>p_x a n"] assms val_ineq UPQ.cfs_monom by fastforce +qed + +lemma val_val_ring_prod: + assumes "a \ \\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "val (a \\<^bsub>Q\<^sub>p\<^esub> b) \ val b" +proof- + have 0: "val (a \\<^bsub>Q\<^sub>p\<^esub> b) = val a + val b" + using assms val_ring_memE[of a] val_mult + by blast + have 1: " val a \ 0" + using assms + by (simp add: val_ring_memE) + then show ?thesis + using assms 0 + by simp +qed + +lemma val_val_ring_prod': + assumes "a \ \\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "val (b \\<^bsub>Q\<^sub>p\<^esub> a) \ val b" + using val_val_ring_prod[of a b] + by (simp add: Qp.m_comm val_ring_memE assms(1) assms(2)) + +lemma val_ring_nat_pow_closed: + assumes "a \ \\<^sub>p" + shows "(a[^](n::nat)) \ \\<^sub>p" + apply(induction n) + apply auto[1] + using Qp.inv_one Z\<^sub>p_mem apply blast + by (metis Qp.nat_pow_Suc Qp.nat_pow_closed val_ring_memE assms image_eqI inc_of_prod to_Zp_closed to_Zp_inc to_Zp_mult) + +lemma val_ringI: + assumes "a \ carrier Q\<^sub>p" + assumes "val a \0" + shows " a \ \\<^sub>p" + apply(rule val_ring_val_criterion) + using assms by auto + +notation UPQ.to_fun (infixl\\\ 70) + +lemma val_gauss_norm_eval: + assumes "g \ carrier Q\<^sub>p_x" + assumes "a \ \\<^sub>p" + shows "val (g \ a) \ gauss_norm g" +proof- + have 0: "g\a = (\\<^bsub>Q\<^sub>p\<^esub>i\{..degree g}. (g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i))" + using val_ring_memE assms to_fun_formula[of g a] by auto + + have 1: "(\i. g i \\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) \ {..degree g} \ carrier Q\<^sub>p" + using assms + by (meson Pi_I val_ring_memE cfs_closed monom_term_car) + then have 2: "val (g\a) \ Min (val ` (\ i. ((g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i))) ` {..degree g})" + using 0 finsum_val_ultrametric[of "\ i. ((g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i))" "{..degree g}" ] + by (metis finite_atMost not_empty_eq_Iic_eq_empty) + have 3: "\ i. val ((g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) = val (g i) + val (a[^]i)" + using assms val_mult + by (simp add: val_ring_memE UPQ.cfs_closed) + have 4: "\ i. val ((g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) \ val (g i)" + proof- + fix i + show "val ((g i)\\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) \ val (g i)" + using val_val_ring_prod'[of "a[^]i" "g i" ] + assms(1) assms(2) val_ring_nat_pow_closed cfs_closed + by simp + qed + have "Min (val ` (\i. g i \\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) ` {..degree g}) \ Min ((\i. val (g i)) ` {..degree g})" + using Min_mono'[of "{..degree g}" "\i. val (g i)" "\i. val (g i \\<^bsub>Q\<^sub>p\<^esub> (a[^]i))" ] 4 2 + by (metis finite_atMost image_image) + then have "Min (val ` (\i. g i \\<^bsub>Q\<^sub>p\<^esub> (a[^]i)) ` {..degree g}) \ Min (val ` g ` {..degree g})" + by (metis image_image) + then have "val (g\a) \ Min (val ` g ` {..degree g})" + using 2 + by (meson atMost_iff atMost_subset_iff in_mono) + then show ?thesis + by (simp add: \val (g\a) \ Min (val ` g ` {..degree g})\ gauss_norm_def) +qed + +lemma positive_gauss_norm_eval: + assumes "g \ carrier Q\<^sub>p_x" + assumes "gauss_norm g \ 0" + assumes "a \ \\<^sub>p" + shows "(g\a) \ \\<^sub>p" + apply(rule val_ring_val_criterion[of "g\a"]) + using assms val_ring_memE + using UPQ.to_fun_closed apply blast + using assms val_gauss_norm_eval[of g a] by auto + +lemma positive_gauss_norm_valuation_ring_coeffs: + assumes "g \ carrier Q\<^sub>p_x" + assumes "gauss_norm g \ 0" + shows "g n \ \\<^sub>p" + apply(rule val_ringI) + using cfs_closed assms(1) apply blast + using gauss_normE[of g n] assms by auto + +lemma val_ring_cfs_imp_nonneg_gauss_norm: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\n. g n \ \\<^sub>p" + shows "gauss_norm g \ 0" + by(rule gauss_norm_geqI, rule assms, rule val_ring_memE, rule assms) + +lemma val_of_add_pow: + assumes "a \ carrier Q\<^sub>p" + shows "val ([(n::nat)]\a) \ val a" +proof- + have 0: "[(n::nat)]\a = ([n]\\)\a" + using assms Qp.add_pow_ldistr Qp.cring_simprules(12) Qp.one_closed by presburger + have 1: "val ([(n::nat)]\a) = val ([n]\\) + val a" + unfolding 0 by(rule val_mult, simp, rule assms) + show ?thesis unfolding 1 using assms + by (simp add: val_of_nat_inc) +qed + +lemma gauss_norm_pderiv: + assumes "g \ carrier (UP Q\<^sub>p)" + shows "gauss_norm g \ gauss_norm (pderiv g)" + apply(rule gauss_norm_geqI) + using UPQ.pderiv_closed assms apply blast + using gauss_normE pderiv_cfs val_of_add_pow + by (smt UPQ.cfs_closed assms dual_order.trans) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Mapping Polynomials with Value Ring Coefficients to Polynomials over $\mathbb{Z}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition to_Zp_poly where +"to_Zp_poly g = (\n. to_Zp (g n))" + +lemma to_Zp_poly_closed: + assumes "g \ carrier Q\<^sub>p_x" + assumes "gauss_norm g \ 0" + shows "to_Zp_poly g \ carrier (UP Z\<^sub>p)" +proof- + have "to_Zp_poly g \ up Z\<^sub>p" + apply(rule mem_upI) + unfolding to_Zp_poly_def + using cfs_closed[of g ] assms(1) to_Zp_closed[of ] apply blast + proof- + have "\n. bound \\<^bsub>Q\<^sub>p\<^esub> n g" + using UPQ.deg_leE assms(1) by auto + then obtain n where n_def: " bound \\<^bsub>Q\<^sub>p\<^esub> n g" + by blast + then have " bound \\<^bsub>Z\<^sub>p\<^esub> n (\n. to_Zp (g n))" + unfolding bound_def + by (simp add: to_Zp_zero) + then show "\n. bound \\<^bsub>Z\<^sub>p\<^esub> n (\n. to_Zp (g n))" + by blast + qed + then show ?thesis using UP_def[of Z\<^sub>p] + by simp +qed + +definition poly_inc where +"poly_inc g = (\n::nat. \ (g n))" + +lemma poly_inc_closed: + assumes "g \ carrier (UP Z\<^sub>p)" + shows "poly_inc g \ carrier Q\<^sub>p_x" +proof- + have "poly_inc g \ up Q\<^sub>p" + proof(rule mem_upI) + show "\n. poly_inc g n \ carrier Q\<^sub>p" + proof- fix n + have "g n \ carrier Z\<^sub>p" + using assms UP_def + by (simp add: UP_def mem_upD) + then show "poly_inc g n \ carrier Q\<^sub>p" + using assms poly_inc_def[of g] inc_def[of "g n" ] inc_closed + by force + qed + show "\n. bound \\<^bsub>Q\<^sub>p\<^esub> n (poly_inc g)" + proof- + obtain n where n_def: " bound \\<^bsub>Z\<^sub>p\<^esub> n g" + using assms bound_def[of "\\<^bsub>Z\<^sub>p\<^esub>" _ g]Zp.cring_axioms UP_cring.deg_leE[of Z\<^sub>p g] + unfolding UP_cring_def + by metis + then have " bound \\<^bsub>Q\<^sub>p\<^esub> n (poly_inc g)" + unfolding poly_inc_def bound_def + by (metis Qp.nat_inc_zero Zp.nat_inc_zero inc_of_nat) + then show ?thesis by blast + qed + qed + then show ?thesis + by (simp add: \poly_inc g \ up Q\<^sub>p\ UP_def) +qed + +lemma poly_inc_inverse_right: + assumes "g \ carrier (UP Z\<^sub>p)" + shows "to_Zp_poly (poly_inc g) = g" +proof- + have 0: "\n. g n \ carrier Z\<^sub>p" + by (simp add: Zp.cfs_closed assms) + show ?thesis + unfolding to_Zp_poly_def poly_inc_def + proof + fix n + show "to_Zp (\ (g n)) = g n" + using 0 inc_to_Zp + by auto + qed +qed + +lemma poly_inc_inverse_left: + assumes "g \ carrier Q\<^sub>p_x" + assumes "gauss_norm g \0" + shows "poly_inc (to_Zp_poly g) = g" +proof + fix x + show "poly_inc (to_Zp_poly g) x = g x" + using assms unfolding poly_inc_def to_Zp_poly_def + by (simp add: positive_gauss_norm_valuation_ring_coeffs to_Zp_inc) +qed + +lemma poly_inc_plus: + assumes "f \ carrier (UP Z\<^sub>p)" + assumes "g \ carrier (UP Z\<^sub>p)" + shows "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> g) = poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc g" +proof + fix n + have 0: "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> g) n = \ (f n \\<^bsub>Z\<^sub>p\<^esub> g n)" + unfolding poly_inc_def using assms by auto + have 1: "(poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc g) n = poly_inc f n \ poly_inc g n" + by(rule cfs_add, rule poly_inc_closed, rule assms, rule poly_inc_closed, rule assms) + show "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> g) n = (poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc g) n" + unfolding 0 1 unfolding poly_inc_def + apply(rule inc_of_sum) + using assms apply (simp add: Zp.cfs_closed; fail) + using assms by (simp add: Zp.cfs_closed) +qed + +lemma poly_inc_monom: + assumes "a \ carrier Z\<^sub>p" + shows "poly_inc (monom (UP Z\<^sub>p) a m) = monom (UP Q\<^sub>p) (\ a) m" +proof fix n + show "poly_inc (monom (UP Z\<^sub>p) a m) n = monom (UP Q\<^sub>p) (\ a) m n" + apply(cases "m = n") + using assms cfs_monom[of "\ a"] Zp.cfs_monom[of a] unfolding poly_inc_def + apply (simp add: inc_closed; fail) + using assms cfs_monom[of "\ a"] Zp.cfs_monom[of a] unfolding poly_inc_def + by (metis Qp.nat_mult_zero Zp_nat_inc_zero inc_closed inc_of_nat) +qed + +lemma poly_inc_times: + assumes "f \ carrier (UP Z\<^sub>p)" + assumes "g \ carrier (UP Z\<^sub>p)" + shows "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> g) = poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc g" + apply(rule UP_ring.poly_induct3[of Z\<^sub>p]) + apply (simp add: Zp.is_UP_ring; fail) + using assms apply blast +proof- + fix p q + assume A: "q \ carrier (UP Z\<^sub>p)" "p \ carrier (UP Z\<^sub>p)" + "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> p) = poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc p" + "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> q) = poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc q" + have 0: "(f \\<^bsub>UP Z\<^sub>p\<^esub> (p \\<^bsub>UP Z\<^sub>p\<^esub> q)) = (f \\<^bsub>UP Z\<^sub>p\<^esub> p) \\<^bsub>UP Z\<^sub>p\<^esub> (f \\<^bsub>UP Z\<^sub>p\<^esub> q)" + using assms(1) A + by (simp add: Zp.P.r_distr) + have 1: "poly_inc (p \\<^bsub>UP Z\<^sub>p\<^esub> q) = poly_inc p \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc q" + by(rule poly_inc_plus, rule A, rule A) + show "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> (p \\<^bsub>UP Z\<^sub>p\<^esub> q)) = poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc (p \\<^bsub>UP Z\<^sub>p\<^esub> q)" + unfolding 0 1 using A poly_inc_closed poly_inc_plus + by (simp add: UPQ.P.r_distr assms(1)) +next + fix a fix n::nat + assume A: "a \ carrier Z\<^sub>p" + show "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> monom (UP Z\<^sub>p) a n) = + poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc (monom (UP Z\<^sub>p) a n)" + proof + fix m + show "poly_inc (f \\<^bsub>UP Z\<^sub>p\<^esub> monom (UP Z\<^sub>p) a n) m = + (poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc (monom (UP Z\<^sub>p) a n)) m" + proof(cases "m < n") + case True + have T0: "(f \\<^bsub>UP Z\<^sub>p\<^esub> monom (UP Z\<^sub>p) a n) m = \\<^bsub>Z\<^sub>p\<^esub>" + using True Zp.cfs_monom_mult[of f a m n] A assms + by blast + have T1: "poly_inc (monom (UP Z\<^sub>p) a n) = (monom (UP Q\<^sub>p) (\ a) n)" + by(rule poly_inc_monom , rule A) + show ?thesis + unfolding T0 T1 using True + by (metis A Q\<^sub>p_def T0 UPQ.cfs_monom_mult Zp_def assms(1) inc_closed padic_fields.to_Zp_zero padic_fields_axioms poly_inc_closed poly_inc_def to_Zp_inc zero_in_val_ring) + next + case False + then have F0: "m \ n" + using False by simp + have F1: "(f \\<^bsub>UP Z\<^sub>p\<^esub> monom (UP Z\<^sub>p) a n) m = a \\<^bsub>Z\<^sub>p\<^esub> f (m - n)" + using Zp.cfs_monom_mult_l' F0 A assms by simp + have F2: "poly_inc (monom (UP Z\<^sub>p) a n) = monom (UP Q\<^sub>p) (\ a) n " + by(rule poly_inc_monom, rule A) + have F3: "(poly_inc f \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc (monom (UP Z\<^sub>p) a n)) m + = (\ a) \ (poly_inc f (m -n))" + using UPQ.cfs_monom_mult_l' F0 A assms poly_inc_closed + by (simp add: F2 inc_closed) + show ?thesis + unfolding F3 unfolding poly_inc_def F1 + apply(rule inc_of_prod, rule A) + using assms Zp.cfs_closed by blast + qed + qed +qed + +lemma poly_inc_one: +"poly_inc (\\<^bsub>UP Z\<^sub>p\<^esub>) = \\<^bsub>UP Q\<^sub>p\<^esub>" +apply(rule ext) + unfolding poly_inc_def + using inc_of_one inc_of_zero + by simp + +lemma poly_inc_zero: +"poly_inc (\\<^bsub>UP Z\<^sub>p\<^esub>) = \\<^bsub>UP Q\<^sub>p\<^esub>" +apply(rule ext) + unfolding poly_inc_def + using inc_of_one inc_of_zero + by simp + +lemma poly_inc_hom: +"poly_inc \ ring_hom (UP Z\<^sub>p) (UP Q\<^sub>p)" + apply(rule ring_hom_memI) + apply(rule poly_inc_closed, blast) + apply(rule poly_inc_times, blast, blast) + apply(rule poly_inc_plus, blast, blast) + by(rule poly_inc_one) + +lemma poly_inc_as_poly_lift_hom: + assumes "f \ carrier (UP Z\<^sub>p)" + shows "poly_inc f = poly_lift_hom Z\<^sub>p Q\<^sub>p \ f" + apply(rule ext) + unfolding poly_inc_def + using Zp.poly_lift_hom_cf[of Q\<^sub>p \ f] assms UPQ.R_cring local.inc_is_hom + by blast + +lemma poly_inc_eval: + assumes "g \ carrier (UP Z\<^sub>p)" + assumes "a \ carrier Z\<^sub>p" + shows "to_function Q\<^sub>p (poly_inc g) (\ a) = \ (to_function Z\<^sub>p g a)" +proof- + have 0: "poly_inc g = poly_lift_hom Z\<^sub>p Q\<^sub>p \ g" + using assms poly_inc_as_poly_lift_hom[of g] by blast + have 1: "to_function Q\<^sub>p (poly_lift_hom Z\<^sub>p Q\<^sub>p \ g) (\ a) = \ (to_function Z\<^sub>p g a)" + using Zp.poly_lift_hom_eval[of Q\<^sub>p \ g a] assms inc_is_hom + unfolding to_fun_def Zp.to_fun_def + using UPQ.R_cring by blast + show ?thesis unfolding 0 1 + by blast +qed + +lemma val_ring_poly_eval: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "\ i. f i \ \\<^sub>p" + shows "\x. x \ \\<^sub>p \ f \ x \ \\<^sub>p" + apply(rule positive_gauss_norm_eval, rule assms) + apply(rule val_ring_cfs_imp_nonneg_gauss_norm) + using assms by auto + +lemma Zp_res_of_pow: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ carrier Z\<^sub>p" + assumes "a n = b n" + shows "(a[^]\<^bsub>Z\<^sub>p\<^esub>(k::nat)) n = (b[^]\<^bsub>Z\<^sub>p\<^esub>(k::nat)) n" + apply(induction k) + using assms Group.nat_pow_0 to_Zp_one apply metis + using Zp.geometric_series_id[of a b] Zp_residue_mult_zero(1) assms(1) assms(2) assms(3) + pow_closed res_diff_zero_fact'' res_diff_zero_fact(1) by metis + +lemma to_Zp_nat_pow: + assumes "a \ \\<^sub>p" + shows "to_Zp (a[^](n::nat)) = (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(n::nat)" + apply(induction n) + using assms Group.nat_pow_0 to_Zp_one apply metis + using assms to_Zp_mult[of a] Qp.m_comm Qp.nat_pow_Suc val_ring_memE pow_suc to_Zp_closed val_ring_nat_pow_closed + by metis + +lemma to_Zp_res_of_pow: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "to_Zp a n = to_Zp b n" + shows "to_Zp (a[^](k::nat)) n = to_Zp (b[^](k::nat)) n" + using assms val_ring_memE Zp_res_of_pow to_Zp_closed to_Zp_nat_pow by presburger + +lemma poly_eval_cong: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\i. g i \ \\<^sub>p" + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "to_Zp a k = to_Zp b k" + shows "to_Zp (g \ a) k = to_Zp (g \ b) k" +proof- + have "(\i. g i \ \\<^sub>p) \ to_Zp (g \ a) k = to_Zp (g \ b) k" + proof(rule UPQ.poly_induct[of g]) + show " g \ carrier (UP Q\<^sub>p)" + using assms by blast + show "\p. p \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p p = 0 \ (\i. p i \ \\<^sub>p) \ to_Zp (p \ a) k = to_Zp (p \ b) k" + proof fix p assume A: "p \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p p = 0" "\i. p i \ \\<^sub>p" + obtain c where c_def: "c \ carrier Q\<^sub>p \ p = up_ring.monom (UP Q\<^sub>p) c 0" + using A + by (metis UPQ.zcf_degree_zero UPQ.cfs_closed UPQ.trms_of_deg_leq_0 UPQ.trms_of_deg_leq_degree_f) + have p_eq: "p = up_ring.monom (UP Q\<^sub>p) c 0" + using c_def by blast + have p_cfs: "p 0 = c" + unfolding p_eq using c_def UP_ring.cfs_monom[of Q\<^sub>p c 0 0] UPQ.P_is_UP_ring by presburger + have c_closed: "c \ \\<^sub>p" + using p_cfs A(3) by blast + have 0: "(p \ a) = c" + unfolding p_eq using c_def assms by (meson UPQ.to_fun_const val_ring_memE(2)) + have 1: "(p \ b) = c" + unfolding p_eq using c_def assms UPQ.to_fun_const val_ring_memE(2) by presburger + show " to_Zp (p \ a) k = to_Zp (p \ b) k" + unfolding 0 1 by blast + qed + show "\p. (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p) \ to_Zp (q \ a) k = to_Zp (q \ b) k) \ + p \ carrier (UP Q\<^sub>p) \ 0 < deg Q\<^sub>p p \ (\i. p i \ \\<^sub>p) \ to_Zp (p \ a) k = to_Zp (p \ b) k" + proof + fix p assume A: "(\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p) \ to_Zp (q \ a) k = to_Zp (q \ b) k)" + "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p " " \i. p i \ \\<^sub>p" + obtain q where q_def: "q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ p = UPQ.ltrm p \\<^bsub>UP Q\<^sub>p\<^esub>q" + by (metis A(2) A(3) UPQ.ltrm_closed UPQ.ltrm_decomp UPQ.UP_a_comm) + have 0: "\i. p i = q i \ UPQ.ltrm p i" + using q_def A + by (metis Qp.a_ac(2) UPQ.ltrm_closed UPQ.UP_car_memE(1) UPQ.cfs_add) + have 1: "\i. q i \ \\<^sub>p" + proof fix i + show "q i \ \\<^sub>p" + apply(cases "i < deg Q\<^sub>p p") + using 0[of i] A(4) A(2) q_def + using UPQ.ltrm_closed UPQ.P.a_ac(2) UPQ.trunc_cfs UPQ.trunc_closed UPQ.trunc_simps(1) + apply (metis Qp.r_zero UPQ.ltrm_cfs UPQ.cfs_closed UPQ.deg_leE) + using q_def + by (metis (no_types, opaque_lifting) A(2) A(4) UPQ.P.add.m_closed UPQ.coeff_of_sum_diff_degree0 UPQ.deg_leE UPQ.equal_deg_sum UPQ.equal_deg_sum' \\thesis. (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ p = up_ring.monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p) \\<^bsub>UP Q\<^sub>p\<^esub> q \ thesis) \ thesis\ lessI linorder_neqE_nat) + qed + have 2: "UPQ.lcf p \ \\<^sub>p" + using A(4) by blast + have 3: "UPQ.ltrm p \ a = UPQ.lcf p \ a[^] deg Q\<^sub>p p" + apply(rule UP_cring.to_fun_monom) unfolding UP_cring_def + using Qp.cring apply blast + using A UPQ.lcf_closed apply blast + using assms val_ring_memE(2) by blast + have 4: "UPQ.ltrm p \ b = UPQ.lcf p \ b[^] deg Q\<^sub>p p" + apply(rule UP_cring.to_fun_monom) unfolding UP_cring_def + using Qp.cring apply blast + using A UPQ.lcf_closed apply blast + using assms val_ring_memE(2) by blast + have p_eq: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.ltrm p" + using q_def by (metis A(2) UPQ.ltrm_closed UPQ.UP_a_comm) + have 5: "p \ a = q \ a \ UPQ.lcf p \ a[^] deg Q\<^sub>p p" + using assms val_ring_memE(2) p_eq q_def UPQ.to_fun_plus[of q "UPQ.ltrm p" a] + by (metis "3" A(2) UPQ.ltrm_closed UPQ.to_fun_plus) + have 6: "p \ b = q \ b \ UPQ.lcf p \ b[^] deg Q\<^sub>p p" + using assms val_ring_memE(2) p_eq q_def UPQ.to_fun_plus[of q "UPQ.ltrm p" a] + by (metis "4" A(2) UPQ.ltrm_closed UPQ.to_fun_plus) + have 7: "UPQ.lcf p \ b[^] deg Q\<^sub>p p \ \\<^sub>p" + apply(rule val_ring_times_closed) + using "2" apply linarith + by(rule val_ring_nat_pow_closed, rule assms) + have 8: "UPQ.lcf p \ a[^] deg Q\<^sub>p p \ \\<^sub>p" + apply(rule val_ring_times_closed) + using "2" apply linarith + by(rule val_ring_nat_pow_closed, rule assms) + have 9: "q \ a \ \\<^sub>p" + using q_def 1 assms(3) val_ring_poly_eval by blast + have 10: "q \ b \ \\<^sub>p" + using q_def 1 assms(4) val_ring_poly_eval by blast + have 11: "to_Zp (p \ a) = to_Zp (q \ a) \\<^bsub>Z\<^sub>p\<^esub> to_Zp (UPQ.ltrm p \ a)" + using 5 8 9 to_Zp_add 3 by presburger + have 12: "to_Zp (p \ b) = to_Zp (q \ b) \\<^bsub>Z\<^sub>p\<^esub> to_Zp (UPQ.ltrm p \ b)" + using 6 10 7 to_Zp_add 4 by presburger + have 13: "to_Zp (p \ a) k = to_Zp (q \ a) k \\<^bsub>Zp_res_ring k\<^esub> to_Zp (UPQ.ltrm p \ a) k" + unfolding 11 using residue_of_sum by blast + have 14: "to_Zp (p \ b) k = to_Zp (q \ b) k \\<^bsub>Zp_res_ring k\<^esub> to_Zp (UPQ.ltrm p \ b) k" + unfolding 12 using residue_of_sum by blast + have 15: "to_Zp (UPQ.ltrm p \ a) k = to_Zp (UPQ.ltrm p \ b) k" + proof(cases "k = 0") + case True + have T0: "to_Zp (UPQ.ltrm p \ a) \ carrier Z\<^sub>p" + unfolding 3 using 8 to_Zp_closed val_ring_memE(2) by blast + have T1: "to_Zp (UPQ.ltrm p \ b) \ carrier Z\<^sub>p" + unfolding 4 using 7 to_Zp_closed val_ring_memE(2) by blast + show ?thesis unfolding True using T0 T1 padic_integers.p_res_ring_0 + by (metis p_res_ring_0' residues_closed) + next + case False + have k_pos: "k > 0" + using False by presburger + have 150: "to_Zp (p (deg Q\<^sub>p p) \ a [^] deg Q\<^sub>p p) = to_Zp (p (deg Q\<^sub>p p)) \\<^bsub>Z\<^sub>p\<^esub> to_Zp( a [^] deg Q\<^sub>p p)" + apply(rule to_Zp_mult) + using "2" apply blast + by(rule val_ring_nat_pow_closed, rule assms) + have 151: "to_Zp (p (deg Q\<^sub>p p) \ b [^] deg Q\<^sub>p p) = to_Zp (p (deg Q\<^sub>p p)) \\<^bsub>Z\<^sub>p\<^esub> to_Zp( b [^] deg Q\<^sub>p p)" + apply(rule to_Zp_mult) + using "2" apply blast + by(rule val_ring_nat_pow_closed, rule assms) + have 152: "to_Zp (p (deg Q\<^sub>p p) \ a [^] deg Q\<^sub>p p) k = to_Zp (p (deg Q\<^sub>p p)) k \\<^bsub>Zp_res_ring k\<^esub> to_Zp( a [^] deg Q\<^sub>p p) k" + unfolding 150 using residue_of_prod by blast + have 153: "to_Zp (p (deg Q\<^sub>p p) \ b [^] deg Q\<^sub>p p) k = to_Zp (p (deg Q\<^sub>p p)) k \\<^bsub>Zp_res_ring k\<^esub> to_Zp( b [^] deg Q\<^sub>p p) k" + unfolding 151 using residue_of_prod by blast + have 154: "to_Zp( a [^] deg Q\<^sub>p p) k = to_Zp a k [^]\<^bsub>Zp_res_ring k\<^esub> deg Q\<^sub>p p" + proof- + have 01: "\m::nat. to_Zp (a[^]m) k = to_Zp a k [^]\<^bsub>Zp_res_ring k\<^esub> m" + proof- + fix m::nat show "to_Zp (a [^] m) k = to_Zp a k [^]\<^bsub>Zp_res_ring k\<^esub> m" + proof- + have 00: "to_Zp (a[^]m) = to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub> m" + using assms to_Zp_nat_pow[of a "m"] by blast + have 01: "to_Zp a \ carrier Z\<^sub>p" + using assms to_Zp_closed val_ring_memE(2) by blast + have 02: "to_Zp a k \ carrier (Zp_res_ring k)" + using 01 residues_closed by blast + have 03: "cring (Zp_res_ring k)" + using k_pos padic_integers.R_cring padic_integers_axioms by blast + have 01: "(to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub> m) k = (to_Zp a) k [^]\<^bsub>Zp_res_ring k\<^esub> m" + apply(induction m) + using 01 02 apply (metis Group.nat_pow_0 k_pos residue_of_one(1)) + using residue_of_prod[of "to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub> m" "to_Zp a" k] 01 02 03 + proof - + fix ma :: nat + assume "(to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub> ma) k = to_Zp a k [^]\<^bsub>Zp_res_ring k\<^esub> ma" + then show "(to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub> Suc ma) k = to_Zp a k [^]\<^bsub>Zp_res_ring k\<^esub> Suc ma" + by (metis (no_types) Group.nat_pow_Suc residue_of_prod) + qed + show ?thesis unfolding 00 01 by blast + qed + qed + thus ?thesis by blast + qed + have 155: "to_Zp( b [^] deg Q\<^sub>p p) k = to_Zp b k [^]\<^bsub>Zp_res_ring k\<^esub> deg Q\<^sub>p p" + using assms by (metis "154" to_Zp_res_of_pow) + show ?thesis + unfolding 3 4 152 153 154 155 assms by blast + qed + show "to_Zp (p \ a) k = to_Zp (p \ b) k" + unfolding 13 14 15 using A 1 q_def by presburger + qed + qed + thus ?thesis using assms by blast +qed + +lemma to_Zp_poly_eval: + assumes "g \ carrier Q\<^sub>p_x" + assumes "gauss_norm g \ 0" + assumes "a \ \\<^sub>p" + shows "to_Zp (to_function Q\<^sub>p g a) = to_function Z\<^sub>p (to_Zp_poly g) (to_Zp a)" +proof- + obtain h where h_def: "h = to_Zp_poly g" + by blast + obtain b where b_def: "b = to_Zp a" + by blast + have h_poly_inc: "poly_inc h = g" + unfolding h_def using assms + by (simp add: poly_inc_inverse_left) + have b_inc: "\ b = a" + unfolding b_def using assms + by (simp add: to_Zp_inc) + have h_closed: "h \ carrier (UP Z\<^sub>p)" + unfolding h_def using assms + by (simp add: to_Zp_poly_closed) + have b_closed: "b \ carrier Z\<^sub>p" + unfolding b_def using assms + by (simp add: to_Zp_closed val_ring_memE) + have 0: "to_function Q\<^sub>p (poly_inc h) (\ b) = \ (to_function Z\<^sub>p h b)" + apply(rule poly_inc_eval) + using h_def assms apply (simp add: to_Zp_poly_closed; fail) + unfolding b_def using assms + by (simp add: to_Zp_closed val_ring_memE) + have 1: "to_Zp (to_function Q\<^sub>p (poly_inc h) (\ b)) = to_function Z\<^sub>p h b" + unfolding 0 + using h_closed b_closed Zp.to_fun_closed Zp.to_fun_def inc_to_Zp by auto + show ?thesis + using 1 unfolding h_poly_inc b_inc + unfolding h_def b_def by blast +qed + +lemma poly_eval_equal_val: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\x. g x \ \\<^sub>p" + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "val (g \ a) < eint n" + assumes "to_Zp a n = to_Zp b n" + shows "val (g \ b) = val (g \ a)" +proof- + have "(\x. g x \ \\<^sub>p) \ to_Zp (g \ b) n = to_Zp (g \ a) n" + proof(rule poly_induct[of g]) + show "g \ carrier (UP Q\<^sub>p)" + by (simp add: assms(1)) + show "\p. p \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p p = 0 \ (\x. p x \ \\<^sub>p) \ to_Zp (p \ b) n = to_Zp (p \ a) n" + proof fix p assume A: "p \ carrier (UP Q\<^sub>p)" " deg Q\<^sub>p p = 0 " "\x. p x \ \\<^sub>p " + show "to_Zp (p \ b) n = to_Zp (p \ a) n" + using A by (metis val_ring_memE UPQ.to_fun_ctrm UPQ.trms_of_deg_leq_0 UPQ.trms_of_deg_leq_degree_f assms(3) assms(4)) + qed + show "\p. (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\x. q x \ \\<^sub>p) \ to_Zp (q \ b) n = to_Zp (q \ a) n) \ + p \ carrier (UP Q\<^sub>p) \ 0 < deg Q\<^sub>p p \ (\x. p x \ \\<^sub>p) \ to_Zp (p \ b) n = to_Zp (p \ a) n" + proof fix p assume IH: "(\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\x. q x \ \\<^sub>p) \ to_Zp (q \ b) n = to_Zp (q \ a) n)" + assume A: "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p" "\x. p x \ \\<^sub>p" + show "to_Zp (p \ b) n = to_Zp (p \ a) n" + proof- + obtain q where q_def: "q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ + p = q \\<^bsub>UP Q\<^sub>p\<^esub> ltrm p" + using A by (meson UPQ.ltrm_decomp) + have p_eq: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> ltrm p" + using q_def by blast + have "\x. q x \ \\<^sub>p" proof fix x + have px: "p x = (q \\<^bsub>UP Q\<^sub>p\<^esub> ltrm p) x" + using p_eq by simp + show "q x \ \\<^sub>p" + proof(cases "x \ deg Q\<^sub>p q") + case True + then have "p x = q x" + unfolding px using q_def A + by (smt UPQ.ltrm_closed UPQ.P.add.right_cancel UPQ.coeff_of_sum_diff_degree0 UPQ.deg_ltrm UPQ.trunc_cfs UPQ.trunc_closed UPQ.trunc_simps(1) less_eq_Suc_le nat_neq_iff not_less_eq_eq) + then show ?thesis using A + by blast + next + case False + then show ?thesis + using q_def UPQ.deg_eqI eq_imp_le nat_le_linear zero_in_val_ring + by (metis (no_types, lifting) UPQ.coeff_simp UPQ.deg_belowI) + qed + qed + then have 0: " to_Zp (q \ b) n = to_Zp (q \ a) n" + using IH q_def by blast + have 1: "to_Zp (ltrm p \ b) n = to_Zp (ltrm p \ a) n" + proof- + have 10: "(ltrm p \ b) = (p (deg Q\<^sub>p p)) \ b[^] (deg Q\<^sub>p p)" + using assms A by (meson val_ring_memE UPQ.to_fun_monom) + have 11: "(ltrm p \ a) = (p (deg Q\<^sub>p p)) \ a[^] (deg Q\<^sub>p p)" + using assms A by (meson val_ring_memE UPQ.to_fun_monom) + have 12: "to_Zp (b[^] (deg Q\<^sub>p p)) n = to_Zp (a[^] (deg Q\<^sub>p p)) n" + using to_Zp_res_of_pow assms by metis + have 13: "p (deg Q\<^sub>p p) \ \\<^sub>p" + using A(3) by blast + have 14: "b[^] (deg Q\<^sub>p p) \ \\<^sub>p" + using assms(4) val_ring_nat_pow_closed by blast + have 15: "a[^] (deg Q\<^sub>p p) \ \\<^sub>p" + using assms(3) val_ring_nat_pow_closed by blast + have 16: "(ltrm p \ b) \ \\<^sub>p" + by (simp add: "10" "13" "14" val_ring_times_closed) + have 17: "to_Zp (ltrm p \ b) n = to_Zp (p (deg Q\<^sub>p p)) n \\<^bsub>Zp_res_ring n\<^esub> to_Zp (b[^] (deg Q\<^sub>p p)) n" + using 10 13 14 15 16 assms residue_of_prod to_Zp_mult by presburger + have 18: "(ltrm p \ a) \ \\<^sub>p" + by (simp add: "11" "15" A(3) val_ring_times_closed) + have 19: "to_Zp (ltrm p \ a) n = to_Zp (p (deg Q\<^sub>p p)) n \\<^bsub>Zp_res_ring n\<^esub> to_Zp (a[^] (deg Q\<^sub>p p)) n" + using 10 13 14 15 16 17 18 assms residue_of_prod to_Zp_mult 11 by presburger + show ?thesis using 12 17 19 by presburger + qed + have 2: "p (deg Q\<^sub>p p) \ \\<^sub>p" + using A(3) by blast + have 3: "(ltrm p \ b) \ \\<^sub>p" + using 2 assms + by (metis A(1) Q\<^sub>p_def val_ring_memE val_ring_memE UPQ.ltrm_closed Zp_def \_def + gauss_norm_monom padic_fields.positive_gauss_norm_eval padic_fields_axioms) + have 4: "(ltrm p \ a) \ \\<^sub>p" + using 2 assms + by (metis A(1) Q\<^sub>p_def val_ring_memE val_ring_memE UPQ.ltrm_closed Zp_def \_def + gauss_norm_monom padic_fields.positive_gauss_norm_eval padic_fields_axioms) + have 5: "(q \ b) \ \\<^sub>p" + using \\x. q x \ \\<^sub>p\ assms(4) q_def + by (metis gauss_norm_coeff_norm positive_gauss_norm_eval val_ring_memE(1)) + have 6: "(q \ a) \ \\<^sub>p" + using \\x. q x \ \\<^sub>p\ assms(3) q_def + by (metis gauss_norm_coeff_norm positive_gauss_norm_eval val_ring_memE(1)) + have 7: "to_Zp (p \ b) = to_Zp (ltrm p \ b) \\<^bsub>Z\<^sub>p\<^esub> to_Zp (q \ b)" + using 5 3 q_def by (metis (no_types, lifting) A(1) val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus add_comm assms(4) to_Zp_add) + have 8: "to_Zp (p \ a) = to_Zp (ltrm p \ a) \\<^bsub>Z\<^sub>p\<^esub> to_Zp (q \ a)" + using 4 6 q_def by (metis (no_types, lifting) A(1) val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus add_comm assms(3) to_Zp_add) + have 9: "to_Zp (p \ b) \ carrier Z\<^sub>p" + using A assms by (meson val_ring_memE UPQ.to_fun_closed to_Zp_closed) + have 10: "to_Zp (p \ a) \ carrier Z\<^sub>p" + using A assms val_ring_memE UPQ.to_fun_closed to_Zp_closed by presburger + have 11: "to_Zp (p \ b) n = to_Zp (ltrm p \ b) n \\<^bsub>Zp_res_ring n\<^esub> to_Zp (q \ b) n" + using 7 9 5 3 residue_of_sum by presburger + have 12: "to_Zp (p \ a) n = to_Zp (ltrm p \ a) n \\<^bsub>Zp_res_ring n\<^esub> to_Zp (q \ a) n" + using 8 6 4 residue_of_sum by presburger + show ?thesis using 0 11 12 q_def assms + using "1" by presburger + qed + qed + qed + have "(\x. g x \ \\<^sub>p) " + using assms by blast + hence 0: "to_Zp (g \ b) n = to_Zp (g \ a) n" + using \(\x. g x \ \\<^sub>p) \ to_Zp (g \ b) n = to_Zp (g \ a) n\ by blast + have 1: "g \ a \ \\<^sub>p" + using assms(1) assms(2) assms(3) + by (metis gauss_norm_coeff_norm positive_gauss_norm_eval val_ring_memE(1)) + have 2: "g \ b \ \\<^sub>p" + using assms(1) assms(2) assms(4) + by (metis gauss_norm_coeff_norm positive_gauss_norm_eval val_ring_memE(1)) + have 3: "val (g \ b) < eint n" + proof- + have P0: "to_Zp (g \ a) \ carrier Z\<^sub>p" + using 1 val_ring_memE to_Zp_closed by blast + have P1: "to_Zp (g \ b) \ carrier Z\<^sub>p" + using 2 val_ring_memE to_Zp_closed by blast + have P2: "val_Zp (to_Zp (g \ a)) < n" + using 1 assms to_Zp_val by presburger + have P3: "to_Zp (g \ a) \ \\<^bsub>Z\<^sub>p\<^esub>" + using P2 P0 unfolding val_Zp_def by (metis P2 infinity_ilessE val_Zp_def) + have P4: "(to_Zp (g \ a)) n \ 0" + using 1 P2 P3 above_ord_nonzero[of "to_Zp (g \ a)" n] + by (metis P0 eint.inject less_eintE val_ord_Zp) + then have "to_Zp (g \ b) n \ 0" + using 0 by linarith + then have "val_Zp (to_Zp (g \ b)) < n" + using P1 P0 + by (smt below_val_Zp_zero eint_ile eint_ord_simps(1) eint_ord_simps(2) nonzero_imp_ex_nonzero_res residue_of_zero(2) zero_below_val_Zp) + then show ?thesis using 2 + by (metis to_Zp_val) + qed + thus ?thesis using 0 1 2 assms val_ring_equal_res_imp_equal_val[of "g \ b" "g \ a" n] by blast +qed + +lemma to_Zp_poly_monom: + assumes "a \ \\<^sub>p" + shows "to_Zp_poly (monom (UP Q\<^sub>p) a n) = monom (UP Z\<^sub>p) (to_Zp a) n" + unfolding to_Zp_poly_def + apply(rule ext) + using assms cfs_monom[of a n] Zp.cfs_monom[of "to_Zp a" n] + by (simp add: to_Zp_closed to_Zp_zero val_ring_memE(2)) + +lemma to_Zp_poly_add: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "gauss_norm f \ 0" + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "gauss_norm g \ 0" + shows "to_Zp_poly (f \\<^bsub>UP Q\<^sub>p\<^esub> g) = to_Zp_poly f \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly g" +proof- + obtain F where F_def: "F = to_Zp_poly f" + by blast + obtain G where G_def: "G = to_Zp_poly g" + by blast + have F_closed: "F \ carrier (UP Z\<^sub>p)" + unfolding F_def using assms + by (simp add: to_Zp_poly_closed) + have G_closed: "G \ carrier (UP Z\<^sub>p)" + unfolding G_def using assms + by (simp add: to_Zp_poly_closed) + have F_inc: "poly_inc F = f" + using assms unfolding F_def + using poly_inc_inverse_left by blast + have G_inc: "poly_inc G = g" + using assms unfolding G_def + by (simp add: poly_inc_inverse_left) + have 0: "poly_inc (F \\<^bsub>UP Z\<^sub>p\<^esub> G) = poly_inc F \\<^bsub>UP Q\<^sub>p\<^esub> poly_inc G" + using F_closed G_closed + by (simp add: poly_inc_plus) + have 1: "to_Zp_poly (poly_inc (F \\<^bsub>UP Z\<^sub>p\<^esub> G)) = F \\<^bsub>UP Z\<^sub>p\<^esub> G" + using G_closed F_closed + by (simp add: poly_inc_inverse_right) + show ?thesis + using 1 unfolding F_inc G_inc 0 unfolding F_def G_def + by blast +qed + +lemma to_Zp_poly_zero: +"to_Zp_poly (\\<^bsub>UP Q\<^sub>p\<^esub>) = \\<^bsub>UP Z\<^sub>p\<^esub>" + unfolding to_Zp_poly_def + apply(rule ext) + by (simp add: to_Zp_zero) + +lemma to_Zp_poly_one: +"to_Zp_poly (\\<^bsub>UP Q\<^sub>p\<^esub>) = \\<^bsub>UP Z\<^sub>p\<^esub>" + unfolding to_Zp_poly_def + apply(rule ext) + by (metis Zp.UP_one_closed poly_inc_inverse_right poly_inc_one to_Zp_poly_def) + +lemma val_ring_add_pow: + assumes "a \ carrier Q\<^sub>p" + 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 to_Zp_poly_pderiv: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "gauss_norm g \ 0" + shows "to_Zp_poly (pderiv g) = Zp.pderiv (to_Zp_poly g)" +proof- + have 0: "gauss_norm g \ 0 \ to_Zp_poly (pderiv g) = Zp.pderiv (to_Zp_poly g)" + proof(rule poly_induct, rule assms, rule) + fix p + assume A: " p \ carrier (UP Q\<^sub>p)" + "deg Q\<^sub>p p = 0" + "0 \ gauss_norm p" + obtain a where a_def: "a \ \\<^sub>p \ p = monom (UP Q\<^sub>p) a 0" + using A + by (metis UPQ.ltrm_deg_0 positive_gauss_norm_valuation_ring_coeffs) + have p_eq: "p = monom (UP Q\<^sub>p) a 0" + using a_def by blast + have 0: "to_Zp_poly p = monom (UP Z\<^sub>p) (to_Zp a) 0" + unfolding p_eq + apply(rule to_Zp_poly_monom) + using a_def by blast + have 1: "UPQ.pderiv (monom (UP Q\<^sub>p) a 0) = \\<^bsub>UP Q\<^sub>p\<^esub>" + using A(1) A(2) UPQ.pderiv_deg_0 p_eq by blast + have 2: "Zp.pderiv (monom (UP Z\<^sub>p) (to_Zp a) 0) = \\<^bsub>UP Z\<^sub>p\<^esub>" + apply(rule Zp.pderiv_deg_0) + apply(rule Zp.monom_closed, rule to_Zp_closed) + using a_def + apply (simp add: val_ring_memE(2); fail) + apply(cases "to_Zp a = \\<^bsub>Z\<^sub>p\<^esub>") + apply (simp; fail) + apply(rule Zp.deg_monom, blast) + using a_def + by (simp add: to_Zp_closed val_ring_memE(2)) + show "to_Zp_poly (UPQ.pderiv p) = Zp.pderiv (to_Zp_poly p)" + unfolding 0 unfolding p_eq + unfolding 1 2 to_Zp_poly_zero by blast + next + fix p + assume A: "\q. q \ carrier (UP Q\<^sub>p) \ + deg Q\<^sub>p q < deg Q\<^sub>p p \ + 0 \ gauss_norm q \ + to_Zp_poly (UPQ.pderiv q) = Zp.pderiv (to_Zp_poly q)" + "p \ carrier (UP Q\<^sub>p)" + " 0 < deg Q\<^sub>p p" + show "0 \ gauss_norm p \ to_Zp_poly (UPQ.pderiv p) = Zp.pderiv (to_Zp_poly p)" + proof + assume B: "0 \ gauss_norm p" + obtain q where q_def: "q = trunc p" + by blast + have p_eq: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> ltrm p" + by (simp add: A(2) UPQ.trunc_simps(1) q_def) + have q_gauss_norm: "gauss_norm q \ 0" + unfolding q_def + apply(rule gauss_norm_geqI) + using A apply (simp add: UPQ.trunc_closed; fail) + using trunc_cfs[of p] A gauss_normE + proof - + fix n :: nat + have f1: "\ = q (deg Q\<^sub>p p)" + by (simp add: UPQ.deg_leE UPQ.trunc_closed UPQ.trunc_degree \0 < deg Q\<^sub>p p\ \p \ carrier (UP Q\<^sub>p)\ q_def) + have "\n. 0 \ val (p n)" + by (meson B \p \ carrier (UP Q\<^sub>p)\ eint_ord_trans gauss_normE) + then show "0 \ val (Cring_Poly.truncate Q\<^sub>p p n)" + using f1 by (metis (no_types) Qp.nat_mult_zero UPQ.ltrm_closed UPQ.coeff_of_sum_diff_degree0 UPQ.deg_ltrm UPQ.trunc_closed \\n. \p \ carrier (UP Q\<^sub>p); n < deg Q\<^sub>p p\ \ Cring_Poly.truncate Q\<^sub>p p n = p n\ \p \ carrier (UP Q\<^sub>p)\ nat_neq_iff p_eq q_def val_of_nat_inc) + qed + have 0: "to_Zp_poly (UPQ.pderiv q) = Zp.pderiv (to_Zp_poly q)" + using A q_def q_gauss_norm + by (simp add: UPQ.trunc_closed UPQ.trunc_degree) + have 1: "UPQ.pderiv (monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p)) = + monom (UP Q\<^sub>p) ([deg Q\<^sub>p p] \ p (deg Q\<^sub>p p)) (deg Q\<^sub>p p - 1)" + apply(rule pderiv_monom) + using A by (simp add: UPQ.UP_car_memE(1)) + have 2: "Zp.pderiv (monom (UP Z\<^sub>p) (to_Zp (p (deg Q\<^sub>p p))) (deg Q\<^sub>p p)) = + monom (UP Z\<^sub>p) ([deg Q\<^sub>p p] \\<^bsub>Z\<^sub>p\<^esub> to_Zp ( p (deg Q\<^sub>p p))) (deg Q\<^sub>p p - 1)" + using A Zp.pderiv_monom[of "to_Zp ( p (deg Q\<^sub>p p))" "deg Q\<^sub>p p"] + by (simp add: UPQ.lcf_closed to_Zp_closed) + have 3: "to_Zp_poly (UPQ.pderiv (monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p))) = monom (UP Z\<^sub>p) (to_Zp ([deg Q\<^sub>p p] \ p (deg Q\<^sub>p p))) (deg Q\<^sub>p p - 1)" + unfolding 1 apply(rule to_Zp_poly_monom) + apply(rule val_ring_memI) + apply (simp add: A(2) UPQ.UP_car_memE(1); fail) + apply(rule val_ring_add_pow) + using A + apply (simp add: UPQ.lcf_closed; fail) + using B A + by (simp add: positive_gauss_norm_valuation_ring_coeffs val_ring_memE(1)) + have 4: "to_Zp_poly (ltrm p) = monom (UP Z\<^sub>p) (to_Zp (p (deg Q\<^sub>p p))) (deg Q\<^sub>p p)" + apply(rule to_Zp_poly_monom) using A + by (simp add: B positive_gauss_norm_valuation_ring_coeffs) + have 5: "to_Zp_poly (UPQ.pderiv (ltrm p)) = Zp.pderiv (to_Zp_poly (ltrm p))" + unfolding 3 4 2 + by (simp add: A(2) B positive_gauss_norm_valuation_ring_coeffs to_Zp_nat_add_pow) + have 6: "pderiv p = pderiv q \\<^bsub>UP Q\<^sub>p\<^esub> pderiv (ltrm p)" + using p_eq + by (metis A(2) UPQ.ltrm_closed UPQ.pderiv_add UPQ.trunc_closed p_eq q_def) + have 7: "to_Zp_poly p = to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (ltrm p)" + using p_eq + by (metis (no_types, lifting) A(2) B UPQ.ltrm_closed UPQ.cfs_closed UPQ.trunc_closed gauss_norm_monom positive_gauss_norm_valuation_ring_coeffs q_def q_gauss_norm to_Zp_poly_add val_ring_memE(1)) + have 8: "to_Zp_poly (pderiv p) = + to_Zp_poly (UPQ.pderiv q) \\<^bsub>UP Z\<^sub>p\<^esub> + to_Zp_poly (UPQ.pderiv (monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p)))" + unfolding 6 apply(rule to_Zp_poly_add) + apply (simp add: A(2) UPQ.pderiv_closed UPQ.trunc_closed q_def; fail) + apply (metis A(2) UPQ.cfs_closed UPQ.pderiv_cfs UPQ.trunc_closed gauss_norm_coeff_norm positive_gauss_norm_valuation_ring_coeffs q_def q_gauss_norm val_ring_add_pow val_ring_memE(1)) + apply (simp add: A(2) UPQ.UP_car_memE(1) UPQ.pderiv_closed; fail) + apply(rule eint_ord_trans[of _ "gauss_norm (monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p))"]) + apply (simp add: A(2) B UPQ.cfs_closed gauss_norm_monom positive_gauss_norm_valuation_ring_coeffs val_ring_memE(1); fail) + apply(rule gauss_norm_pderiv) + using A(2) UPQ.ltrm_closed by blast + have 9: "Zp.pderiv (to_Zp_poly p) = Zp.pderiv (to_Zp_poly q) \\<^bsub>UP Z\<^sub>p\<^esub> + Zp.pderiv (to_Zp_poly (monom (UP Q\<^sub>p) (p (deg Q\<^sub>p p)) (deg Q\<^sub>p p)))" + unfolding 7 apply(rule Zp.pderiv_add) + apply(rule to_Zp_poly_closed) + apply (simp add: A(2) UPQ.trunc_closed q_def; fail) + apply (simp add: q_gauss_norm; fail) + apply(rule to_Zp_poly_closed) + apply (simp add: A(2) UPQ.UP_car_memE(1); fail) + by (simp add: A(2) B UPQ.cfs_closed gauss_norm_monom positive_gauss_norm_valuation_ring_coeffs val_ring_memE(1)) + show "to_Zp_poly (UPQ.pderiv p) = Zp.pderiv (to_Zp_poly p)" + unfolding 9 8 5 0 by blast + qed + qed + thus ?thesis using assms by blast +qed + +lemma val_p_int_pow: +"val (\[^]k) = eint (k)" + by (simp add: ord_p_pow_int p_intpow_closed(2)) + +definition int_gauss_norm where +"int_gauss_norm g = (SOME n::int. eint n = gauss_norm g)" + +lemma int_gauss_norm_eq: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + shows "eint (int_gauss_norm g) = gauss_norm g" +proof- + have 0: "gauss_norm g < \" + using assms by (simp add: gauss_norm_prop) + then show ?thesis unfolding int_gauss_norm_def + using assms + by fastforce +qed + +lemma int_gauss_norm_smult: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + assumes "a \ nonzero Q\<^sub>p" + shows "int_gauss_norm (a \\<^bsub>UP Q\<^sub>p\<^esub> g) = ord a + int_gauss_norm g" + using gauss_norm_smult[of g a] int_gauss_norm_eq val_ord assms + by (metis (no_types, opaque_lifting) Qp.nonzero_closed UPQ.UP_smult_closed UPQ.cfs_zero + eint.distinct(2) eint.inject gauss_norm_coeff_norm local.val_zero plus_eint_simps(1)) + +definition normalize_poly where +"normalize_poly g = (if g = \\<^bsub>UP Q\<^sub>p\<^esub> then g else (\[^](- int_gauss_norm g)) \\<^bsub>Q\<^sub>p_x\<^esub> g)" + +lemma normalize_poly_zero: +"normalize_poly \\<^bsub>UP Q\<^sub>p\<^esub> = \\<^bsub>UP Q\<^sub>p\<^esub>" + unfolding normalize_poly_def by simp + +lemma normalize_poly_nonzero_eq: + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + assumes "g \ carrier (UP Q\<^sub>p)" + shows "normalize_poly g = (\[^](- int_gauss_norm g)) \\<^bsub>UP Q\<^sub>p\<^esub> g" + using assms unfolding normalize_poly_def by simp + +lemma int_gauss_norm_normalize_poly: + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + assumes "g \ carrier (UP Q\<^sub>p)" + shows "int_gauss_norm (normalize_poly g) = 0" + using normalize_poly_nonzero_eq int_gauss_norm_smult assms + by (simp add: ord_p_pow_int p_intpow_closed(2)) + +lemma normalize_poly_closed: + assumes "g \ carrier (UP Q\<^sub>p)" + shows "normalize_poly g \ carrier (UP Q\<^sub>p)" + using assms unfolding normalize_poly_def + by (simp add: p_intpow_closed(1)) + +lemma normalize_poly_nonzero: + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + assumes "g \ carrier (UP Q\<^sub>p)" + shows "normalize_poly g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + using assms normalize_poly_nonzero_eq + by (metis (no_types, lifting) UPQ.UP_smult_one UPQ.module_axioms UPQ.smult_r_null module.smult_assoc1 p_intpow_closed(1) p_intpow_inv') + +lemma gauss_norm_normalize_poly: + assumes "g \ \\<^bsub>UP Q\<^sub>p\<^esub>" + assumes "g \ carrier (UP Q\<^sub>p)" + shows "gauss_norm (normalize_poly g) = 0" +proof- + have 0: "eint (int_gauss_norm (normalize_poly g)) = gauss_norm (normalize_poly g)" + by(rule int_gauss_norm_eq, rule normalize_poly_closed, rule assms, + rule normalize_poly_nonzero, rule assms, rule assms) + show ?thesis + using 0 int_gauss_norm_normalize_poly assms + by (simp add: zero_eint_def) +qed + +lemma taylor_term_eval_eq: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "x \ carrier Q\<^sub>p" + assumes "t \ carrier Q\<^sub>p" + assumes "\j. i \ j \ val (UPQ.taylor_term x f i \ t) < val (UPQ.taylor_term x f j \ t) " + shows "val (f \ t) = val (UPQ.taylor_term x f i \ t)" +proof- + have 0: "f = finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) {..deg Q\<^sub>p f}" + by(rule UPQ.taylor_term_sum[of f "deg Q\<^sub>p f" x], rule assms, blast, rule assms) + show ?thesis + proof(cases "i \ {..deg Q\<^sub>p f}") + case True + have T0: "finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) {..deg Q\<^sub>p f} = UPQ.taylor_term x f i \\<^bsub>UP Q\<^sub>p\<^esub> finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})" + apply(rule UPQ.P.finsum_remove[of "{..deg Q\<^sub>p f}" "UPQ.taylor_term x f" i]) + by(rule UPQ.taylor_term_closed, rule assms, rule assms, blast, rule True) + have T1: "f = UPQ.taylor_term x f i \\<^bsub>UP Q\<^sub>p\<^esub> finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})" + using 0 T0 by metis + have T2: "finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i}) \ carrier (UP Q\<^sub>p)" + apply(rule UPQ.P.finsum_closed) + using UPQ.taylor_term_closed assms(1) assms(2) by blast + have T3: "UPQ.taylor_term x f i \ carrier (UP Q\<^sub>p)" + by(rule UPQ.taylor_term_closed, rule assms, rule assms ) + obtain g where g_def: "g = f" + by blast + have T4: "g = UPQ.taylor_term x f i \\<^bsub>UP Q\<^sub>p\<^esub> finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})" + unfolding g_def by(rule T1) + have g_closed: "g \ carrier (UP Q\<^sub>p)" + unfolding g_def by(rule assms) + have T5: "g \ t = UPQ.taylor_term x f i \ t \ ( finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})) \ t" + unfolding T4 by(rule UPQ.to_fun_plus, rule T2, rule T3, rule assms) + have T6: "( finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})) \ t = + ( finsum Q\<^sub>p (\i. UPQ.taylor_term x f i \ t) ({..deg Q\<^sub>p f} - {i}))" + apply(rule UPQ.to_fun_finsum, blast) + using assms UPQ.taylor_term_closed apply blast + using assms by blast + have T7: "\j. j \ {..deg Q\<^sub>p f} - {i} \ val (UPQ.taylor_term x f j \ t) > val (UPQ.taylor_term x f i \ t)" + using assms by (metis Diff_iff singletonI) + have T8: "val (( finsum (UP Q\<^sub>p) (UPQ.taylor_term x f) ({..deg Q\<^sub>p f} - {i})) \ t) > val (UPQ.taylor_term x f i \ t)" + unfolding T6 + apply(rule finsum_val_ultrametric'') + using UPQ.taylor_term_closed assms + apply (metis (no_types, lifting) Pi_I UPQ.to_fun_closed) + apply blast + using assms T7 apply blast + using assms(4)[of "Suc i"] using eint_ord_simps(4) + assms(4) eint_ord_code(6) g_def gr_implies_not_zero less_one by smt + have T9: "val (g \ t) = val (UPQ.taylor_term x f i \ t)" + unfolding T5 using T8 T2 T3 + by (metis (no_types, lifting) Qp.add.m_comm UPQ.to_fun_closed assms(3) val_ultrametric_noteq) + show ?thesis using T9 unfolding g_def by blast + next + case False + have "i > deg Q\<^sub>p f" + using False by simp + hence "i > deg Q\<^sub>p (UPQ.taylor x f)" + using assms UPQ.taylor_deg by presburger + hence F0: "UPQ.taylor x f i = \" + using assms UPQ.taylor_closed UPQ.deg_leE by blast + have F1: "(UPQ.taylor_term x f i \ t) = \" + using UPQ.to_fun_taylor_term[of f t x i] + unfolding F0 + using assms Qp.cring_simprules(2) Qp.cring_simprules(4) Qp.integral_iff Qp.nat_pow_closed by presburger + show ?thesis + using assms(4)[of "Suc i"] unfolding F1 + by (metis eint_ord_code(6) local.val_zero n_not_Suc_n) + qed +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Hensel's Lemma for \p\-adic fields\ +(**************************************************************************************************) +(**************************************************************************************************) + +theorem hensels_lemma: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "a \ \\<^sub>p" + assumes "gauss_norm f \ 0" + assumes "val (f\a) > 2*val ((pderiv f)\a)" + shows "\!\ \ \\<^sub>p. f\\ = \ \ val (a \ \) > val ((pderiv f)\a)" +proof- + have a_closed: "a \ carrier Q\<^sub>p" + using assms val_ring_memE by auto + have f_nonzero: "f \ \\<^bsub>UP Q\<^sub>p\<^esub>" + proof(rule ccontr) + assume N: "\ f \ \\<^bsub>UP Q\<^sub>p\<^esub>" + then have 0: "pderiv f = \\<^bsub>UP Q\<^sub>p\<^esub>" + using UPQ.deg_zero UPQ.pderiv_deg_0 by blast + have 1: "f = \\<^bsub>UP Q\<^sub>p\<^esub>" + using N by auto + have 2: "eint 2 * val (UPQ.pderiv \\<^bsub>UP Q\<^sub>p\<^esub> \ a) = \" + by (simp add: UPQ.to_fun_zero local.a_closed local.val_zero) + show False using assms a_closed + unfolding 2 1 + using eint_ord_simps(6) by blast + qed + obtain h where h_def: "h = to_Zp_poly f" + by blast + have h_closed: "h \ carrier (UP Z\<^sub>p)" + unfolding h_def using assms + by (simp add: to_Zp_poly_closed) + have h_deriv: "Zp.pderiv h = to_Zp_poly (pderiv f)" + unfolding h_def + using to_Zp_poly_pderiv[of f] assms by auto + have 0: "to_Zp (f\a) = to_function Z\<^sub>p h (to_Zp a)" + unfolding h_def + using assms a_closed + by (simp add: UPQ.to_fun_def to_Zp_poly_eval) + have 1: "to_Zp ((pderiv f)\a) = to_function Z\<^sub>p (Zp.pderiv h) (to_Zp a)" + unfolding h_deriv + using assms a_closed UPQ.pderiv_closed UPQ.to_fun_def eint_ord_trans gauss_norm_pderiv to_Zp_poly_eval + by presburger + have 2: "val (f\a) = val_Zp (to_function Z\<^sub>p h (to_Zp a))" + proof- + have 20: "f\a \ \\<^sub>p" + using assms positive_gauss_norm_eval by blast + have 21: "val (f\a) = val_Zp (to_Zp (f\a))" + using 20 by (simp add: to_Zp_val) + show ?thesis unfolding 21 0 by blast + qed + have 3: "val ((pderiv f)\a) = val_Zp ( to_function Z\<^sub>p (Zp.pderiv h) (to_Zp a))" + proof- + have 30: "(pderiv f)\a \ \\<^sub>p" + using positive_gauss_norm_eval assms gauss_norm_pderiv + by (meson UPQ.pderiv_closed eint_ord_trans) + have 31: "val ((pderiv f)\a) = val_Zp (to_Zp ((pderiv f)\a))" + using 30 by (simp add: to_Zp_val) + show ?thesis unfolding 31 1 by blast + qed + have 4: "\!\. \ \ carrier Z\<^sub>p \ + Zp.to_fun (to_Zp_poly f) \ = \\<^bsub>Z\<^sub>p\<^esub> \ + val_Zp (Zp.to_fun (Zp.pderiv (to_Zp_poly f)) (to_Zp a)) + < val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> \)" + apply(rule hensels_lemma') + using h_closed h_def apply blast + using assms local.a_closed to_Zp_closed apply blast + using assms unfolding 2 3 h_def Zp.to_fun_def by blast + obtain \ where \_def: "\ \ carrier Z\<^sub>p \ + Zp.to_fun (to_Zp_poly f) \ = \\<^bsub>Z\<^sub>p\<^esub> \ + val_Zp (Zp.to_fun (Zp.pderiv (to_Zp_poly f)) (to_Zp a)) + < val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> \) + \ (\x. x \ carrier Z\<^sub>p \ + Zp.to_fun (to_Zp_poly f) x = \\<^bsub>Z\<^sub>p\<^esub> \ + val_Zp (Zp.to_fun (Zp.pderiv (to_Zp_poly f)) (to_Zp a)) + < val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> x) \ x = \)" + using 4 by blast + obtain \ where \_def: "\ = \ \" + by blast + have \_closed: "\ \ \\<^sub>p" + using \_def unfolding \_def by simp + have 5: "(Zp.to_fun (to_Zp_poly f) \) = to_Zp (f\\)" + using \_closed to_Zp_poly_eval[of f \] assms + unfolding \_def UPQ.to_fun_def + by (simp add: Zp.to_fun_def \_def inc_to_Zp) + have 6: "to_Zp (f\\) = \\<^bsub>Z\<^sub>p\<^esub>" + using 5 \_def by auto + have \_closed: "\ \ \\<^sub>p" + unfolding \_def using \_def by simp + have 7: "(f\\) = \" + using 6 assms unfolding \_def + by (metis \_closed \_def inc_of_zero positive_gauss_norm_eval to_Zp_inc) + have 8: "\ = to_Zp \" + unfolding \_def using \_def + by (simp add: inc_to_Zp) + have 9: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> \ = to_Zp (a \ \)" + unfolding 8 using assms(2) \_closed + by (simp add: to_Zp_minus) + have 10: "val (a \ \) = val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> \)" + unfolding 9 using \_closed assms(2) + to_Zp_val val_ring_minus_closed by presburger + have 11: "val (a \ \) > val ((pderiv f)\a)" + using \_def unfolding 9 10 3 h_def + by (simp add: Zp.to_fun_def) + have 12: "\ \ \\<^sub>p \ f \ \ = \ \ val (UPQ.pderiv f \ a) < val (a \ \)" + using "11" "7" \_closed by linarith + have 13: "\x. x\ \\<^sub>p \ f \ x = \ \ val (UPQ.pderiv f \ a) < val (a \ x) + \ x = \" + proof(rule, rule) + fix x assume A: "x \ \\<^sub>p \ f \ x = \ \ val (UPQ.pderiv f \ a) < val (a \ x)" + obtain y where y_def: "y = to_Zp x" + by blast + have y_closed: "y \ carrier Z\<^sub>p" + unfolding y_def using A + by (simp add: to_Zp_closed val_ring_memE(2)) + have eval: "Zp.to_fun (to_Zp_poly f) y = \\<^bsub>Z\<^sub>p\<^esub>" + unfolding y_def using A assms + by (metis UPQ.to_fun_def Zp.to_fun_def to_Zp_poly_eval to_Zp_zero) + have 0: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> y = to_Zp (a \ x)" + unfolding y_def using A assms + by (simp add: to_Zp_minus) + have q: " val_Zp (Zp.to_fun (Zp.pderiv (to_Zp_poly f)) (to_Zp a)) = val (UPQ.pderiv f \ a)" + by (simp add: "3" Zp.to_fun_def h_def) + have 1: "y \ carrier Z\<^sub>p \ + Zp.to_fun (to_Zp_poly f) y = \\<^bsub>Z\<^sub>p\<^esub> \ + val_Zp (Zp.to_fun (Zp.pderiv (to_Zp_poly f)) (to_Zp a)) + < val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> y)" + unfolding 0 eval Zp.to_fun_def h_def + apply(intro conjI y_closed) + using eval Zp.to_fun_def apply (simp; fail) + using A unfolding 0 eval Zp.to_fun_def h_def 3 + using assms(2) to_Zp_val val_ring_minus_closed by presburger + have 2: "y = \" + using 1 \_def by blast + show "x = \" + using y_def unfolding 2 8 using A \_closed + by (metis to_Zp_inc) + qed + show "\!\. \ \ \\<^sub>p \ f \ \ = \ \ val (UPQ.pderiv f \ a) < val (a \ \)" + using 12 13 by metis +qed + +lemma nth_root_poly_root_fixed: + assumes "(n::nat) > 1" + assumes "a \ \\<^sub>p" + assumes "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > 2* val ([n]\\)" + shows "(\! b \ \\<^sub>p. (b[^]n) = a \ val (b \ \) > val ([n]\\))" +proof- + obtain f where f_def: "f = up_ring.monom (UP Q\<^sub>p) \ n \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) a 0" + by blast + have f_closed: "f \ carrier (UP Q\<^sub>p)" + unfolding f_def apply(rule UPQ.P.ring_simprules) + apply (simp; fail) using assms + by (simp add: val_ring_memE(2)) + have 0: "UPQ.pderiv (up_ring.monom (UP Q\<^sub>p) a 0) = \\<^bsub>UP Q\<^sub>p\<^esub>" + using assms + by (simp add: val_ring_memE(2)) + have 1: "UPQ.pderiv (up_ring.monom (UP Q\<^sub>p) (\) n) = (up_ring.monom (UP Q\<^sub>p) ([n]\\) (n-1)) " + using UPQ.pderiv_monom by blast + have 2: "up_ring.monom (UP Q\<^sub>p) \ n \ carrier (UP Q\<^sub>p)" + by simp + have 3: "up_ring.monom (UP Q\<^sub>p) a 0 \ carrier (UP Q\<^sub>p)" + using assms val_ring_memE by simp + have 4: "UPQ.pderiv f = up_ring.monom (UP Q\<^sub>p) ([n] \ \) (n - 1) \\<^bsub>UP Q\<^sub>p\<^esub> \\<^bsub>UP Q\<^sub>p\<^esub>" + using 2 3 assms val_ring_memE UPQ.pderiv_minus[of "up_ring.monom (UP Q\<^sub>p) \ n" "up_ring.monom (UP Q\<^sub>p) a 0"] + unfolding f_def 0 1 by blast + have 5: "UPQ.pderiv f = (up_ring.monom (UP Q\<^sub>p) ([n]\\) (n-1))" + unfolding 4 a_minus_def by simp + have a_closed: "a \ carrier Q\<^sub>p" + using assms val_ring_memE by blast + have 6: "UPQ.pderiv f \ \ = [n]\\ \ \[^](n-1)" + unfolding 5 using a_closed + by (simp add: UPQ.to_fun_monom) + have 7: "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > val \" + proof- + have "eint 2 * val ([n] \ \) \ 0" + by (meson eint_ord_trans eint_pos_int_times_ge val_of_nat_inc zero_less_numeral) + thus ?thesis + using assms unfolding val_one + by (simp add: Q\<^sub>p_def) + qed + hence 8: "val a = val \" + using a_closed + by (metis Qp.cring_simprules(6) ultrametric_equal_eq') + have 9:"val (a [^] (n - 1)) = 0" + by (simp add: "8" local.a_closed val_zero_imp_val_pow_zero) + have 10: "val ([n]\\ \ \[^](n-1)) = val ([n]\\)" + unfolding val_one 9 by simp + have 11: "0 \ gauss_norm f" + proof- + have p0: "gauss_norm (up_ring.monom (UP Q\<^sub>p) \ n) \ 0" + using gauss_norm_monom by simp + have p1: "gauss_norm (up_ring.monom (UP Q\<^sub>p) a 0) \ 0" + using gauss_norm_monom assms val_ring_memE by simp + have p2: "min (gauss_norm (up_ring.monom (UP Q\<^sub>p) \ n)) (gauss_norm (up_ring.monom (UP Q\<^sub>p) a 0)) \ 0" + using p0 p1 by simp + have p3: "0 \ gauss_norm + (up_ring.monom (UP Q\<^sub>p) \ n \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) a 0)" + using gauss_norm_ultrametric'[of "up_ring.monom (UP Q\<^sub>p) \ n" "up_ring.monom (UP Q\<^sub>p) a 0"] + p2 "2" "3" eint_ord_trans by blast + show ?thesis using p3 unfolding f_def by simp + qed + have 12: "\\. \ \ \\<^sub>p \ f \ \ = \[^]n \ a" + unfolding f_def using a_closed + by (simp add: UPQ.to_fun_const UPQ.to_fun_diff UPQ.to_fun_monic_monom val_ring_memE(2)) + have 13: "\!\. \ \ \\<^sub>p \ f \ \ = \ \ val (UPQ.pderiv f \ \) < val (\ \ \)" + apply(rule hensels_lemma, rule f_closed, rule one_in_val_ring, rule 11) + unfolding 6 10 + using a_closed assms 12[of \] assms(3) + by (simp add: one_in_val_ring) + have 14: "\\. \ \ \\<^sub>p \ \[^]n = a \ f \ \ = \" + unfolding f_def using a_closed 12 f_def val_ring_memE(2) by auto + have 15: "val (UPQ.pderiv f \ \) = val ([n]\\)" + unfolding 6 10 by auto + have 16: "\\. \ \ \\<^sub>p \ val (\ \ \) = val (\ \ \)" + proof- + have 17: "\\. \ \ \\<^sub>p \ (\ \ \) = \ (\ \ \)" + using val_ring_memE + by (meson Qp.minus_a_inv Qp.one_closed) + show "\\. \ \ \\<^sub>p \ val (\ \ \) = val (\ \ \)" + unfolding 17 + using Qp.minus_closed Qp.one_closed val_minus val_ring_memE(2) by presburger + qed + show ?thesis using 13 unfolding 15 using 14 16 Qp.one_closed val_ring_memE(2) by metis +qed + +lemma mod_zeroE: + assumes "(a::int) mod k = 0" + shows "\l. a = l*k" + using assms + using Groups.mult_ac(2) by blast + +lemma to_Zp_poly_closed': + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\i. g i \ \\<^sub>p" + shows "to_Zp_poly g \ carrier (UP Z\<^sub>p)" +proof(rule to_Zp_poly_closed) + show "g \ carrier (UP Q\<^sub>p)" + using assms(1) by blast + show "0 \ gauss_norm g" + proof- + have "\i. val (g i) \ 0" + using assms val_ring_memE by blast + thus ?thesis unfolding gauss_norm_def + by (metis gauss_norm_coeff_norm gauss_norm_def) + qed +qed + +lemma to_Zp_poly_eval_to_Zp: + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "\i. g i \ \\<^sub>p" + assumes "a \ \\<^sub>p" + shows "to_function Z\<^sub>p (to_Zp_poly g) (to_Zp a) = to_Zp (g \ a)" +proof- + have "(\i. g i \ \\<^sub>p) \ to_function Z\<^sub>p (to_Zp_poly g) (to_Zp a) = to_Zp (g \ a)" + apply(rule UPQ.poly_induct[of g]) using assms apply blast + proof + fix p assume A: "p \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p p = 0" "\i. p i \ \\<^sub>p" + obtain c where c_def: "c \ carrier Q\<^sub>p \ p = up_ring.monom (UP Q\<^sub>p) c 0" + using A by (metis UPQ.ltrm_deg_0 val_ring_memE(2)) + have 0: "to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c 0) = up_ring.monom (UP Z\<^sub>p) (to_Zp c) 0" + unfolding to_Zp_poly_def proof fix n show " to_Zp (up_ring.monom (UP Q\<^sub>p) c 0 n) = up_ring.monom (UP Z\<^sub>p) (to_Zp c) 0 n" + using UP_ring.cfs_monom[of Z\<^sub>p "to_Zp c" 0 n] UP_ring.cfs_monom[of Q\<^sub>p c 0 n] to_Zp_closed[of c ] + unfolding UP_ring_def + apply(cases "0 = n") + using UPQ.cfs_monom Zp.cfs_monom c_def apply presburger + using UPQ.cfs_monom Zp.cfs_monom c_def + using to_Zp_zero by presburger + qed + have p_eq: "p = up_ring.monom (UP Q\<^sub>p) c 0" + using c_def by blast + have 1: "(up_ring.monom (UP Q\<^sub>p) c 0 \ a) = c" + using UPQ.to_fun_to_poly[of c a] c_def assms val_ring_memE + unfolding to_polynomial_def by blast + show "to_function Z\<^sub>p (to_Zp_poly p) (to_Zp a) = to_Zp (p \ a)" + using c_def assms(3) val_ring_memE(2)[of a] + UP_cring.to_fun_to_poly[of Z\<^sub>p "to_Zp c" "to_Zp a"] + unfolding p_eq 0 1 Zp.to_fun_def to_polynomial_def + using Zp.UP_cring_axioms to_Zp_closed by blast + next + show "\p. (\q. q \ carrier (UP Q\<^sub>p) \ + deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p) \ to_function Z\<^sub>p (to_Zp_poly q) (to_Zp a) = to_Zp (q \ a)) \ + p \ carrier (UP Q\<^sub>p) \ 0 < deg Q\<^sub>p p \ (\i. p i \ \\<^sub>p) \ to_function Z\<^sub>p (to_Zp_poly p) (to_Zp a) = to_Zp (p \ a)" + proof fix p + assume A: "(\q. q \ carrier (UP Q\<^sub>p) \ + deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p) \ to_function Z\<^sub>p (to_Zp_poly q) (to_Zp a) = to_Zp (q \ a))" + "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p" "\i. p i \ \\<^sub>p" + show "to_function Z\<^sub>p (to_Zp_poly p) (to_Zp a) = to_Zp (p \ a)" + proof- + obtain q where q_def: "q = truncate Q\<^sub>p p" + by blast + have q_closed: "q \ carrier (UP Q\<^sub>p)" + unfolding q_def by(rule UPQ.trunc_closed, rule A) + obtain c where c_def: "c = UPQ.lcf p" + by blast + obtain n where n_def: "n = deg Q\<^sub>p p" + by blast + have 0: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) c n" + unfolding c_def n_def q_def + using A(2) UPQ.trunc_simps(1) by blast + have 1: "up_ring.monom (UP Q\<^sub>p) c n \ carrier (UP Q\<^sub>p)" + using A(2) UPQ.ltrm_closed c_def n_def by blast + have 2: "p \ a = q \ a \ (c \ a[^]n)" + unfolding 0 using assms val_ring_memE + by (metis "1" A(4) UPQ.to_fun_monom UPQ.to_fun_plus c_def q_closed) + have 3: "\i. i < n \ q i = p i" + unfolding n_def q_def + using A(2) UPQ.trunc_cfs by blast + have 4: "deg Q\<^sub>p q < n" + unfolding n_def q_def using A + using UPQ.trunc_degree by presburger + have 5: "\i. i \ n \ i > deg Q\<^sub>p q" + using A[of ] less_le_trans[of "deg Q\<^sub>p q" "deg Q\<^sub>p p"] unfolding q_def n_def + using "4" n_def q_def by blast + have 6: "\i. i \ n \ q i = \" + using q_closed 5 UPQ.deg_leE by blast + have 7: "(\i. q i \ \\<^sub>p) \ to_function Z\<^sub>p (to_Zp_poly q) (to_Zp a) = to_Zp (q \ a)" + apply(rule A) unfolding q_def + using q_closed q_def apply blast + using "4" n_def q_def by blast + have 8: "(\i. q i \ \\<^sub>p)" + proof fix i show "q i \ \\<^sub>p" apply(cases "i < n") + using 3 A(4) apply blast using 6[of i] + by (metis less_or_eq_imp_le linorder_neqE_nat zero_in_val_ring) + qed + have 9: "to_function Z\<^sub>p (to_Zp_poly q) (to_Zp a) = to_Zp (q \ a)" + using 7 8 by blast + have 10: "to_Zp_poly p = to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)" + proof fix x + have 100: "to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n) = (up_ring.monom (UP Z\<^sub>p) (to_Zp c) n)" + using to_Zp_poly_monom[of c] A(4) c_def by blast + have 101: "deg Z\<^sub>p (to_Zp_poly q) \ n-1" + apply(rule UP_cring.deg_leqI) + unfolding UP_cring_def using Zp.R_cring apply auto[1] + using to_Zp_poly_closed' 8 q_closed apply blast + unfolding to_Zp_poly_def using 4 6 + by (simp add: to_Zp_zero) + have 102: "(to_Zp_poly q) \ carrier (UP Z\<^sub>p)" + apply(rule to_Zp_poly_closed', rule q_closed) using 8 by blast + have 103: "deg Z\<^sub>p (to_Zp_poly q) < n" + using 101 4 by linarith + have T0: "(to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)) x = + (to_Zp_poly q x) \\<^bsub>Z\<^sub>p\<^esub> (to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n) x)" + apply(rule UP_ring.cfs_add) + unfolding UP_ring_def apply (simp add: Zp.is_ring) + apply(rule 102) unfolding 100 apply(rule UP_ring.monom_closed) + unfolding UP_ring_def apply (simp add: Zp.is_ring) + apply(rule to_Zp_closed ) unfolding c_def + using A(2) UPQ.UP_car_memE(1) by blast + have c_closed: "c \ \\<^sub>p" + unfolding c_def using A(4) by blast + have to_Zp_c_closed: "to_Zp c \ carrier Z\<^sub>p" + using c_closed to_Zp_closed val_ring_memE(2) by blast + show "to_Zp_poly p x = (to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)) x" + proof(cases "x < n") + case True + have T1: "(to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n) x) = \\<^bsub>Z\<^sub>p\<^esub>" + using True UP_ring.cfs_monom[of Z\<^sub>p] unfolding UP_ring_def + by (metis "100" A(2) UPQ.cfs_closed UPQ.deg_leE Zp.is_ring c_def n_def to_Zp_closed to_Zp_zero) + have T2: "to_Zp (p x) = to_Zp (q x)" using 3[of x] True by smt + have T3: "to_Zp (p x) \ carrier Z\<^sub>p" + apply(rule to_Zp_closed) using A(2) UPQ.UP_car_memE(1) by blast + show ?thesis using T3 + unfolding T0 unfolding T1 unfolding to_Zp_poly_def T2 + using Zp.cring_simprules(8) add_comm by presburger + next + case False + have F: "q x = \ " + using False + by (metis "6" less_or_eq_imp_le linorder_neqE_nat) + have F': "(to_Zp_poly q) x = \\<^bsub>Z\<^sub>p\<^esub>" + unfolding to_Zp_poly_def F using to_Zp_zero by blast + show "to_Zp_poly p x = (to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)) x" + proof(cases "x = n") + case True + have T1: "to_Zp (p x) \ carrier Z\<^sub>p" + apply(rule to_Zp_closed) + using A(2) UPQ.UP_car_memE(1) by blast + have T2: "(to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n) x) = to_Zp c" + unfolding 100 using UP_ring.cfs_monom[of Z\<^sub>p "to_Zp c" n n] unfolding UP_ring_def True + using Zp.is_ring to_Zp_c_closed by presburger + show ?thesis using to_Zp_c_closed unfolding T0 F' T2 unfolding to_Zp_poly_def True c_def n_def + using Zp.cring_simprules(8) by presburger + next + case FF: False + have F0: "p x = \" + using FF False unfolding n_def + using A(2) UPQ.UP_car_memE(2) linorder_neqE_nat by blast + have F1: "q x = \" + using FF False F by linarith + have F2: "(up_ring.monom (UP Q\<^sub>p) c n) x = \" + using FF False A(2) UPQ.cfs_closed UPQ.cfs_monom c_def by presburger + show ?thesis unfolding T0 unfolding to_Zp_poly_def F0 F1 F2 + using Zp.r_zero Zp.zero_closed to_Zp_zero by presburger + qed + qed + qed + have 11: "deg Z\<^sub>p (to_Zp_poly q) \ n-1" + apply(rule UP_cring.deg_leqI) + unfolding UP_cring_def using Zp.R_cring apply auto[1] + using to_Zp_poly_closed' 8 q_closed apply blast + unfolding to_Zp_poly_def using 4 6 + by (smt diff_commute diff_diff_cancel less_one less_or_eq_imp_le linorder_neqE_nat to_Zp_zero zero_less_diff) + have 12: "(to_Zp_poly q) \ carrier (UP Z\<^sub>p)" + apply(rule to_Zp_poly_closed', rule q_closed) using 8 by blast + have 13: "deg Z\<^sub>p (to_Zp_poly q) < n" + using 11 4 by linarith + have 14: "to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n) = (up_ring.monom (UP Z\<^sub>p) (to_Zp c) n)" + using to_Zp_poly_monom[of c] A(4) c_def by blast + have 15: "Zp.to_fun (to_Zp_poly q \\<^bsub>UP Z\<^sub>p\<^esub> to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)) (to_Zp a)= + Zp.to_fun (to_Zp_poly q) (to_Zp a) \\<^bsub>Z\<^sub>p\<^esub> Zp.to_fun (to_Zp_poly (up_ring.monom (UP Q\<^sub>p) c n)) (to_Zp a)" + apply(rule Zp.to_fun_plus) + unfolding 14 apply(rule UP_ring.monom_closed) + unfolding UP_ring_def using Zp.is_ring apply auto[1] + apply(rule to_Zp_closed) unfolding c_def + using A(2) UPQ.cfs_closed apply blast + using 12 apply blast + apply(rule to_Zp_closed) using assms val_ring_memE by blast + have 16: "to_Zp (q \ a \ c \ a [^] n) = to_Zp (q \ a) \\<^bsub>Z\<^sub>p\<^esub> to_Zp (c \ a [^] n)" + apply(rule to_Zp_add) + apply(rule val_ring_poly_eval, rule q_closed) + using "8" apply blast + apply(rule assms) + apply(rule val_ring_times_closed) + unfolding c_def using A(4) apply blast + by(rule val_ring_nat_pow_closed, rule assms) + have 17: " to_function Z\<^sub>p (up_ring.monom (UP Z\<^sub>p) (to_Zp c) n) (to_Zp a) = to_Zp (c \ a [^] n)" + proof- + have 170: "to_Zp (c \ a [^] n) = to_Zp c \\<^bsub>Z\<^sub>p\<^esub> to_Zp (a [^] n)" + apply(rule to_Zp_mult[of c "a[^]n"]) + unfolding c_def using A(4) apply blast + by(rule val_ring_nat_pow_closed, rule assms) + have 171: "to_Zp (a [^] n) = (to_Zp a [^]\<^bsub>Z\<^sub>p\<^esub>n)" + by(rule to_Zp_nat_pow, rule assms) + have 172: "to_Zp c \ carrier Z\<^sub>p " + apply(rule to_Zp_closed) unfolding c_def + using A(2) UPQ.UP_car_memE(1) by blast + have 173: "to_Zp a \ carrier Z\<^sub>p " + apply(rule to_Zp_closed) using assms val_ring_memE by blast + show ?thesis + using 172 173 Zp.to_fun_monom[of "to_Zp c" "to_Zp a" n] unfolding Zp.to_fun_def 170 171 + by blast + qed + show ?thesis + using 15 unfolding Zp.to_fun_def 10 2 16 9 unfolding 14 17 + by blast + qed + qed + qed + thus ?thesis using assms by blast +qed + +lemma inc_nat_pow: + assumes "a \ carrier Z\<^sub>p" + shows "\ ([(n::nat)] \\<^bsub>Z\<^sub>p\<^esub>a) = [n]\(\ a)" + apply(induction n) + apply (metis Q\<^sub>p_def Qp.int_inc_zero Qp.nat_mult_zero Zp.add.nat_pow_0 Zp_int_inc_zero' \_def frac_inc_of_int) + unfolding Qp.add.nat_pow_Suc Zp.add.nat_pow_Suc + using Zp_nat_mult_closed assms inc_of_sum by presburger + +lemma poly_inc_pderiv: + assumes "g \ carrier (UP Z\<^sub>p)" + shows "poly_inc (Zp.pderiv g) = UPQ.pderiv (poly_inc g)" +proof fix x + have 0: "UPQ.pderiv (poly_inc g) x = [Suc x] \ poly_inc g (Suc x)" + apply(rule UPQ.pderiv_cfs[of "poly_inc g" x]) + by(rule poly_inc_closed, rule assms) + have 1: "Zp.pderiv g x = [Suc x] \\<^bsub>Z\<^sub>p\<^esub> g (Suc x)" + by(rule Zp.pderiv_cfs[of g x], rule assms) + show "poly_inc (Zp.pderiv g) x = UPQ.pderiv (poly_inc g) x" + unfolding 0 unfolding poly_inc_def 1 apply(rule inc_nat_pow) + using Zp.UP_car_memE(1) assms by blast +qed + +lemma Zp_hensels_lemma: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Z\<^sub>p" + assumes "Zp.to_fun (Zp.pderiv f) a \ \\<^bsub>Z\<^sub>p\<^esub>" + assumes "Zp.to_fun f a \ \\<^bsub>Z\<^sub>p \<^esub>" + assumes "val_Zp (Zp.to_fun f a) > eint 2 * val_Zp (Zp.to_fun (Zp.pderiv f) a)" + obtains \ where + "Zp.to_fun f \ = \\<^bsub>Z\<^sub>p\<^esub>" and "\ \ carrier Z\<^sub>p" + "val_Zp (a \\<^bsub>Z\<^sub>p\<^esub> \) > val_Zp (Zp.to_fun (Zp.pderiv f) a)" + "val_Zp (a \\<^bsub>Z\<^sub>p\<^esub> \) = val_Zp (divide (Zp.to_fun f a) (Zp.to_fun (Zp.pderiv f) a))" + "val_Zp (Zp.to_fun (Zp.pderiv f) \) = val_Zp (Zp.to_fun (Zp.pderiv f) a)" +proof- + have "hensel p f a" + using assms + by (simp add: Zp_def hensel.intro hensel_axioms.intro padic_integers_axioms) + then show ?thesis + using hensel.full_hensels_lemma[of p f a] that + unfolding Zp_def + by blast +qed + +end +end diff --git a/thys/Padic_Field/Padic_Field_Powers.thy b/thys/Padic_Field/Padic_Field_Powers.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Padic_Field_Powers.thy @@ -0,0 +1,11440 @@ +theory Padic_Field_Powers + imports Ring_Powers Padic_Field_Polynomials Generated_Boolean_Algebra + Padic_Field_Topology + + +begin + +text\This theory is intended to develop the necessary background on subsets of powers of a $p$-adic +field to prove Macintyre's quantifier elimination theorem. In particular, we define semi-algebraic +subsets of $\mathbb{Q}_p^n$, semi-algebraic functions $\mathbb{Q}_p^n \to \mathbb{Q}_p$, and semi- +algebraic mappings $\mathbb{Q}_p^n \to \mathbb{Q}_p^m$ for arbitrary $n, m \in \mathbb{N}$. In +addition we prove that many common sets and functions are semi-algebraic. We are closely following +the paper \cite{denef1986} by Denef, where an algebraic proof of Mactinyre's theorem is developed.\ + +(**************************************************************************************************) +(**************************************************************************************************) +section\Cartesian Powers of $p$-adic Fields\ +(**************************************************************************************************) +(**************************************************************************************************) +lemma list_tl: +"tl (t#x) = x" + using List.list.sel(3) by auto + +lemma list_hd: +"hd (t#x) = t" + unfolding List.list.sel(1) by auto + + +sublocale padic_fields < cring_coord_rings Q\<^sub>p "UP Q\<^sub>p" + unfolding cring_coord_rings_axioms_def cring_coord_rings_def + using Qp.zero_not_one UPQ.R_cring + apply (simp add: UPQ.is_UP_cring) + by auto + +sublocale padic_fields < Qp: domain_coord_rings Q\<^sub>p "UP Q\<^sub>p" + unfolding domain_coord_rings_def cring_coord_rings_axioms_def cring_coord_rings_def + using Qp.domain_axioms Qp.zero_not_one UPQ.R_cring + apply (simp add: UPQ.UP_cring_axioms) + by auto + +context padic_fields +begin +no_notation Zp.to_fun (infixl\\\ 70) +no_notation ideal_prod (infixl "\\" 80) + +notation +evimage (infixr "\\" 90) and +euminus_set ("_ \<^sup>c\" 70) + + +type_synonym padic_tuple = "padic_number list" +type_synonym padic_function = "padic_number \ padic_number" +type_synonym padic_nary_function = "padic_tuple \ padic_number" +type_synonym padic_function_tuple = "padic_nary_function list" +type_synonym padic_nary_function_poly = "nat \ padic_nary_function" + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Polynomials over $\mathbb{Q}_p$ and Polynomial Maps\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma last_closed': + assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "t \ carrier Q\<^sub>p" + using assms last_closed[of n "x@[t]" Q\<^sub>p] + by (metis (full_types) cartesian_power_car_memE gr0I last_snoc + length_append_singleton less_not_refl zero_less_Suc) + +lemma segment_in_car': + assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" +proof- + have 0: "length x = n" + by (metis Suc_inject assms cartesian_power_car_memE length_append_singleton) + have "set x \ set (x@[t])" + by (metis rotate1.simps(2) set_rotate1 set_subset_Cons) + then have 1: "set x \ carrier Q\<^sub>p" + using assms cartesian_power_car_memE''[of "x@[t]" Q\<^sub>p "Suc n"] + by blast + show ?thesis + using 0 1 assms cartesian_power_car_memI[of x n Q\<^sub>p] + by blast +qed + +lemma Qp_zero: +"Q\<^sub>p\<^bsup>0\<^esup> = nil_ring" + unfolding cartesian_power_def + by simp + +lemma Qp_zero_carrier: +"carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" + by (simp add: Qp_zero) + +text\Abbreviation for constant polynomials\ + +abbreviation(input) Qp_to_IP where +"Qp_to_IP k \ Qp.indexed_const k" + +lemma Qp_to_IP_car: + assumes "k \ carrier Q\<^sub>p" + shows "Qp_to_IP k \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms + unfolding coord_ring_def + using Qp.indexed_const_closed by blast + +lemma(in cring_coord_rings) smult_closed: + assumes "a \ carrier R" + assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q \ carrier (R[\\<^bsub>n\<^esub>])" + using assms unfolding coord_ring_def + using Pring_smult_closed + by (simp add: R.Pring_smult_closed) + + +lemma Qp_poly_smult_cfs: + assumes "a \ carrier Q\<^sub>p" + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "(a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) m = a \ (P m)" + using assms unfolding coord_ring_def + using Qp.Pring_smult_cfs by blast + +lemma Qp_smult_r_distr: + assumes "a \ carrier Q\<^sub>p" + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q) = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q)" + using assms unfolding coord_ring_def + using Qp.Pring_smult_r_distr by blast + +lemma Qp_smult_l_distr: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "(a \ b) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (b \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P)" + using assms unfolding coord_ring_def + using Qp.Pring_smult_l_distr by blast + +abbreviation(input) Qp_funs where +"Qp_funs n \ Fun\<^bsub>n\<^esub> Q\<^sub>p" + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Evaluation of Polynomials in $\mathbb{Q}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + + +abbreviation(input) Qp_ev where +"Qp_ev P q \ (eval_at_point Q\<^sub>p q P)" + +lemma Qp_ev_one: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \" unfolding coord_ring_def + by (metis Qp.Pring_one eval_at_point_const Qp.one_closed assms) + +lemma Qp_ev_zero: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \"unfolding coord_ring_def + by (metis Qp.Pring_zero eval_at_point_const Qp.zero_closed assms) + +lemma Qp_eval_pvar_pow: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "k < n" + assumes "(m::nat) \ 0" + shows "Qp_ev ((pvar Q\<^sub>p k)[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> m) a = ((a!k)[^]m)" + by (metis eval_at_point_nat_pow eval_pvar assms(1) assms(2) pvar_closed) + +text\composition of polynomials over $\mathbb{Q}_p$\ + +definition Qp_poly_comp where +"Qp_poly_comp m fs = poly_compose (length fs) m fs" + +text\lemmas about polynomial maps\ + +lemma Qp_is_poly_tupleI: + assumes "\i. i < length fs\ fs!i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + shows "is_poly_tuple m fs" + unfolding is_poly_tuple_def using assms + using cartesian_power_car_memE'' cartesian_power_car_memI' by blast + +lemma Qp_is_poly_tuple_append: + assumes "is_poly_tuple m fs" + assumes "is_poly_tuple m gs" + shows "is_poly_tuple m (fs@gs)" +proof(rule Qp_is_poly_tupleI) + show "\i. i < length (fs @ gs) \ (fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + proof- fix i assume A: "i < length (fs @ gs)" + show "(fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + apply(cases "i < length fs") + using assms is_poly_tupleE[of m fs i] nth_append[of fs gs i] + apply presburger + proof- + assume B: "\ i < length fs" + then have C: "length fs \ i \ i < length (fs @ gs)" + using A not_le + by blast + then have "i - length fs < length gs" + using length_append[of fs gs] + by linarith + then show ?thesis + using A assms is_poly_tupleE[of m gs "i - length fs"] nth_append[of fs gs i] B + by presburger + qed + qed +qed + +lemma Qp_poly_mapE: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "j < m" + shows "(poly_map n fs as)!j \ carrier Q\<^sub>p" + using assms poly_map_closed cartesian_power_car_memE' by blast + +lemma Qp_poly_mapE': + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "length (poly_map n fs as) = length fs" + unfolding poly_map_def + using Qp.cring_axioms poly_tuple_evalE' + by (metis assms restrict_def) + +lemma Qp_poly_mapE'': + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "n \ 0" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "j < m" + shows "(poly_map n fs as)!j = (Qp_ev (fs!j) as)" + using assms + unfolding poly_map_def poly_tuple_eval_def + by (metis (no_types, lifting) nth_map restrict_apply') + +lemma poly_map_apply: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "poly_map n fs as = poly_tuple_eval fs as" + unfolding poly_map_def restrict_def + by (simp add: assms) + +lemma poly_map_pullbackI: + assumes "is_poly_tuple n fs" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "poly_map n fs as \ S" + shows "as \ poly_map n fs \\<^bsub>n\<^esub> S" + using assms poly_map_apply + by blast + +lemma poly_map_pullbackI': + assumes "is_poly_tuple n fs" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "poly_map n fs as \ S" + shows "as \ ((poly_map n fs) -` S)" + by (simp add: assms(3)) + +text\lemmas about polynomial composition\ +lemma poly_compose_ring_hom: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + shows "(ring_hom_ring (Q\<^sub>p[\\<^bsub>n\<^esub>]) (Q\<^sub>p[\\<^bsub>m\<^esub>]) (Qp_poly_comp m fs))" + unfolding Qp_poly_comp_def + by (simp add: assms(1) assms(2) poly_compose_ring_hom) + +lemma poly_compose_closed: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "(Qp_poly_comp m fs f) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + using Qp.cring_axioms assms + unfolding Qp_poly_comp_def + using poly_compose_closed by blast + + +lemma poly_compose_add: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" + using Qp.cring_axioms assms poly_compose_add + unfolding is_poly_tuple_def Qp_poly_comp_def + by blast + +lemma poly_compose_mult: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" + using Qp.cring_axioms assms poly_compose_mult + unfolding is_poly_tuple_def Qp_poly_comp_def + by blast + +lemma poly_compose_const: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "a \ carrier Q\<^sub>p" + shows "Qp_poly_comp m fs (Qp_to_IP a) = Qp_to_IP a" + using Qp.cring_axioms assms poly_compose_const + unfolding is_poly_tuple_def Qp_poly_comp_def + by metis + +lemma Qp_poly_comp_eval: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "Qp_ev (Qp_poly_comp m fs f) as = Qp_ev f (poly_map m fs as)" +proof- + have "(restrict (poly_tuple_eval fs) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) as) = poly_tuple_eval fs as" + unfolding restrict_def + by (simp add: assms) + thus ?thesis + using assms Qp.cring_axioms poly_compose_eval + unfolding Qp_poly_comp_def poly_map_def + by presburger +qed + + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Mapping Univariate Polynomials to Multivariable Polynomials in One Variable\ +(**************************************************************************************************) +(**************************************************************************************************) + + +abbreviation(input) to_Qp_x where +"to_Qp_x \ (IP_to_UP (0::nat) :: (nat multiset \ padic_number) \ nat \ padic_number)" + +abbreviation(input) from_Qp_x where +"from_Qp_x \ UP_to_IP Q\<^sub>p (0::nat)" + +lemma from_Qp_x_closed: + assumes "q \ carrier Q\<^sub>p_x" + shows "from_Qp_x q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" + using assms UP_to_IP_closed unfolding coord_ring_def + by (metis One_nat_def lessThan_0 lessThan_Suc) + +lemma to_Qp_x_closed: + assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" + shows "to_Qp_x q \ carrier Q\<^sub>p_x" + using assms Qp.IP_to_UP_closed[of q "0::nat"] Qp.cring_axioms + unfolding coord_ring_def + by (metis One_nat_def lessThan_0 lessThan_Suc) + +lemma to_Qp_x_from_Qp_x: + assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" + shows "from_Qp_x (to_Qp_x q) = q" + using assms UP_to_IP_inv[of q "0::nat"] Qp.Pring_car + unfolding coord_ring_def + by (metis One_nat_def lessThan_0 lessThan_Suc) + +lemma from_Qp_x_to_Qp_x: + assumes "q \ carrier Q\<^sub>p_x" + shows "to_Qp_x (from_Qp_x q) = q" + by (meson UPQ.IP_to_UP_inv assms) + +text\ring hom properties of these maps\ + +lemma to_Qp_x_ring_hom: +"to_Qp_x \ ring_hom (Q\<^sub>p[\\<^bsub>1\<^esub>]) Q\<^sub>p_x" + using IP_to_UP_ring_hom[of "0::nat"] ring_hom_ring.homh + unfolding coord_ring_def + by (metis One_nat_def lessThan_0 lessThan_Suc) + +lemma from_Qp_x_ring_hom: +"from_Qp_x \ ring_hom Q\<^sub>p_x (Q\<^sub>p[\\<^bsub>1\<^esub>])" + using UP_to_IP_ring_hom ring_hom_ring.homh + unfolding coord_ring_def + by (metis One_nat_def lessThan_0 lessThan_Suc) + + +lemma from_Qp_x_add: + assumes "a \ carrier Q\<^sub>p_x" + assumes "b \ carrier Q\<^sub>p_x" + shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" + by (metis (mono_tags, lifting) assms(1) assms(2) from_Qp_x_ring_hom ring_hom_add) + +lemma from_Qp_x_mult: + assumes "a \ carrier Q\<^sub>p_x" + assumes "b \ carrier Q\<^sub>p_x" + shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" + by (metis assms(1) assms(2) from_Qp_x_ring_hom ring_hom_mult) + +text\equivalence of evaluation maps\ + +lemma Qp_poly_Qp_x_eval: + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" + assumes "a \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + shows "Qp_ev P a = (to_Qp_x P)\(Qp.to_R a)" +proof- + have 0: "(IP_to_UP 0 P) \ (a ! 0) = ((IP_to_UP 0 P) \ (if 0 < length a then a ! 0 else \))" + using assms + by (metis (mono_tags, lifting) cartesian_power_car_memE gr0I zero_neq_one) + have 1: "closed_fun Q\<^sub>p (\n. if n < length a then a ! n else \)" + proof + fix n + show "(if n < length a then a ! n else \) \ carrier Q\<^sub>p" + apply(cases "n < length a") + using assms + apply (metis cartesian_power_car_memE cartesian_power_car_memE') + by (meson Qp.cring_axioms cring.cring_simprules(2)) + qed + have 2: " P \ Pring_set Q\<^sub>p {0::nat}" + using assms unfolding coord_ring_def + by (metis Qp.Pring_car UPQ.UP_to_IP_closed assms(1) to_Qp_x_closed to_Qp_x_from_Qp_x) + have 3: "total_eval Q\<^sub>p (\i. if i < length a then a ! i else \) P = IP_to_UP 0 P \ (if 0 < length a then a ! 0 else \)" + using 1 2 assms IP_to_UP_poly_eval[of P "0::nat" "(\i. if i < length a then a ! i else \)" ] + UPQ.to_fun_def by presburger + then show ?thesis + using 0 + unfolding eval_at_point_def + by blast +qed + +lemma Qp_x_Qp_poly_eval: + assumes "P \ carrier Q\<^sub>p_x" + assumes "a \ carrier Q\<^sub>p" + shows "P \ a = Qp_ev (from_Qp_x P) (to_R1 a)" +proof- + have "Qp_ev (from_Qp_x P) (to_R1 a) = (to_Qp_x (from_Qp_x P))\(Qp.to_R (Qp.to_R1 a))" + using Qp_poly_Qp_x_eval assms(1) assms(2) from_Qp_x_closed Qp.to_R1_closed by blast + then show ?thesis using assms + by (metis UPQ.IP_to_UP_inv Qp.to_R_to_R1) +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\$n^{th}$-Power Sets over $\mathbb{Q}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition P_set where +"P_set (n::nat) = {a \ nonzero Q\<^sub>p. (\y \ carrier Q\<^sub>p . (y[^] n) = a)}" + +lemma P_set_carrier: +"P_set n \ carrier Q\<^sub>p" + unfolding P_set_def nonzero_def + by blast + +lemma P_set_memI: + assumes "a \ carrier Q\<^sub>p" + assumes "a \ \" + assumes "b \ carrier Q\<^sub>p" + assumes "b[^](n::nat) = a" + shows "a \ P_set n" + unfolding P_set_def + using assms + by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) + +lemma P_set_nonzero: +"P_set n \ nonzero Q\<^sub>p" + unfolding P_set_def by blast + + +lemma P_set_nonzero': + assumes "a \ P_set n" + shows "a \ nonzero Q\<^sub>p" + "a \ carrier Q\<^sub>p" + using assms P_set_nonzero P_set_carrier + apply blast using assms P_set_carrier by blast + +lemma P_set_one: + assumes "n \ 0" + shows "\ \ P_set (n::nat)" +proof- + have 0: "\[^]n = \" + using Qp.nat_pow_one by blast + have 1: "\ \ carrier Q\<^sub>p" + by blast + then show ?thesis + using one_nonzero unfolding P_set_def + using 0 by blast +qed + +lemma zeroth_P_set: +"P_set 0 = {\}" +proof + show "P_set 0 \ {\}" + unfolding P_set_def + proof + fix x + assume "x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^](0::nat)) = a}" + then have "\y\carrier Q\<^sub>p. (y[^](0::nat)) = x" + by blast + then obtain a where a_def: "a \ carrier Q\<^sub>p \ (a[^](0::nat)) = x" + by blast + then show "x \ {\}" + using Qp.nat_pow_0 by blast + qed + show "{\} \ P_set 0" + using P_set_memI[of \ \ 0] Qp.nat_pow_one Qp.one_closed local.one_neq_zero + by blast +qed + +lemma P_set_mult_closed: + assumes "n \ 0" + assumes "a \ P_set n" + assumes "b \ P_set n" + shows "a \ b \ P_set n" +proof- + obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ (a0 [^] n = a)" + using assms(2) + unfolding P_set_def + by blast + obtain b0 where b0_def: "b0 \ carrier Q\<^sub>p \ (b0 [^] n = b)" + using assms(3) + unfolding P_set_def + by blast + have "(a0 \ b0) [^] n = a0 [^] n \ b0 [^] n" + using a0_def b0_def Qp.nat_pow_distrib by blast + then have 0: "a \ b = (a0 \ b0) [^] n" + using a0_def b0_def by blast + have 1: "a0 \ b0 \ carrier Q\<^sub>p" + by (meson Qp.cring_axioms a0_def b0_def cring.cring_simprules(5)) + have 2: "a \ b \ nonzero Q\<^sub>p" + using assms nonzero_is_submonoid unfolding P_set_def + by (metis (no_types, lifting) "0" "1" Qp.integral Qp_nat_pow_nonzero a0_def b0_def mem_Collect_eq not_nonzero_Qp) + then show ?thesis + using 0 1 assms + unfolding P_set_def by blast +qed + +lemma P_set_inv_closed: + assumes "a \ P_set n" + shows "inv a \ P_set n" +proof(cases "n = 0") + case True + then show ?thesis + using assms zeroth_P_set + by (metis Qp.inv_one singletonD) +next + case False + then show ?thesis proof- + obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ a0[^]n = a" + using assms P_set_def[of n] by blast + have "a0 \ nonzero Q\<^sub>p" + apply(rule ccontr) + proof- + assume "a0 \ nonzero Q\<^sub>p " + then have "a0 = \" + using a0_def + by (meson not_nonzero_Qp) + then show False using a0_def assms + by (metis (mono_tags, lifting) False P_set_def Qp.cring_axioms \a0 \ nonzero Q\<^sub>p\ + cring_def mem_Collect_eq neq0_conv ring.pow_zero) + qed + then have "(inv a0)[^]n = inv a" + using a0_def \a0 \ carrier Q\<^sub>p \ (a0[^]n) = a\ \a0 \ nonzero Q\<^sub>p\ Units_nonzero + monoid.nat_pow_of_inv[of Q\<^sub>p a n] Qp.nat_pow_of_inv Units_eq_nonzero by presburger + then show ?thesis + by (metis P_set_memI Qp.nat_pow_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero \a0 \ nonzero Q\<^sub>p\ a0_def inv_in_frac(1) inv_in_frac(2)) + qed +qed + +lemma P_set_val: + assumes "a \ P_set (n::nat)" + shows "(ord a) mod n = 0" +proof(cases "n = 0") + case True + then show ?thesis + using assms zeroth_P_set + by (metis mod_by_0 of_nat_0 ord_one singletonD) +next + case False + then show ?thesis + proof- + obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^] n) = a" + using assms P_set_def by blast + have an: "a \ nonzero Q\<^sub>p" + using P_set_def assms by blast + have bn: "b \ nonzero Q\<^sub>p" + proof(rule ccontr) + assume "b \ nonzero Q\<^sub>p" + then have "b = \\<^bsub> Q\<^sub>p\<^esub>" + using b_def not_nonzero_Qp + by metis + then have "(b[^] n) = \" + using False Qp.cring_axioms cring_def ring.pow_zero + by blast + then show False + using b_def an Qp.not_nonzero_memI by blast + qed + then have "ord a = n * (ord b)" + using b_def an nonzero_nat_pow_ord + by blast + then show ?thesis + by (metis mod_mult_self1_is_0) + qed +qed + + +lemma P_set_pow: + assumes "n > 0" + assumes "s \ P_set n" + shows "s[^]k \ P_set (n*k)" +proof- + obtain y where y_def: "y \ carrier Q\<^sub>p \ y[^]n = s" + using assms unfolding P_set_def by blast + then have 0: "y \ nonzero Q\<^sub>p" + using assms P_set_nonzero'(1) Qp_nonzero_nat_pow by blast + have 1: "y[^](n*k) = s[^] k" + using 0 y_def assms Qp.nat_pow_pow by blast + hence 2: "s[^]k \ nonzero Q\<^sub>p" + using 0 by (metis Qp_nat_pow_nonzero) + thus ?thesis unfolding P_set_def using 1 y_def by blast +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Semialgebraic Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + In this section we introduce the notion of a $p$-adic semialgebraic set. Intuitively, these are + the subsets of $\mathbb{Q}_p^n$ which are definable by first order quantifier-free formulas in + the standard first-order language of rings, with an additional relation symbol included for the + relation $\text{ val}(x) \leq \text{ val}(y)$, interpreted according to the definiton of the + $p$-adic valuation on $\mathbb{Q}_p$. In fact, by Macintyre's quantifier elimination theorem + for the first-order theory of $\mathbb{Q}_p$ in this language, one can equivalently remove the + ``quantifier-free" clause from the latter definition. The definition we give here is also + equivalent, and due to Denef in \cite{denef1986}. The given definition here is desirable mainly + for its utility in producing a proof of Macintyre's theorem, which is our overarching goal. +\ + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Defining Semialgebraic Sets\ + (********************************************************************************************) + (********************************************************************************************) + +definition basic_semialg_set where +"basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). \y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)}" + +lemma basic_semialg_set_zero_set: + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "Qp_ev P q = \" + assumes "n \ 0" + shows "q \ basic_semialg_set (m::nat) (n::nat) P" +proof- + have "\ = (\[^]n)" + using assms(4) Qp.nat_pow_zero by blast + then show ?thesis + unfolding basic_semialg_set_def + using assms Qp.cring_axioms cring.cring_simprules(2) + by blast +qed + +lemma basic_semialg_set_def': + assumes "n \ 0" + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + shows "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ (P_set n)}" +proof + show "basic_semialg_set m n P \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" + proof + fix x + assume A: "x \ basic_semialg_set m n P" + show "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" + apply(cases "Qp_ev P x = \") + using A basic_semialg_set_def apply blast + unfolding basic_semialg_set_def P_set_def + proof + assume A0: "Qp_ev P x \ \" + have A1: " \y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n)" + using A basic_semialg_set_def + by blast + have A2: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A basic_semialg_set_def + by blast + show " x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (Qp_ev P x = \ \ Qp_ev P x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a})" + by (metis (mono_tags, lifting) A1 A2 Qp.nonzero_memI assms(2) eval_at_point_closed mem_Collect_eq) + qed + qed + show "{q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} \ basic_semialg_set m n P" + proof + fix x + assume A: " x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" + then have A':"x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + by blast + show "x \ basic_semialg_set m n P" + using A A' + apply(cases "Qp_ev P x = \") + using assms basic_semialg_set_zero_set[of P m x n] + apply blast + proof- + assume B: "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} " + assume B': "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assume B'': "Qp_ev P x \ \ " + show "x \ basic_semialg_set m n P" + unfolding basic_semialg_set_def P_set_def + proof + have "\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n) " + using A nonzero_def [of Q\<^sub>p] unfolding P_set_def + proof - + assume "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a}}" + then have "Qp_ev P x \ nonzero Q\<^sub>p \ (\r. r \ carrier Q\<^sub>p \ (r[^]n) = Qp_ev P x)" + using B'' by blast + then show ?thesis + by blast + qed + then show "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n))" + using B' + by blast + qed + qed + qed +qed + +lemma basic_semialg_set_memI: + assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "y \ carrier Q\<^sub>p" + assumes "Qp_ev P q = (y[^]n)" + shows "q \ basic_semialg_set m n P" + using assms(1) assms(2) assms(3) basic_semialg_set_def + by blast + +lemma basic_semialg_set_memE: + assumes "q \ basic_semialg_set m n P" + shows "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + "\y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)" + using assms basic_semialg_set_def apply blast + using assms basic_semialg_set_def by blast + +definition is_basic_semialg :: "nat \ ((nat \ int) \ (nat \ int)) set list set \ bool" where +"is_basic_semialg m S \ (\ (n::nat) \ 0. (\ P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]). S = basic_semialg_set m n P))" + +abbreviation(input) basic_semialgs where +"basic_semialgs m \ {S. (is_basic_semialg m S)}" + +definition semialg_sets where +"semialg_sets n = gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" + +lemma carrier_is_semialg: +"(carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ semialg_sets n " + unfolding semialg_sets_def + using gen_boolean_algebra.universe by blast + +lemma empty_set_is_semialg: +" {} \ semialg_sets n" + using carrier_is_semialg[of n] + unfolding semialg_sets_def using gen_boolean_algebra.complement + by (metis Diff_cancel) + +lemma semialg_intersect: + assumes "A \ semialg_sets n" + assumes "B \ semialg_sets n" + shows "(A \ B) \ semialg_sets n " + using assms(1) assms(2) gen_boolean_algebra_intersect semialg_sets_def + by blast + +lemma semialg_union: + assumes "A \ semialg_sets n" + assumes "B \ semialg_sets n" + shows "(A \ B) \ semialg_sets n " + using assms gen_boolean_algebra.union semialg_sets_def + by blast + +lemma semialg_complement: + assumes "A \ semialg_sets n" + shows "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) \ semialg_sets n " + using assms gen_boolean_algebra.complement semialg_sets_def + by blast + +lemma semialg_zero: + assumes "A \ semialg_sets 0" + shows "A = {[]} \ A = {}" + using assms + unfolding semialg_sets_def cartesian_power_def +proof- + assume A0: " A \ gen_boolean_algebra (carrier (RDirProd_list (R_list 0 Q\<^sub>p))) (basic_semialgs 0)" + show " A = {[]} \ A = {}" + proof- + have "A \ {[]} \ A = {}" + proof + assume A1: "A \ {[]}" + show "A = {}" + proof- + have "(R_list 0 Q\<^sub>p) = []" + by simp + then have "(carrier (RDirProd_list (R_list 0 Q\<^sub>p))) = {[]}" + using RDirProd_list_nil + by simp + then show ?thesis + using A0 A1 + by (metis gen_boolean_algebra_subset subset_singletonD) + qed + qed + then show ?thesis + by linarith + qed +qed + +lemma basic_semialg_is_semialg: + assumes "is_basic_semialg n A" + shows "A \ semialg_sets n" + by (metis (no_types, lifting) assms gen_boolean_algebra.simps inf_absorb1 + is_basic_semialg_def mem_Collect_eq basic_semialg_set_def + semialg_sets_def subsetI) + +lemma basic_semialg_is_semialg': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "m \0" + assumes "A = basic_semialg_set n m f" + shows "A \ semialg_sets n" + using assms basic_semialg_is_semialg is_basic_semialg_def + by blast + +definition is_semialgebraic where +"is_semialgebraic n S = (S \ semialg_sets n)" + +lemma is_semialgebraicE: + assumes "is_semialgebraic n S" + shows "S \ semialg_sets n" + using assms is_semialgebraic_def by blast + +lemma is_semialgebraic_closed: + assumes "is_semialgebraic n S" + shows "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using is_semialgebraicE[of n S] unfolding semialg_sets_def + using assms gen_boolean_algebra_subset is_semialgebraicE semialg_sets_def + by blast + +lemma is_semialgebraicI: + assumes "S \ semialg_sets n" + shows "is_semialgebraic n S" + by (simp add: assms is_semialgebraic_def) + +lemma basic_semialg_is_semialgebraic: + assumes "is_basic_semialg n A" + shows "is_semialgebraic n A" + using assms basic_semialg_is_semialg is_semialgebraicI by blast + +lemma basic_semialg_is_semialgebraic': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "m \0" + assumes "A = basic_semialg_set n m f" + shows "is_semialgebraic n A" + using assms(1) assms(2) assms(3) basic_semialg_is_semialg' is_semialgebraicI by blast + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Algebraic Sets over $p$-adic Fields\ + (********************************************************************************************) + (********************************************************************************************) + +lemma p_times_square_not_square: + assumes "a \ nonzero Q\<^sub>p" + shows "\ \ (a [^] (2::nat)) \ P_set (2::nat)" +proof + assume A: "\ \ (a[^](2::nat)) \ P_set (2::nat)" + then have "\ \ (a[^](2::nat)) \ nonzero Q\<^sub>p" + unfolding P_set_def + by blast + then obtain b where b_def: "b \ carrier Q\<^sub>p \ b[^](2::nat) = \ \ (a[^](2::nat))" + using A P_set_def by blast + have "b \ nonzero Q\<^sub>p" + apply(rule ccontr) using b_def assms + by (metis A P_set_nonzero'(1) Qp.nat_pow_zero not_nonzero_Qp zero_neq_numeral) + then have LHS: "ord (b[^](2::nat)) = 2* (ord b)" + using nonzero_nat_pow_ord + by presburger + have "ord( \ \ (a[^](2::nat))) = 1 + 2* ord a" + using assms nonzero_nat_pow_ord Qp_nat_pow_nonzero ord_mult ord_p p_nonzero + by presburger + then show False + using b_def LHS + by presburger +qed + +lemma p_times_square_not_square': + assumes "a \ carrier Q\<^sub>p" + shows "\ \ (a [^] (2::nat)) = \ \ a = \" + by (metis Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero assms p_nonzero) + +lemma zero_set_semialg_set: + assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_ev q a = \ \( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev q a) [^] (2::nat)) = y[^](2::nat)) " +proof + show "Qp_ev q a = \ \ \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" + proof- + assume "Qp_ev q a = \" + then have "\ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" + by (metis Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null zero_neq_numeral) + then have "\ \ carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" + using Qp.cring_axioms cring.cring_simprules(2) + by blast + then show "\y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" + by blast + qed + show " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat)) \ Qp_ev q a = \" + proof- + assume A: " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat))" + then obtain b where b_def: "b\carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (b[^](2::nat))" + by blast + show "Qp_ev q a = \" + proof(rule ccontr) + assume " Qp_ev q a \ \" + then have " Qp_ev q a \ nonzero Q\<^sub>p" using assms eval_at_point_closed[of a n q] nonzero_def + proof - + have "Qp_ev q a \ carrier Q\<^sub>p" + using \\a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>); q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ \ + Qp_ev q a \ carrier Q\<^sub>p\ \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ + by fastforce + then have "Qp_ev q a \ {r \ carrier Q\<^sub>p. r \ \}" + using \Qp_ev q a \ \\ by force + then show ?thesis + by (metis nonzero_def ) + qed + then have "\ \ (Qp_ev q a[^](2::nat)) \ nonzero Q\<^sub>p" + by (metis Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero not_nonzero_Qp p_nonzero p_times_square_not_square') + then have "\ \ (Qp_ev q a[^](2::nat)) \ P_set (2::nat)" + using b_def + unfolding P_set_def + by blast + then show False + using \Qp_ev q a \ nonzero Q\<^sub>p\ p_times_square_not_square + by blast + qed + qed +qed + +lemma alg_as_semialg: + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" + shows "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" +proof + have 00: "\x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ Qp_ev q x = \ \ (Qp_ev P x) [^] (2::nat)" + using assms eval_at_point_smult MP.nat_pow_closed Qp.int_inc_closed eval_at_point_nat_pow + by presburger + show "V\<^bsub>Q\<^sub>p\<^esub> n P \ basic_semialg_set n 2 q" + proof + fix x + assume A: "x \ V\<^bsub>Q\<^sub>p\<^esub> n P " + show "x \ basic_semialg_set n (2::nat) q " + proof- + have P: "Qp_ev P x = \" + using A zero_setE(2) + by blast + have "Qp_ev q x = \" + proof- + have "Qp_ev q x = \ \ (Qp_ev (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) x)" + using assms eval_at_point_smult[of x n "(P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" \] basic_semialg_set_def + by (meson A MP.nat_pow_closed Qp.int_inc_closed zero_setE(1)) + then show ?thesis + by (metis A P Qp.int_inc_closed Qp.integral_iff Qp.nat_pow_zero Qp.zero_closed assms(1) + eval_at_point_nat_pow neq0_conv zero_less_numeral zero_setE(1)) + qed + then have 0: "Qp_ev q x = \ \ Qp_ev q x \ P_set (2::nat)" + by blast + have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A zero_setE(1) + by blast + then show ?thesis using 0 basic_semialg_set_def' + by (metis (no_types, opaque_lifting) Qp.nat_pow_zero Qp.zero_closed + \eval_at_point Q\<^sub>p x q = \\ basic_semialg_set_memI zero_neq_numeral) + qed + qed + show "basic_semialg_set n 2 q \ V\<^bsub>Q\<^sub>p\<^esub> n P" + proof + fix x + assume A: "x \ basic_semialg_set n 2 q" + have 0: "\ Qp_ev q x \ P_set 2" + proof + assume "Qp_ev q x \ P_set 2" + then have 0: "Qp_ev q x \ nonzero Q\<^sub>p \ (\y \ carrier Q\<^sub>p . (y[^] (2::nat)) = Qp_ev q x)" + using P_set_def by blast + have "( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev P x) [^] (2::nat)) = y[^](2::nat))" + proof- + obtain y where y_def: "y \ carrier Q\<^sub>p \ (y[^] (2::nat)) = Qp_ev q x" + using 0 by blast + have "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A basic_semialg_set_memE(1) by blast + then have "Qp_ev q x = \ \ ((Qp_ev P x) [^] (2::nat))" + using assms eval_at_point_scalar_mult 00 by blast + then have "(y[^] (2::nat)) = \ \ ((Qp_ev P x) [^] (2::nat))" + using y_def by blast + then show ?thesis using y_def by blast + qed + then have "Qp_ev P x = \" + by (metis (no_types, lifting) A assms(1) basic_semialg_set_def mem_Collect_eq zero_set_semialg_set) + then have "Qp_ev q x = \" + using assms eval_at_point_smult + by (metis "00" A Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null basic_semialg_set_memE(1) zero_neq_numeral) + then show False + using 0 Qp.not_nonzero_memI by blast + qed + show " x \ V\<^bsub>Q\<^sub>p\<^esub> n P" + apply(rule zero_setI) + using A basic_semialg_set_memE(1) apply blast + using A 0 00[of x] + by (metis assms(1) basic_semialg_set_memE(1) basic_semialg_set_memE(2) zero_set_semialg_set) + qed +qed + +lemma is_zero_set_imp_basic_semialg: + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "S = zero_set Q\<^sub>p n P" + shows "is_basic_semialg n S" + unfolding is_basic_semialg_def +proof- + obtain q where q_def: "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" + by blast + have 0: "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" + using alg_as_semialg[of P n q] q_def assms(1) by linarith + have "(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms(1) + by blast + then have "\ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms q_def Qp.int_inc_closed local.smult_closed by blast + then have 1: "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + by (metis q_def ) + then show "\m. m \ 0 \ (\P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). S = basic_semialg_set n m P)" + using 0 assms + by (metis zero_neq_numeral) +qed + +lemma is_zero_set_imp_semialg: + assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "S = zero_set Q\<^sub>p n P" + shows "is_semialgebraic n S" + using assms(1) assms(2) basic_semialg_is_semialg is_semialgebraicI is_zero_set_imp_basic_semialg + by blast + +text\Algebraic sets are semialgebraic\ + +lemma is_algebraic_imp_is_semialg: + assumes "is_algebraic Q\<^sub>p n S" + shows "is_semialgebraic n S" + proof(rule is_semialgebraicI) + obtain ps where ps_def: "finite ps \ ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ S = affine_alg_set Q\<^sub>p n ps" + using is_algebraicE + by (metis assms) + have "ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n ps \ semialg_sets n" + apply(rule finite.induct[of ps]) + apply (simp add: ps_def) + using affine_alg_set_empty[of n] + apply (simp add: carrier_is_semialg) + proof + fix A a + assume IH: "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n A \ semialg_sets n" + assume P: "insert a A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + have "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using P by blast + then + show "affine_alg_set Q\<^sub>p n (insert a A) \ semialg_sets n" + using IH P semialg_intersect[of "affine_alg_set Q\<^sub>p n A" n "affine_alg_set Q\<^sub>p n {a}" ] + is_zero_set_imp_semialg affine_alg_set_insert[of n a A] + by (metis Int_commute affine_alg_set_singleton insert_subset is_semialgebraicE) + qed + then show "S \ semialg_sets n" + using ps_def by blast + qed + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Basic Lemmas about the Semialgebraic Predicate\ + (********************************************************************************************) + (********************************************************************************************) + + +text\Finite and cofinite sets are semialgebraic\ + +lemma finite_is_semialg: + assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "finite F" + shows "is_semialgebraic n F" + using Qp.finite_sets_are_algebraic is_algebraic_imp_is_semialg[of n F] + assms(1) assms(2) + by blast + +definition is_cofinite where +"is_cofinite n F = finite (ring_pow_comp Q\<^sub>p n F)" + +lemma is_cofiniteE: + assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "is_cofinite n F" + shows "finite (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)" + using assms(2) is_cofinite_def + by (simp add: ring_pow_comp_def) + +lemma complement_is_semialg: + assumes "is_semialgebraic n F" + shows "is_semialgebraic n ((carrier (Q\<^sub>p\<^bsup>n\<^esup>)) - F)" + using assms is_semialgebraic_def semialg_complement by blast + +lemma cofinite_is_semialgebraic: + assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "is_cofinite n F" + shows "is_semialgebraic n F" + using assms ring_pow_comp_inv[of F Q\<^sub>p n] complement_is_semialg[of n "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] + finite_is_semialg[of "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] is_cofiniteE[of F] + by (simp add: ring_pow_comp_def) + +lemma diff_is_semialgebraic: + assumes "is_semialgebraic n A" + assumes "is_semialgebraic n B" + shows "is_semialgebraic n (A - B)" + apply(rule is_semialgebraicI) + using assms unfolding semialg_sets_def + using gen_boolean_algebra_diff is_semialgebraicE semialg_sets_def + by blast + +lemma intersection_is_semialg: + assumes "is_semialgebraic n A" + assumes "is_semialgebraic n B" + shows "is_semialgebraic n (A \ B)" + using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_intersect + by blast + +lemma union_is_semialgebraic: + assumes "is_semialgebraic n A" + assumes "is_semialgebraic n B" + shows "is_semialgebraic n (A \ B)" + using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_union by blast + +lemma carrier_is_semialgebraic: +"is_semialgebraic n (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" +using carrier_is_semialg + by (simp add: carrier_is_semialg is_semialgebraic_def) + +lemma empty_is_semialgebraic: +"is_semialgebraic n {}" + by (simp add: empty_set_is_semialg is_semialgebraic_def) + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\One-Dimensional Semialgebraic Sets\ + (********************************************************************************************) + (********************************************************************************************) + +definition one_var_semialg where +"one_var_semialg S = ((to_R1 ` S) \ (semialg_sets 1))" + +definition univ_basic_semialg_set where +"univ_basic_semialg_set (m::nat) P = {a \ carrier Q\<^sub>p. (\y \ carrier Q\<^sub>p. (P \ a = (y[^]m)))}" + +text\Equivalence of univ\_basic\_semialg\_sets and semialgebraic subsets of $\mathbb{Q}^1$ \ + +lemma univ_basic_semialg_set_to_semialg_set: + assumes "P \ carrier Q\<^sub>p_x" + assumes "m \ 0" + shows "to_R1 ` (univ_basic_semialg_set m P) = basic_semialg_set 1 m (from_Qp_x P)" +proof + show "(\a. [a]) ` univ_basic_semialg_set m P \ basic_semialg_set 1 m (from_Qp_x P)" + proof fix x + assume A: "x \ (\a. [a]) ` univ_basic_semialg_set m P" + then obtain b y where by_def:"b \ carrier Q\<^sub>p \ y \ carrier Q\<^sub>p \ (P \ b) = (y[^]m) \ x = [b]" + unfolding univ_basic_semialg_set_def + by blast + then have "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using A Qp.to_R1_closed[of b] + unfolding univ_basic_semialg_set_def + by blast + then show "x \ basic_semialg_set 1 m (from_Qp_x P)" + using by_def Qp_x_Qp_poly_eval assms + unfolding basic_semialg_set_def + by blast + qed + show "basic_semialg_set 1 m (from_Qp_x P) \ (\a. [a]) ` univ_basic_semialg_set m P" + proof + fix x + assume A: "x \ basic_semialg_set 1 m (from_Qp_x P)" + then obtain b where b_def: "b \ carrier Q\<^sub>p \ x = [b]" + unfolding basic_semialg_set_def + by (metis (mono_tags, lifting) mem_Collect_eq Qp.to_R1_to_R Qp.to_R_pow_closed) + obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev (from_Qp_x P) [b] = (y[^]m))" + using A b_def + unfolding basic_semialg_set_def + by blast + have " P \ b = (y[^]m)" + using assms y_def b_def Qp_x_Qp_poly_eval by blast + then show " x \ (\a. [a]) ` univ_basic_semialg_set m P" + using y_def b_def + unfolding basic_semialg_set_def univ_basic_semialg_set_def + by blast + qed +qed + +definition is_univ_semialgebraic where +"is_univ_semialgebraic S = (S \ carrier Q\<^sub>p \ is_semialgebraic 1 (to_R1 ` S))" + +lemma is_univ_semialgebraicE: + assumes "is_univ_semialgebraic S" + shows "is_semialgebraic 1 (to_R1 ` S)" + using assms is_univ_semialgebraic_def by blast + +lemma is_univ_semialgebraicI: + assumes "is_semialgebraic 1 (to_R1 ` S)" + shows "is_univ_semialgebraic S" +proof- + have "S \ carrier Q\<^sub>p" + proof fix x assume "x \ S" + then have "(to_R1 x) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using assms + by (smt Collect_mono_iff gen_boolean_algebra_subset image_def is_semialgebraicE mem_Collect_eq semialg_sets_def Qp.to_R1_carrier) + then show "x \ carrier Q\<^sub>p" + using assms + by (metis nth_Cons_0 Qp.to_R_pow_closed) + qed + then show ?thesis + using assms + unfolding is_univ_semialgebraic_def + by blast +qed + +lemma univ_basic_semialg_set_is_univ_semialgebraic: + assumes "P \ carrier Q\<^sub>p_x" + assumes "m \ 0" + shows "is_univ_semialgebraic (univ_basic_semialg_set m P)" + using assms + by (metis (mono_tags, lifting) basic_semialg_is_semialgebraic' + from_Qp_x_closed is_univ_semialgebraic_def mem_Collect_eq subsetI + univ_basic_semialg_set_def univ_basic_semialg_set_to_semialg_set) + +lemma intersection_is_univ_semialgebraic: + assumes "is_univ_semialgebraic A" + assumes "is_univ_semialgebraic B" + shows "is_univ_semialgebraic (A \ B)" + using assms intersection_is_semialg[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] + unfolding is_univ_semialgebraic_def + by (metis le_infI1 Qp.to_R1_intersection) + +lemma union_is_univ_semialgebraic: + assumes "is_univ_semialgebraic A" + assumes "is_univ_semialgebraic B" + shows "is_univ_semialgebraic (A \ B)" + using assms union_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] + unfolding is_univ_semialgebraic_def + by (metis Un_subset_iff image_Un) + +lemma diff_is_univ_semialgebraic: + assumes "is_univ_semialgebraic A" + assumes "is_univ_semialgebraic B" + shows "is_univ_semialgebraic (A - B)" + using assms diff_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] + unfolding is_univ_semialgebraic_def + by (smt Diff_subset subset_trans Qp.to_R1_diff) + +lemma finite_is_univ_semialgebraic: + assumes "A \ carrier Q\<^sub>p" + assumes "finite A" + shows "is_univ_semialgebraic A" + using assms finite_is_semialg[of "((\a. [a]) ` A)" ] to_R1_finite[of A] + unfolding is_univ_semialgebraic_def + by (metis Qp.to_R1_carrier Qp.to_R1_subset) + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Defining the $p$-adic Valuation Semialgebraically\ + (********************************************************************************************) + (********************************************************************************************) + + +lemma Qp_square_root_criterion0: + assumes "p \ 2" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val a \ val b" + assumes "a \ \" + assumes "b \ \" + assumes "val a \ 0" + shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" +proof- + have 0: "(to_Zp a) \ carrier Z\<^sub>p" + using assms(2) to_Zp_closed + by blast + have 1: "(to_Zp b) \ carrier Z\<^sub>p" + using assms(3) to_Zp_closed + by blast + have 2: "a \ \\<^sub>p" + using val_ring_val_criterion assms(2) assms(5) assms(7) by blast + have 3: "b \ \\<^sub>p" + using assms val_ring_val_criterion[of b] dual_order.trans by blast + have 4: "val_Zp (to_Zp b) = val b" + using 3 Zp_def \_def padic_fields.to_Zp_val padic_fields_axioms by blast + have 5: "val_Zp (to_Zp a) = val a" + using Q\<^sub>p_def Zp_def assms(2) assms(7) padic_fields.Qp_val_ringI padic_fields.to_Zp_val padic_fields_axioms + by blast + have "\y \ carrier Z\<^sub>p. (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" + using 0 1 2 4 5 assms Zp_square_root_criterion[of "(to_Zp a)" "(to_Zp b)"] + by (metis "3" to_Zp_inc to_Zp_zero zero_in_val_ring) + then obtain y where y_def: "y \ carrier Z\<^sub>p \ (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" + by blast + have 6: "a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \ \b[^](2::nat) = ((\ y) [^] (2::nat))" + proof- + have 0: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = ((\ y) [^] (2::nat))" + using Qp_nonzero_nat_pow nat_pow_closed inc_pow nat_inc_zero inc_is_hom \_def y_def ring_hom_nat_pow[of Z\<^sub>p Q\<^sub>p \ y 2] + Q\<^sub>p_def Qp.ring_axioms Zp.ring_axioms + by blast + have 1: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = \ ((to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" + using y_def by presburger + have 2: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = \ ((to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ ( \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" + using "1" Zp.m_closed Zp_int_inc_closed assms(2) assms(3) inc_of_sum pow_closed to_Zp_closed by presburger + hence 3: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = (\ (to_Zp a))[^](2::nat) \ (\ \

) \ \ ((to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" + using Qp_nonzero_nat_pow nat_pow_closed inc_pow nat_inc_zero inc_is_hom \_def y_def ring_hom_nat_pow[of Z\<^sub>p Q\<^sub>p \ _ 2] + Q\<^sub>p_def Qp.ring_axioms Zp.ring_axioms Zp_int_inc_closed assms(2) assms(3) inc_of_prod pow_closed to_Zp_closed + by metis + then show ?thesis + using "0" "4" val_ring_ord_criterion assms(2) assms(3) assms(4) assms(5) + assms(6) assms(7) inc_pow not_nonzero_Zp ord_of_nonzero(1) p_inc to_Zp_closed to_Zp_inc + by (metis to_Zp_zero val_pos val_ringI zero_in_val_ring) + qed + have "(\ y) \ carrier Q\<^sub>p" + using frac_closed local.inc_def y_def inc_closed by blast + then show ?thesis + using 6 + by blast +qed + +lemma eint_minus_ineq': + assumes "(a::eint) \ b" + shows "a -b \ 0" + by (metis assms eint_minus_ineq eint_ord_simps(3) idiff_infinity idiff_self order_trans top.extremum_unique top_eint_def) + +lemma Qp_square_root_criterion: + assumes "p \ 2" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "ord b \ ord a" + assumes "a \ \" + assumes "b \ \" + shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" +proof- + have "\k::int. k \ min (ord a) (ord b) \ k mod 2 = 0" + proof- + let ?k = "if (min (ord a) (ord b)) mod 2 = 0 then min (ord a) (ord b) else (min (ord a) (ord b)) - 1" + have "?k \ min (ord a) (ord b) \ ?k mod 2 = 0" + apply(cases "(min (ord a) (ord b)) mod 2 = 0 ") + apply presburger + by presburger + then show ?thesis + by meson + qed + then obtain k where k_def: "k \ min (ord a) (ord b) \ k mod 2 = 0" + by meson + obtain a0 where a0_def: "a0 = (\[^](-k)) \ a" + by blast + obtain b0 where b0_def: "b0 = (\[^](-k)) \ b" + by blast + have 0: "a0 \ nonzero Q\<^sub>p" + using Qp.cring_axioms Qp.field_axioms Ring.integral a0_def assms(2) assms(5) cring_simprules(5) + not_nonzero_Qp p_intpow_closed(1) p_nonzero + by (metis Qp_int_pow_nonzero cring.cring_simprules(5)) + have 1: "val a0 = val a - k" + using a0_def assms(2) assms(5) val_mult p_nonzero p_intpow_closed(1) + by (metis Qp.m_comm Qp_int_pow_nonzero p_intpow_inv'' val_fract val_p_int_pow) + have 11: "val b0 = val b - k" + using assms(3) assms(6) b0_def val_mult p_nonzero p_intpow_closed(1) + by (metis Qp.m_lcomm Qp.one_closed Qp.r_one Qp_int_pow_nonzero p_intpow_inv'' val_fract val_p_int_pow) + have A: "val a \ k" + using k_def val_ord assms by (smt eint_ord_simps(1) not_nonzero_Qp) + have B: "val b \ k" + using k_def val_ord assms by (smt eint_ord_simps(1) not_nonzero_Qp) + then have 2: "val a0 \ 0" + using A 1 assms k_def eint_minus_ineq eint_ord_code(5) local.eint_minus_ineq' by presburger + have 3: "val a0 \ val b0" + using 1 11 assms + by (metis eint.distinct(2) eint_minus_ineq eint_ord_simps(1) val_def) + have 4: "a0 \ \" + using a0_def "0" Qp.nonzero_memE(2) by blast + have 5: "b0 \ \" + using b0_def + by (metis "4" Qp.integral_iff a0_def assms(2) assms(3) assms(6) p_intpow_closed(1)) + have "\y \ carrier Q\<^sub>p. a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat) = (y [^] (2::nat))" + using Qp_square_root_criterion0[of a0 b0] assms 2 3 4 5 b0_def a0_def Qp.m_closed p_intpow_closed(1) + by metis + then obtain y where y_def: " y \ carrier Q\<^sub>p \ a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat) = (y [^] (2::nat))" + by blast + then have 6: " (\[^] (2 * k)) \ (a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat)) = (\[^] (2 * k)) \ (y [^] (2::nat))" + by presburger + then have 8: "((\[^] (2 * k)) \ (a0[^](2::nat))) \\<^bsub>Q\<^sub>p\<^esub>((\[^] (2 * k)) \ (\\b0[^](2::nat))) = (\[^] (2 * k)) \ (y [^] (2::nat))" + using 6 Qp.r_distr[of "(a0[^](2::nat))" " (\\b0[^](2::nat))" "(\[^] (2 * k))"] + by (metis Qp.add.int_pow_closed Qp.m_closed Qp.nat_pow_closed Qp.one_closed a0_def assms(2) assms(3) b0_def p_inc p_intpow_closed(1) y_def) + have 9: "(\[^](int 2*k)) = (\[^]k)[^](2::nat)" + using Qp_int_nat_pow_pow[of \ k 2] + by (metis mult_of_nat_commute p_nonzero) + then have "((\[^]k)[^](2::nat) \ (a0[^](2::nat))) \\<^bsub>Q\<^sub>p\<^esub> (\[^]k)[^](2::nat) \ (\\b0[^](2::nat)) = (\[^]k)[^](2::nat) \ (y [^] (2::nat))" + by (metis "8" int_eq_iff_numeral) + then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\[^]k)[^](2::nat)) \ (\\b0[^](2::nat)) = ((\[^]k)[^](2::nat)) \ (y [^] (2::nat))" + by (metis Qp.cring_axioms a0_def assms(2) comm_monoid.nat_pow_distrib cring.cring_simprules(5) cring_def p_intpow_closed(1)) + then have 10: "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\[^]k)[^](2::nat)) \ (\\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" + using comm_monoid.nat_pow_distrib y_def + by (metis Qp.comm_monoid_axioms p_intpow_closed(1)) + then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((((\[^]k)[^](2::nat)) \ \)\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" + using 10 monoid.m_assoc[of Q\<^sub>p "((\[^]k)[^](2::nat))" \ " b0[^](2::nat)"] + by (metis Qp.int_inc_closed Qp.m_assoc Qp.m_closed Qp.nat_pow_closed assms(3) b0_def p_intpow_closed(1)) + then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\ \ ((\[^]k)[^](2::nat)) )\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" + by (metis Qp.group_commutes_pow Qp.int_inc_closed Qp.m_comm p_intpow_closed(1)) + then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ (((\[^]k)[^](2::nat)) \b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" + by (metis "10" Qp.int_inc_closed Qp.m_closed Qp.m_lcomm Qp.nat_pow_closed assms(3) b0_def p_intpow_closed(1)) + then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ ((\[^]k) \b0)[^](2::nat) = ((\[^]k) \ y) [^] (2::nat)" + by (metis Qp.m_closed Qp.nat_pow_distrib assms(3) b0_def p_intpow_closed(1)) + then have "a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ b[^](2::nat) = ((\[^]k) \ y) [^] (2::nat)" + by (metis Qp.l_one Qp.m_assoc a0_def assms(2) assms(3) b0_def p_intpow_closed(1) p_intpow_inv) + then show ?thesis + by (meson Qp.cring_axioms cring.cring_simprules(5) p_intpow_closed(1) y_def) +qed + +lemma Qp_val_ring_alt_def0: + assumes "a \ nonzero Q\<^sub>p" + assumes "ord a \ 0" + shows "\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" +proof- + have "\y \ carrier Z\<^sub>p. \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub> (3::nat))\\<^bsub>Z\<^sub>p\<^esub> ((to_Zp a) [^]\<^bsub>Z\<^sub>p\<^esub> (4::nat)) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" + using padic_integers.Zp_semialg_eq[of p "to_Zp a"] prime assms to_Zp_def + by (metis (no_types, lifting) Qp.nonzero_closed Qp.not_nonzero_memI Zp_def val_ring_ord_criterion not_nonzero_Zp padic_integers_axioms to_Zp_closed to_Zp_inc to_Zp_zero zero_in_val_ring) + then obtain y where y_def: "y \ carrier Z\<^sub>p \ \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub> (3::nat))\\<^bsub>Z\<^sub>p\<^esub> ((to_Zp a) [^]\<^bsub>Z\<^sub>p\<^esub> (4::nat)) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" + by blast + then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((\ y)[^](2::nat))" + using Group.nat_pow_0 Group.nat_pow_Suc nonzero_def + val_ring_ord_criterion assms inc_of_nonzero inc_of_prod inc_of_sum inc_pow + m_closed nat_inc_closed nat_pow_closed not_nonzero_Zp + numeral_2_eq_2 p_natpow_inc to_Zp_closed to_Zp_inc + by (smt Qp.nonzero_closed Qp.nonzero_memE(2) Zp.monom_term_car p_pow_nonzero(1) pow_closed to_Zp_zero zero_in_val_ring) + then have "(\ y) \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((\ y)[^](2::nat))" + using y_def inc_closed by blast + then show ?thesis + by blast +qed + +text\Defining the valuation semialgebraically for odd primes\ + +lemma P_set_ord_semialg_odd_p: + assumes "p \ 2" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" +proof(cases "a = \") + case True + show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" + proof + show "val b \ val a \ \y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + proof- + assume A: "val b \ val a" + then have "val b \ \" + by (metis True local.val_zero) + then have "b = \" + using assms(3) local.val_zero val_ineq by presburger + then have "(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (\[^](2::nat))" + using True + by (metis Qp.int_inc_zero Qp.int_nat_pow_rep Qp.nonzero_closed Qp.r_null Qp.r_zero assms(3) p_nonzero zero_power2) + then show ?thesis + using \b = \\ assms(3) by blast + qed + show "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat)) \ val b \ val a" + proof- + assume "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + then obtain y where y_def: "y \ carrier Q\<^sub>p \(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + by blast + then have 0: "\ \ (b[^](2::nat)) = (y[^](2::nat))" + by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.int_inc_closed Qp.nat_pow_closed + Qp.not_nonzero_memI Qp_nonzero_nat_pow True assms(2) assms(3) local.monom_term_car not_nonzero_Qp zero_less_numeral) + have "b = \" + apply(rule ccontr) + using 0 assms y_def p_times_square_not_square[of b] + unfolding P_set_def + by (metis (no_types, opaque_lifting) P_set_memI Qp.nat_pow_closed True + \b \ nonzero Q\<^sub>p \ \ \ b [^] 2 \ P_set 2\ not_nonzero_Qp p_times_square_not_square') + then show ?thesis + using eint_ord_code(3) local.val_zero by presburger + qed + qed +next + case False + then show ?thesis + proof(cases "b = \") + case True + then have "(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (a[^](2::nat))" + by (metis Qp.add.l_cancel_one' Qp.int_inc_zero Qp.int_nat_pow_rep Qp.nat_pow_closed Qp.nonzero_closed Qp.r_null assms(2) assms(3) p_nonzero zero_power2) + then have 0: "(\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" + using assms(2) + by blast + have 1: "val a \ val b" + using True assms local.val_zero eint_ord_code(3) by presburger + show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" + using 0 1 + by blast + next + case F: False + show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" + proof + show "val b \ val a \ \y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + proof- + assume "val b \ val a " + then have "ord b \ ord a" + using F False + by (metis eint_ord_simps(1) val_def) + then show "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + using assms Qp_square_root_criterion[of a b] False F + by blast + qed + show "\y\carrier Q\<^sub>p.(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat)) \ val b \ val a" + proof- + assume "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + then obtain y where y_def: "y \ carrier Q\<^sub>p \(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" + by blast + have 0: "ord (a[^](2::nat)) = 2* ord a" + by (metis (mono_tags, opaque_lifting) False Suc_1 assms(2) int_eq_iff_numeral nat_numeral + nonzero_nat_pow_ord not_nonzero_Qp) + have 1: "ord (\ \ (b[^](2::nat))) = 1 + 2* ord b" + proof- + have 0: "ord (\ \ (b[^](2::nat))) = ord \ + ord (b[^](2::nat))" + using F Qp_nat_pow_nonzero assms(3) not_nonzero_Qp ord_mult p_nonzero + by metis + have 1: "ord (b[^](2::nat)) = 2* ord b" + using F assms + by (metis (mono_tags, opaque_lifting) Suc_1 int_eq_iff_numeral nat_numeral + nonzero_nat_pow_ord not_nonzero_Qp) + then show ?thesis + using "0" ord_p + by linarith + qed + show "val b \ val a" + proof(rule ccontr) + assume "\ val b \ val a" + then have "val b \ val a \ val a \ val b" + by (metis linear) + then have "ord a > ord b" + using F False assms + by (metis \\ val a \ val b\ eint_ord_simps(1) le_less not_less_iff_gr_or_eq val_def) + then have "ord (a[^](2::nat)) > ord (\ \ (b[^](2::nat)))" + using 0 1 + by linarith + then have "ord ((a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat))) = ord (\ \ (b[^](2::nat)))" + by (meson F False Qp.int_inc_closed Qp_nat_pow_nonzero assms(2) assms(3) + local.monom_term_car not_nonzero_Qp ord_ultrametric_noteq p_times_square_not_square') + then have A0: "ord (y[^](2::nat)) = 1 + 2* ord b" + by (metis "1" \y \ carrier Q\<^sub>p \ (a[^]2) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^]2) = (y[^]2)\) + have A1: "(y[^](2::nat)) \ nonzero Q\<^sub>p" + using y_def 0 1 + by (smt F False Qp.nonzero_closed Qp_nat_pow_nonzero assms(2) assms(3) diff_ord_nonzero + local.monom_term_car not_nonzero_Qp p_nonzero p_times_square_not_square') + have A2: "y \ nonzero Q\<^sub>p" + using A1 Qp_nonzero_nat_pow pos2 y_def by blast + have A3: "ord (y[^](2::nat)) = 2* ord y" + using A2 nonzero_nat_pow_ord + by presburger + then show False using A0 + by presburger + qed + qed + qed + qed +qed + +text\Defining the valuation ring semialgebraically for all primes\ + +lemma Qp_val_ring_alt_def: + assumes "a \ carrier Q\<^sub>p" + shows "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" +proof(cases "a = \") + case True + then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = \" + by (metis Qp.add.l_cancel_one' Qp.integral_iff Qp.nat_pow_closed Qp.not_nonzero_memI + Qp.one_closed Qp_nonzero_nat_pow assms not_nonzero_Qp p_natpow_closed(1) zero_less_numeral) + then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (\[^](2::nat))" + using Qp.nat_pow_one by blast + then show ?thesis + using True zero_in_val_ring by blast +next + case False + show "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" + proof + show "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" + using assms Qp_val_ring_alt_def0[of a] False + by (meson not_nonzero_Qp ord_nonneg) + show "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))) \ a \ \\<^sub>p" + proof- + assume "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" + then obtain y where y_def: "y \ carrier Q\<^sub>p \\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + by blast + then have "(\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" + using Qp.ring_simprules + by (smt Qp.nat_pow_closed assms p_natpow_closed(1)) + then have "ord ((\[^](3::nat))\ (a[^](4::nat))) = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" + by presburger + then have "3 + ord (a[^](4::nat)) = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" + by (metis False Qp_nat_pow_nonzero assms not_nonzero_Qp of_nat_numeral ord_mult ord_p_pow_nat p_nonzero) + then have 0: "3 + 4* ord a = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" + using assms False nonzero_nat_pow_ord[of a "(4::nat)"] + by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) + have "ord a \ 0" + proof(rule ccontr) + assume "\ 0 \ ord a" + then have 00: "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) < 0" + using 0 + by linarith + have yn: "y \ nonzero Q\<^sub>p" + apply(rule ccontr) + using y_def 0 + by (metis "00" Qp.not_eq_diff_nonzero Qp.one_closed Qp.one_nonzero Qp.pow_zero + \\ [^] 3 \ a [^] 4 = y [^] 2 \ \\ diff_ord_nonzero less_numeral_extra(3) + local.one_neq_zero not_nonzero_Qp ord_one zero_less_numeral) + then have "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) = ord (y[^](2::nat))" + using y_def ord_ultrametric_noteq''[of "(y[^](2::nat))" "\" ] + by (metis "00" False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_pow_nonzero + Qp.not_eq_diff_nonzero Qp.one_nonzero Qp.r_right_minus_eq \\ [^] 3 \ a [^] 4 = y [^] 2 \ \\ + assms ord_one ord_ultrametric_noteq p_nonzero) + then have "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) = 2* ord y" + using y_def Qp_nat_pow_nonzero Qp_nonzero_nat_pow nonzero_nat_pow_ord[of y "(2::nat)"] yn + by linarith + then have "3 + (4* ord a) = 2* ord y" + using "00" "0" + by linarith + then show False + by presburger + qed + then show "a \ \\<^sub>p" + using False val_ring_ord_criterion assms by blast + qed + qed +qed + +lemma Qp_val_alt_def: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "val b \ val a \ (\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" +proof + show "val a \ val b \ \y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + proof- + assume A: "val a \ val b" + show "\y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + proof(cases "b = \") + case True + then have "a = \" + using A assms(1) val_ineq + by blast + then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (\[^](2::nat))" + by (metis Qp.nat_pow_zero Qp.r_null Qp.r_zero True assms(2) p_natpow_closed(1) zero_neq_numeral) + then show ?thesis + using True A assms(2) + by blast + next + case False + assume B: "b \ \" + show "\y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat)) \ (a[^](4::nat)) = (y[^](2::nat))" + proof(cases "a = \") + case True + then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (b[^](4::nat))" + using Qp.cring_axioms Qp.nat_pow_closed assms(2) cring_def p_natpow_closed(1) ring.pow_zero zero_less_numeral + by (metis Qp.add.l_cancel_one' Qp.integral_iff assms(1)) + then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((b[^](2::nat))[^] (2::nat))" + by (metis Qp_nat_pow_pow assms(2) mult_2_right numeral_Bit0) + then have "(b[^](2::nat)) \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((b[^](2::nat))[^] (2::nat))" + using Qp.nat_pow_closed assms(2) + by blast + then show ?thesis + by blast + next + case False + have F0: "b \ nonzero Q\<^sub>p" + using B assms(2) not_nonzero_Qp + by metis + have F1: "a \ nonzero Q\<^sub>p" + using False assms(1) not_nonzero_Qp + by metis + then have "(a \

b) \ nonzero Q\<^sub>p" + using B + by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) inv_in_frac(3)) + then have "val a \ val b" + using F0 F1 A by blast + then have "val (a \
b) \ 0" + using F0 F1 val_fract assms(1) local.eint_minus_ineq' by presburger + obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat)) = (y[^](2::nat))" + using Qp_val_ring_alt_def0 + by (meson B False Qp.integral Qp.nonzero_closed \(a \
b) \ nonzero Q\<^sub>p\ \0 \ val (a \
b)\ + assms(1) assms(2) inv_in_frac(1) inv_in_frac(2) ord_nonneg val_ringI) + then have "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = + (b[^](4::nat)) \ (y[^](2::nat))" + by presburger + then have F2: "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = + ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" + by (metis Qp.nat_pow_pow assms(2) mult_2_right numeral_Bit0) + have F3: "((b[^](4::nat)) \ \) \\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = + ((b[^](2::nat))[^] (2::nat)) \ (y[^](2::nat))" + proof- + have 0: "(\[^](3::nat)) \ (a \
b[^](4::nat)) \ carrier Q\<^sub>p " + proof- + have "(a \
b[^](4::nat)) \ carrier Q\<^sub>p" + using F0 Qp.nat_pow_closed assms(1) fract_closed Qp_nat_pow_nonzero by presburger + then show ?thesis + by (meson Qp.cring_axioms cring.cring_simprules(5) p_natpow_closed(1)) + qed + have 1: "(b[^](4::nat)) \ carrier Q\<^sub>p" + using Qp.nat_pow_closed assms(2) + by blast + then show ?thesis + using 0 F2 ring.ring_simprules(23)[of Q\<^sub>p "\" "(\[^](3::nat))\ ((a \
b)[^](4::nat))" "(b[^](4::nat))"] + Qp.cring_axioms Qp.nonzero_mult_closed Qp.ring_axioms Qp_nat_pow_nonzero \(a \
b) \ nonzero Q\<^sub>p\ p_nonzero + by blast + qed + have F4: "(b[^](4::nat)) \ carrier Q\<^sub>p" + using Qp.nat_pow_closed assms(2) + by blast + then have "((b[^](4::nat)) \ \) = (b[^](4::nat))" + using Qp.r_one by blast + then have F5: "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = + ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" + using F3 + by presburger + have "((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat)))" + proof- + have 0: "(b[^](4::nat)) \ carrier Q\<^sub>p" + using F4 by blast + have 1: "(\[^](3::nat)) \ carrier Q\<^sub>p" + by blast + have 2: "((a \
b)[^](4::nat)) \ carrier Q\<^sub>p" + using F0 Qp.nat_pow_closed assms(1) fract_closed + by blast + show ?thesis using 0 1 2 monoid.m_assoc[of Q\<^sub>p] comm_monoid.m_comm[of Q\<^sub>p] + using Qp.m_lcomm by presburger + qed + then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat))) = + ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" + using F5 by presburger + then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b \(a \
b))[^](4::nat)) = + ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" + using F0 Qp.nat_pow_distrib assms(1) assms(2) fract_closed by presburger + then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = + ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" + by (metis F0 assms(1) local.fract_cancel_right) + then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = + (((b[^](2::nat))\ y)[^](2::nat))" + using Qp.nat_pow_closed Qp.nat_pow_distrib assms(2) y_def by blast + then have "((b[^](2::nat))\ y) \ carrier Q\<^sub>p \ (b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = + (((b[^](2::nat))\ y)[^](2::nat))" + by (meson Qp.cring_axioms Qp.nat_pow_closed assms(2) cring.cring_simprules(5) y_def) + then show ?thesis + by blast + qed + qed + qed + show "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \ val a \ val b" + proof- + assume A: "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + show "val a \ val b" + proof(cases "a = \") + case True + then show ?thesis + using eint_ord_code(3) local.val_zero by presburger + next + case False + have "b \ \" + proof(rule ccontr) + assume "\ b \ \" + then have "\y \ carrier Q\<^sub>p. (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + using A + by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.nat_pow_closed Qp.nonzero_memE(2) + Qp_nonzero_nat_pow assms(1) assms(2) local.monom_term_car not_nonzero_Qp + p_natpow_closed(1) zero_less_numeral) + then obtain y where y_def: "y \ carrier Q\<^sub>p \ (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + by blast + have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" + proof- + have 00: "(\[^](3::nat)) \ nonzero Q\<^sub>p" + using Qp_nat_pow_nonzero p_nonzero by blast + have 01: "(a[^](4::nat)) \ nonzero Q\<^sub>p" + using False Qp_nat_pow_nonzero assms(1) not_nonzero_Qp Qp.nonzero_memI by presburger + then show ?thesis using ord_mult[of "\[^](3::nat)" "a[^](4::nat)"] + by (metis (no_types, lifting) "00" False assms(1) nonzero_nat_pow_ord + not_nonzero_Qp of_nat_numeral ord_p_pow_nat) + qed + have 1: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 2* (ord y)" + proof- + have "y \ \" + proof(rule ccontr) + assume " \ y \ \" + then have "(\[^](3::nat))\ (a[^](4::nat)) = \" + using y_def Qp.cring_axioms cring_def pos2 ring.pow_zero by blast + then show False + by (metis False Qp.integral Qp.nat_pow_closed Qp.nonzero_pow_nonzero + Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) p_natpow_closed(1) p_nonzero) + qed + then show ?thesis + using y_def + by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) + qed + then show False + using 0 + by presburger + qed + then have F0: "b \ nonzero Q\<^sub>p" + using assms(2) not_nonzero_Qp by metis + + have F1: "a \ nonzero Q\<^sub>p" + using False assms(1) not_nonzero_Qp by metis + obtain y where y_def: "y \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" + using A by blast + show ?thesis + proof(rule ccontr) + assume " \ val a \ val b " + then have F2: "ord a < ord b" + using F0 F1 assms + by (metis False \b \ \\ eint_ord_simps(1) leI val_def) + have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" + using F0 ord_mult F1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_p_pow_nat p_natpow_closed(2) + by presburger + have 1: " ord (b[^](4::nat)) = 4* ord b" + using F0 nonzero_nat_pow_ord + by presburger + have 2: "(4 * (ord b)) > 4 * (ord a)" + using F2 by linarith + have 3: "(4 * (ord b)) \ 3 + 4* ord a" + proof(rule ccontr) + assume "\ (4 * (ord b)) \ 3 + 4* ord a" + then have "(4 * (ord b)) > 3 + 4* ord a" + by linarith + then have 30: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" + using "0" "1" F0 F1 Qp_nat_pow_nonzero Qp.nat_pow_closed assms(1) monom_term_car not_nonzero_Qp ord_ultrametric_noteq + p_natpow_closed(1) p_nonzero + by (metis Qp.integral) + have "y \ nonzero Q\<^sub>p" + proof(rule ccontr) + assume A: "y \ nonzero Q\<^sub>p" + then have "y = \" + using y_def Qp.nonzero_memI by blast + then have "b [^] 4 \ \ [^] 3 \ a [^] 4 = \" + by (smt "0" "1" A F0 False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed + Qp.nonzero_mult_closed Qp.nonzero_pow_nonzero Qp.pow_zero assms(1) diff_ord_nonzero not_nonzero_Qp p_nonzero pos2 y_def) + then show False + by (smt "0" "1" A F0 F1 Qp.integral Qp.nat_pow_closed Qp.nonzero_mult_closed + Qp_nat_pow_nonzero assms(1) diff_ord_nonzero not_nonzero_Qp p_natpow_closed(1) p_nonzero y_def) + qed + then have 31: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 2* ord y" + using nonzero_nat_pow_ord y_def + by presburger + then show False using 30 by presburger + qed + show False + using 2 3 + by presburger + qed + qed + qed +qed + +text\The polynomial in two variables which semialgebraically defines the valuation relation\ + +definition Qp_val_poly where +"Qp_val_poly = (pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" + +lemma Qp_val_poly_closed: +"Qp_val_poly \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" +proof- + have "(pvar Q\<^sub>p 1) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + using local.pvar_closed one_less_numeral_iff semiring_norm(76) by blast + then have 0: "(pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + + using ring.Pring_is_ring[of Q\<^sub>p "{0::nat..2-1}"] + monoid.nat_pow_closed[of "coord_ring Q\<^sub>p 2"] Qp.cring_axioms cring.axioms(1) ring.Pring_is_monoid + by blast + have 1: "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + using local.pvar_closed pos2 by blast + have 2: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + using 1 local.smult_closed p_natpow_closed(1) by blast + then show ?thesis + unfolding Qp_val_poly_def + using 0 by blast +qed + +lemma Qp_val_poly_eval: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "Qp_ev Qp_val_poly [a, b] = (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))" +proof- + have 0: "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + proof(rule cartesian_power_car_memI) + show "length [a, b] = 2" + by simp + have "set [a,b] = {a,b}" + by auto + then show "set [a, b] \ carrier Q\<^sub>p" + using assms + by (simp add: \a \ carrier Q\<^sub>p\ \b \ carrier Q\<^sub>p\) + qed + obtain f where f_def: "f = ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))" + by blast + obtain g where g_def: "g = (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" + by blast + have 1: "Qp_val_poly = f \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> g" + unfolding Qp_val_poly_def + using f_def g_def by blast + have 1: "Qp_ev (pvar Q\<^sub>p (0::nat)) [a,b] = a" + using eval_pvar + by (metis \[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)\ nth_Cons_0 pos2) + have 2: "Qp_ev (pvar Q\<^sub>p (1::nat)) [a,b] = b" + using eval_pvar + by (metis (no_types, lifting) "0" One_nat_def add_diff_cancel_right' assms(2) + cartesian_power_car_memE gr_zeroI less_numeral_extra(1) less_numeral_extra(4) + list.size(4) nth_Cons_pos Qp.to_R1_closed Qp.to_R_to_R1 zero_less_diff) + have 3: "Qp_ev ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (b[^](4::nat))" + by (metis "0" "2" eval_at_point_nat_pow local.pvar_closed one_less_numeral_iff semiring_norm(76)) + have 4: "Qp_ev ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (a[^](4::nat))" + using "0" "1" eval_at_point_nat_pow local.pvar_closed pos2 by presburger + then have 5: "Qp_ev (poly_scalar_mult Q\<^sub>p (\[^](3::nat)) ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))) [a,b] = (\[^](3::nat))\ (a[^](4::nat))" + using eval_at_point_smult[of "[a,b]" 2 "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)" "\[^](3::nat)" ] 2 + by (metis "0" MP.nat_pow_closed eval_at_point_scalar_mult local.pvar_closed p_natpow_closed(1) zero_less_numeral) + then show ?thesis + proof- + have 00: "[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + by (simp add: "0") + have 01: " pvar Q\<^sub>p 1 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + by (meson MP.nat_pow_closed local.pvar_closed one_less_numeral_iff semiring_norm(76)) + have 02: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (pvar Q\<^sub>p 0 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat)) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" + by (meson MP.nat_pow_closed local.pvar_closed local.smult_closed p_natpow_closed(1) zero_less_numeral) + then show ?thesis + unfolding Qp_val_poly_def + using 00 01 02 + by (metis (no_types, lifting) "3" "4" MP.nat_pow_closed eval_at_point_add eval_at_point_smult + local.pvar_closed p_natpow_closed(1) zero_less_numeral) + qed +qed + +lemma Qp_2I: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms + apply (simp add: assms(1) assms(2)) + using assms + by (simp add: assms(1) assms(2)) + +lemma pair_id: + assumes "length as = 2" + shows "as = [as!0, as!1]" + using assms + by (smt One_nat_def diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) + nth_Cons' nth_equalityI numeral_2_eq_2) + +lemma Qp_val_semialg: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "val b \ val a \ [a,b] \ basic_semialg_set 2 (2::nat) Qp_val_poly" +proof + show "val a \ val b \ [a, b] \ basic_semialg_set 2 2 Qp_val_poly" + using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] + unfolding basic_semialg_set_def + by (metis (mono_tags, lifting) assms(1) assms(2) mem_Collect_eq) + show "[a, b] \ basic_semialg_set 2 2 Qp_val_poly \ val a \ val b" + using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] + unfolding basic_semialg_set_def + using assms(1) assms(2) + by blast +qed + +definition val_relation_set where +"val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as!1) \ val (as!0)}" + +lemma val_relation_setE: + assumes "as \ val_relation_set" + shows "as!0 \ carrier Q\<^sub>p \ as!1 \ carrier Q\<^sub>p \ as = [as!0,as!1] \ val (as!1) \ val (as!0)" + using assms unfolding val_relation_set_def + by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff pair_id pos2 semiring_norm(76)) + +lemma val_relation_setI: + assumes "as!0 \ carrier Q\<^sub>p" + assumes "as!1 \ carrier Q\<^sub>p" + assumes "length as = 2" + assumes "val (as!1) \ val(as!0)" + shows "as \ val_relation_set" + unfolding val_relation_set_def using assms Qp_2I[of "as!0" "as!1"] + by (metis (no_types, lifting) mem_Collect_eq pair_id) + +lemma val_relation_semialg: +"val_relation_set = basic_semialg_set 2 (2::nat) Qp_val_poly" +proof + show "val_relation_set \ basic_semialg_set 2 (2::nat) Qp_val_poly" + proof fix as + assume A: "as \ val_relation_set" + have 0: "length as = 2" + unfolding val_relation_set_def + by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq val_relation_set_def) + have 1: "as = [as ! 0, as ! 1]" + by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq pair_id val_relation_set_def) + show "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" + using A 1 val_relation_setE[of as] Qp_val_semialg[of "as!0" "as!1"] + by presburger + qed + show "basic_semialg_set 2 (2::nat) Qp_val_poly \ val_relation_set" + proof + fix as + assume "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" + then show "as \ val_relation_set" + using val_relation_setI[of as] + by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq + one_less_numeral_iff Qp_val_semialg basic_semialg_set_def + val_relation_set_def padic_fields_axioms pair_id pos2 semiring_norm(76)) + qed +qed + +lemma val_relation_is_semialgebraic: +"is_semialgebraic 2 val_relation_set" +proof - +have "{rs \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (rs ! 0) \ val (rs ! 1)} = basic_semialg_set (Suc 1) (Suc 1) Qp_val_poly" +using Suc_1 val_relation_semialg val_relation_set_def by presburger + then show ?thesis + by (metis (no_types) Qp_val_poly_closed Suc_1 basic_semialg_is_semialgebraic' val_relation_set_def zero_neq_numeral) +qed + +lemma Qp_val_ring_is_semialg: + obtains P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" +proof- + obtain P where P_def: "P = (\[^](3::nat)) \\<^bsub>Q\<^sub>p_x \<^esub>(X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat) \\<^bsub>Q\<^sub>p_x\<^esub> \\<^bsub>Q\<^sub>p_x\<^esub>" + by blast + have 0: "P \ carrier Q\<^sub>p_x" + proof- + have 0: "(X_poly Q\<^sub>p) \ carrier Q\<^sub>p_x" + using UPQ.X_closed by blast + then show ?thesis + using P_def UPQ.P.nat_pow_closed p_natpow_closed(1) by blast + qed + have 1: "\\<^sub>p = univ_basic_semialg_set 2 P" + proof + show "\\<^sub>p \ univ_basic_semialg_set 2 P" + proof + fix x + assume A: "x \ \\<^sub>p" + show "x \ univ_basic_semialg_set 2 P" + proof- + have x_car: "x \ carrier Q\<^sub>p" + using A val_ring_memE by blast + then have "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat)))" + using A Qp_val_ring_alt_def[of x] + by blast + then obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat))" + by blast + have "y \ carrier Q\<^sub>p \ P \ x = (y[^](2::nat))" + proof- + have "P \ x = \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat))" + proof- + have "((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat)) \ carrier Q\<^sub>p_x" + using UPQ.monom_closed p_natpow_closed(1) by blast + then have "P \ x = (((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" + using P_def x_car UPQ.to_fun_plus by blast + then have 0: "P \ x = (\[^](3::nat)) \(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" + using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_smult p_natpow_closed(1) x_car by presburger + have "(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) = (x[^](4::nat))" + using UPQ.to_fun_X_pow x_car by blast + then have "P \ x = (\[^](3::nat)) \(x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" + using "0" UPQ.to_fun_one x_car by presburger + then show ?thesis + using y_def Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1) x_car + by presburger + qed + then show ?thesis + using y_def + by blast + qed + then show ?thesis + unfolding univ_basic_semialg_set_def + using x_car + by blast + qed + qed + show "univ_basic_semialg_set 2 P \ \\<^sub>p" + proof fix x + assume A: "x \ univ_basic_semialg_set (2::nat) P" + then obtain y where y_def: "y \ carrier Q\<^sub>p \ (P \ x) = (y[^](2::nat))" + unfolding univ_basic_semialg_set_def + by blast + have x_car: "x \ carrier Q\<^sub>p" + using A + by (metis (no_types, lifting) mem_Collect_eq univ_basic_semialg_set_def) + have 0: "(P \ x) = (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" + using P_def x_car UPQ.UP_one_closed UPQ.monom_closed UPQ.monom_rep_X_pow UPQ.to_fun_monom + UPQ.to_fun_one UPQ.to_fun_plus p_natpow_closed(1) by presburger + have 1: "y \ carrier Q\<^sub>p \ (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ = (y[^](2::nat))" + using "0" y_def + by blast + then show "x \ \\<^sub>p" + using x_car Qp_val_ring_alt_def[of x] y_def + by (metis Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1)) + qed + qed + show ?thesis + using 0 1 that + by blast +qed + +lemma Qp_val_ring_is_univ_semialgebraic: +"is_univ_semialgebraic \\<^sub>p" +proof- + obtain P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" + using Qp_val_ring_is_semialg by blast + then show ?thesis + by (metis univ_basic_semialg_set_is_univ_semialgebraic zero_neq_numeral) +qed + +lemma Qp_val_ring_is_semialgebraic: +"is_semialgebraic 1 (to_R1` \\<^sub>p)" + using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def by blast + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Inverse Images of Semialgebraic Sets by Polynomial Maps\ + (********************************************************************************************) + (********************************************************************************************) + +lemma basic_semialg_pullback: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" + assumes "is_poly_tuple n fs" + assumes "length fs = k" + assumes "S = basic_semialg_set k m f" + assumes "m \0" + shows "poly_map n fs \\<^bsub>n\<^esub> S = basic_semialg_set n m (Qp_poly_comp n fs f)" +proof + show "poly_map n fs \\<^bsub>n\<^esub> S \ basic_semialg_set n m (Qp_poly_comp n fs f)" + proof + fix x + assume A: "x \ poly_map n fs \\<^bsub>n\<^esub> S" + then have 0: "poly_map n fs x \ S" + proof - + have "\n f. {rs. rs \ S} \ {rs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). \r. r \ carrier Q\<^sub>p \ Qp_ev f rs = (r[^](n::nat))}" + by (metis (no_types) Collect_mem_eq \S = basic_semialg_set k m f\ basic_semialg_set_def eq_iff) + then show ?thesis + using A by blast + qed + have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A assms + by (meson evimage_eq) + have "\y \ (carrier Q\<^sub>p). Qp_ev f (poly_map n fs x) = (y[^]m)" + using A 0 assms basic_semialg_set_def + by blast + then have "\y \ (carrier Q\<^sub>p). Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" + using 1 assms Qp_poly_comp_eval + by blast + then show "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" + using "1" basic_semialg_set_def + by blast + qed + show "basic_semialg_set n m (Qp_poly_comp n fs f) \ poly_map n fs \\<^bsub>n\<^esub> S" + proof fix x + assume A: "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" + have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A basic_semialg_set_def + by blast + have 1: "(poly_map n fs x) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + using "0" poly_map_closed assms(2) assms(3) by blast + show "x \ poly_map n fs \\<^bsub>n\<^esub> S" + proof- + have "\y \ carrier Q\<^sub>p. Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" + using A basic_semialg_set_def + by blast + then have 2: "\y \ carrier Q\<^sub>p. Qp_ev f (poly_map n fs x) = (y[^]m)" + using assms Qp_poly_comp_eval + by (metis (no_types, lifting) A basic_semialg_set_def mem_Collect_eq) + have 3: "poly_map n fs x \ S" + using assms 0 1 basic_semialg_set_def[of k m f] "2" + by blast + show ?thesis + using "0" "3" by blast + qed + qed +qed + +lemma basic_semialg_pullback': + assumes "is_poly_tuple n fs" + assumes "length fs = k" + assumes "A \ basic_semialgs k" + shows "poly_map n fs \\<^bsub>n\<^esub> A \ (basic_semialgs n)" +proof- + obtain f m where fm_def: "m \0 \f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>]) \ A = basic_semialg_set k m f" + using assms + by (metis is_basic_semialg_def mem_Collect_eq) + then have "poly_map n fs \\<^bsub>n\<^esub> A = basic_semialg_set n m (Qp_poly_comp n fs f)" + using assms basic_semialg_pullback[of f k n fs A m] + by linarith + then show ?thesis unfolding is_basic_semialg_def + by (metis (mono_tags, lifting) assms(1) assms(2) fm_def mem_Collect_eq poly_compose_closed) +qed + +lemma semialg_pullback: + assumes "is_poly_tuple n fs" + assumes "length fs = k" + assumes "S \ semialg_sets k" + shows "poly_map n fs \\<^bsub>n\<^esub> S \ semialg_sets n" + unfolding semialg_sets_def + apply(rule gen_boolean_algebra.induct[of S "(carrier (Q\<^sub>p\<^bsup>k\<^esup>))" "basic_semialgs k"]) + using assms semialg_sets_def apply blast + apply (metis assms(1) assms(2) carrier_is_semialgebraic evimageI2 extensional_vimage_closed is_semialgebraicE poly_map_closed semialg_sets_def subsetI subset_antisym) + apply (metis Int_absorb2 assms(1) assms(2) basic_semialg_is_semialg basic_semialg_is_semialgebraic basic_semialg_pullback' is_semialgebraic_closed mem_Collect_eq semialg_sets_def) + apply (metis evimage_Un semialg_sets_def semialg_union) + by (metis assms(1) assms(2) carrier_is_semialgebraic diff_is_semialgebraic evimage_Diff extensional_vimage_closed is_semialgebraicE is_semialgebraicI poly_map_closed poly_map_pullbackI semialg_sets_def subsetI subset_antisym) + +lemma pullback_is_semialg: + assumes "is_poly_tuple n fs" + assumes "length fs = k" + assumes "S \ semialg_sets k" + shows "is_semialgebraic n (poly_map n fs \\<^bsub>n\<^esub> S)" + using assms(1) assms(2) assms(3) is_semialgebraicI padic_fields_axioms semialg_pullback + by blast + +text\Equality and inequality sets for a pair of polynomials\ + +definition val_ineq_set where +"val_ineq_set n f g = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" + +lemma poly_map_length : + assumes "length fs = m" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "length (poly_map n fs as) = m" + using assms unfolding poly_map_def poly_tuple_eval_def + by (metis (no_types, lifting) length_map restrict_apply') + +lemma val_ineq_set_pullback: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "val_ineq_set n f g = poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set " +proof + show "val_ineq_set n f g \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" + proof + fix x + assume "x \ val_ineq_set n f g" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ val (Qp_ev f x) \ val (Qp_ev g x)" + by (metis (mono_tags, lifting) mem_Collect_eq val_ineq_set_def) + have 1: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" + unfolding poly_map_def poly_tuple_eval_def using 0 + by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') + have 2: "poly_map n [g,f] x \ val_relation_set" + apply(rule val_relation_setI) + using 1 0 assms apply (metis eval_at_point_closed nth_Cons_0) + using 1 0 assms apply (metis One_nat_def eval_at_point_closed diff_Suc_1 less_numeral_extra(1) nth_Cons_pos Qp.to_R_to_R1) + using poly_map_length assms 0 apply (metis "1" Qp_2I cartesian_power_car_memE eval_at_point_closed) + by (metis "0" "1" One_nat_def nth_Cons_0 nth_Cons_Suc) + have 3: "is_poly_tuple n [g, f]" + using assms + by (smt One_nat_def diff_Suc_1 Qp_is_poly_tupleI length_Suc_conv less_SucE less_one list.size(3) nth_Cons') + then show "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" + using 0 1 2 + by blast + qed + show "poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set \ val_ineq_set n f g" + proof fix x + have 0: "is_poly_tuple n [g, f]" + using Qp_is_poly_tupleI assms + by (metis (no_types, lifting) diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons') + assume A: "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" + then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ poly_map n [g,f] x \ val_relation_set" + using 0 + by (meson evimageD extensional_vimage_closed subsetD) + have 2: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" + by (metis "1" Qp_poly_mapE' length_0_conv poly_map_cons) + show "x \ val_ineq_set n f g" + using 0 1 2 unfolding val_ineq_set_def val_relation_set_def + by (metis (no_types, lifting) "1" list.inject mem_Collect_eq nth_Cons_0 poly_map_apply val_relation_setE) + qed +qed + +lemma val_ineq_set_is_semialg: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "val_ineq_set n f g \ semialg_sets n" +proof- + have 0: "val_relation_set \ semialg_sets 2" + using val_relation_semialg basic_semialg_is_semialg' + by (metis Qp_val_poly_closed zero_neq_numeral) + show ?thesis using val_ineq_set_pullback semialg_pullback[of n "[g,f]" 2 "val_relation_set" ] + by (metis (no_types, lifting) "0" assms(1) assms(2) diff_Suc_1 Qp_is_poly_tupleI + length_Cons less_Suc0 less_SucE list.size(3) nth_Cons_0 nth_Cons_pos numeral_2_eq_2 + zero_neq_numeral) +qed + +lemma val_ineq_set_is_semialgebraic: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n (val_ineq_set n f g)" + using assms(1) assms(2) is_semialgebraicI val_ineq_set_is_semialg by blast + +lemma val_ineq_setI: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "x \ (val_ineq_set n f g)" + shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + "val (Qp_ev f x) \ val (Qp_ev g x)" + using assms unfolding val_ineq_set_def apply blast + using assms unfolding val_ineq_set_def by blast + +lemma val_ineq_setE: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "val (Qp_ev f x) \ val (Qp_ev g x)" + shows "x \ (val_ineq_set n f g)" + using assms unfolding val_ineq_set_def + by blast + +lemma val_ineq_set_is_semialgebraic': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" + using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast + +lemma val_eq_set_is_semialgebraic: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" +proof- + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" + using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def + by blast + have 1: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" + using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def + by blast + have 2: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} \ + {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" + apply(rule equalityI, rule subsetI , rule IntI) unfolding mem_Collect_eq + using le_less apply blast apply (metis order_refl) + apply(rule subsetI, erule IntE) unfolding mem_Collect_eq + by (meson less_le_trans not_less_iff_gr_or_eq) + show ?thesis unfolding 2 apply(rule intersection_is_semialg) + using 0 apply blast using 1 by blast +qed + +lemma equalityI'': + assumes "\x. A x \ B x" + assumes "\x. B x \ A x" + shows "{x. A x} = {x. B x}" + using assms by blast + +lemma val_strict_ineq_set_is_semialgebraic: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = + {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} - {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" + apply(rule equalityI', rule DiffI) unfolding le_less mem_Collect_eq apply blast + unfolding mem_Collect_eq using neq_iff apply blast + apply(erule DiffE) unfolding mem_Collect_eq by blast + show ?thesis unfolding 0 + apply(rule diff_is_semialgebraic) + using assms val_ineq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def apply blast + using assms val_eq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def by blast +qed + +lemma constant_poly_val_exists: + shows "\g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" +proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ val a = c" + by (meson Qp.minus_closed Qp.nonzero_closed dist_nonempty' p_nonzero) + obtain g where g_def: "g = coord_const a" + by blast + show ?thesis using a_def g_def Qp_to_IP_car + by (metis (no_types, opaque_lifting) Qp_to_IP_car a_def eval_at_point_const g_def le_less subset_iff) +qed + +lemma val_ineq_set_is_semialgebraic'': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" +proof- + obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" + using constant_poly_val_exists by blast + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" + apply(rule val_ineq_set_is_semialgebraic') + using assms apply blast using g_def by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" + apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce + show ?thesis using 0 unfolding 1 by blast +qed + +lemma val_ineq_set_is_semialgebraic''': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" +proof- + obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" + using constant_poly_val_exists by blast + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" + apply(rule val_ineq_set_is_semialgebraic') + using g_def apply blast using assms by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" + apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce + show ?thesis using 0 unfolding 1 by blast +qed + +lemma val_eq_set_is_semialgebraic': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" +proof- + obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" + using constant_poly_val_exists by blast + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" + apply(rule val_eq_set_is_semialgebraic) + using assms apply blast using g_def by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" + apply(rule equalityI'') using g_def apply fastforce using g_def by metis + show ?thesis using 0 unfolding 1 by blast +qed + +lemma val_strict_ineq_set_is_semialgebraic': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" +proof- + obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" + using constant_poly_val_exists by blast + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" + apply(rule val_strict_ineq_set_is_semialgebraic) + using assms apply blast using g_def by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" + apply(rule equalityI'') using g_def apply fastforce using g_def + by fastforce + show ?thesis using 0 g_def unfolding 1 + by blast +qed + +lemma val_strict_ineq_set_is_semialgebraic'': + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" +proof- + obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" + using constant_poly_val_exists by blast + have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)}" + apply(rule val_strict_ineq_set_is_semialgebraic) + using g_def apply blast using assms by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" + apply(rule equalityI'') using assms g_def apply fastforce using assms g_def by fastforce + show ?thesis using 0 g_def unfolding 1 + by blast +qed + +lemma(in cring) R1_memE: + assumes "x \ carrier (R\<^bsup>1\<^esup>)" + shows "x = [(hd x)]" + using assms cartesian_power_car_memE + by (metis diff_is_0_eq' hd_conv_nth le_eq_less_or_eq length_0_conv length_tl list.exhaust list.sel(3) normalize.cases nth_Cons_0 zero_neq_one) + +lemma(in cring) R1_memE': + assumes "x \ carrier (R\<^bsup>1\<^esup>)" + shows "hd x \ carrier R" + using R1_memE assms cartesian_power_car_memE[of x R 1] cartesian_power_car_memE'[of x R 1 0] + by (metis hd_conv_nth less_numeral_extra(1) list.size(3) zero_neq_one) + +lemma univ_val_ineq_set_is_univ_semialgebraic: +"is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x \ c}" +proof- + have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c}" + apply(rule val_ineq_set_is_semialgebraic'') + using pvar_closed by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c} = to_R1 ` {x \ carrier Q\<^sub>p. val x \ c}" + proof(rule equalityI') + show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" + proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" + then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" + using Qp.R1_memE Qp.R1_memE' by blast + have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" + using A 0 + by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) + then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" + using A 0 unfolding mem_Collect_eq + by (metis (no_types, lifting) image_iff mem_Collect_eq) + qed + show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" + proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} " + then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a \ c" + by blast + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using cartesian_power_car_memI Qp.to_R1_closed by presburger + then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" + using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) + show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c" + unfolding 1 using a_def 0 by blast + qed + qed + show ?thesis using 0 unfolding 1 + using is_univ_semialgebraicI by blast +qed + +lemma univ_val_strict_ineq_set_is_univ_semialgebraic: +"is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x < c}" +proof- + have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) < c} = to_R1 ` {x \ carrier Q\<^sub>p. val x < c}" + proof(rule equalityI') + show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" + proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" + then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" + using Qp.R1_memE Qp.R1_memE' by blast + have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" + using A 0 + by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) + then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" + using A 0 unfolding mem_Collect_eq + by (metis (no_types, lifting) image_iff mem_Collect_eq) + qed + show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" + proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} " + then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a < c" + by blast + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using cartesian_power_car_memI Qp.to_R1_closed by presburger + then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" + using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) + show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c" + unfolding 1 using a_def 0 by blast + qed + qed + show ?thesis using 0 unfolding 1 + using is_univ_semialgebraicI by blast +qed + +lemma univ_val_eq_set_is_univ_semialgebraic: +"is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x = c}" +proof- + have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c}" + apply(rule val_eq_set_is_semialgebraic') + using pvar_closed by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c} = to_R1 ` {x \ carrier Q\<^sub>p. val x = c}" + proof(rule equalityI') + show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" + proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" + then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" + using Qp.R1_memE Qp.R1_memE' by blast + have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" + using A 0 + by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) + show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" + using A 0 unfolding mem_Collect_eq 1 by blast + qed + show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" + proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} " + then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a = c" + by blast + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using cartesian_power_car_memI Qp.to_R1_closed by presburger + then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" + using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) + show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c" + unfolding 1 using a_def 0 by blast + qed + qed + show ?thesis using 0 unfolding 1 + using is_univ_semialgebraicI by blast +qed + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\One Dimensional $p$-adic Balls are Semialgebraic\ + (********************************************************************************************) + (********************************************************************************************) + +lemma coord_ring_one_def: +"Pring Q\<^sub>p {(0::nat)} = (Q\<^sub>p[\\<^bsub>1\<^esub>])" +proof- + have "{(0::nat)} = {..<1}" + by auto + thus ?thesis + unfolding coord_ring_def + by auto +qed + +lemma times_p_pow_val: + assumes "a \ carrier Q\<^sub>p" + assumes "b = \[^]n \ a" + shows "val b = val a + n" + using val_mult[of "\[^]n" a] assms unfolding assms(2) val_p_int_pow + by (metis add.commute p_intpow_closed(1)) + +lemma times_p_pow_neg_val: + assumes "a \ carrier Q\<^sub>p" + assumes "b = \[^]-n \ a" + shows "val b = val a - n" + by (metis Qp.m_comm Qp_int_pow_nonzero assms(1) assms(2) p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) + +lemma eint_minus_int_pos: + assumes "a - eint n \ 0" + shows "a \ n" + using assms apply(induction a) + apply (metis diff_ge_0_iff_ge eint_ord_simps(1) idiff_eint_eint zero_eint_def) + by simp + +text\\p\-adic balls as pullbacks of polynomial maps\ + +lemma balls_as_pullbacks: + assumes "c \ carrier Q\<^sub>p" + shows "\P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]). to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p)" +proof- + obtain P0 where P0_def: "P0 = (to_poly (\[^](-n))) \\<^bsub>Q\<^sub>p_x\<^esub>((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c)" + by blast + have 0: "P0 \ carrier Q\<^sub>p_x" + proof- + have P0: "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c \ carrier Q\<^sub>p_x" + using UPQ.X_closed UPQ.to_poly_closed assms by blast + have P1: "(to_poly (\[^](-n))) \ carrier Q\<^sub>p_x" + using UPQ.to_poly_closed p_intpow_closed(1) by blast + then show ?thesis + using P0_def P0 P1 + by blast + qed + have 1: "\x. x \ carrier Q\<^sub>p \ P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" + proof- fix x assume A: "x \ carrier Q\<^sub>p" + have P0: "(to_poly (\[^](-n))) \ x = (\[^](-n))" + using A UPQ.to_fun_to_poly p_intpow_closed(1) by blast + have P1: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (x \\<^bsub>Q\<^sub>p\<^esub> c)" + by (metis A UPQ.to_fun_X_minus X_poly_minus_def assms) + have P2: "to_poly (\[^](-n)) \ carrier Q\<^sub>p_x" + using UPQ.to_poly_closed p_intpow_closed(1) by blast + have P3: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ carrier Q\<^sub>p_x" + using UPQ.X_closed UPQ.to_poly_closed assms by blast + have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = to_poly (\[^]- n) \ x \ (((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x)" + using A P0_def P0 P1 P2 P3 to_fun_mult[of "to_poly (\[^](-n))" "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c" x] UPQ.to_fun_mult + by blast + then have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c) " + by (metis P0 P1) + then show "P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" + using P0_def by metis + qed + have 2: " (\a. [a]) ` B\<^bsub>n\<^esub>[c] = poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" + proof + show "(\a. [a]) ` B\<^bsub>n\<^esub>[c] \ poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" + proof + fix x + assume A: "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" + then obtain a where a_def: "x = [a] \ a \ B\<^bsub>n\<^esub>[c]" + by blast + have P0: "P0 \ a \ \\<^sub>p" + proof- + have "B\<^bsub>n\<^esub>[c] \ carrier Q\<^sub>p" + using c_ball_in_Qp by blast + hence a_closed: "a \ carrier Q\<^sub>p" + using a_def by blast + have P0: "P0 \ a = (\[^](-n)) \ (a \ c)" + using 1 a_def c_ballE(1) + by blast + then have P1: "val (P0 \ a) = val (\[^](-n)) + val (a \ c)" + using val_mult[of "\[^]-n" "a \ c"] a_closed assms Qp.minus_closed p_intpow_closed(1) + by presburger + then have P2: "val (P0 \ a) = val (a \\<^bsub>Q\<^sub>p\<^esub> c) - n" + by (metis P0 Qp.m_comm Qp.minus_closed Qp_int_pow_nonzero assms local.a_closed + p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) + have P3: "val (a \\<^bsub>Q\<^sub>p\<^esub> c) \ n" + using a_def c_ballE(2) + by blast + then have "val (P0 \ a) \ -n + n" + using P2 by (metis add.commute diff_conv_add_uminus diff_self local.eint_minus_ineq' zero_eint_def) + then have P4: "val (P0 \ a) \ 0" + by (metis add.commute add.right_inverse zero_eint_def) + have P5: "P0 \ a \ carrier Q\<^sub>p" + using "0" UPQ.to_fun_closed local.a_closed by blast + then show ?thesis using P4 + using val_ring_val_criterion + by blast + qed + have "poly_map 1 [from_Qp_x P0] x = [Qp_ev (from_Qp_x P0) [a]] " + using a_def poly_map_def[of 1 "[from_Qp_x P0]"] poly_tuple_eval_def[of ] + by (metis Qp_poly_mapE' c_ballE(1) length_0_conv poly_map_cons Qp.to_R1_closed) + then have "poly_map 1 [from_Qp_x P0] x = [P0 \ a] " + using Qp_x_Qp_poly_eval[of P0 a] + by (metis "0" a_def c_ballE(1)) + then have P1: "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" + using P0 + by blast + have P2: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using A c_ballE(1) Qp.to_R1_closed + by blast + have P3: "is_poly_tuple 1 [from_Qp_x P0]" + apply(rule Qp_is_poly_tupleI) + by (metis "0" Qp_is_poly_tupleI from_Qp_x_closed gr_implies_not0 is_poly_tupleE is_poly_tuple_Cons list.size(3) zero_neq_one) + show "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p" + using P3 P2 P1 unfolding evimage_def poly_map_def + by blast + qed + have 20: "is_poly_tuple 1 [from_Qp_x P0]" + using 0 UP_to_IP_closed[of P0 "0::nat"] + unfolding is_poly_tuple_def + by (metis (no_types, lifting) empty_set from_Qp_x_closed list.simps(15) singletonD subset_code(1)) + show "poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" + proof fix x assume A: "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" + have A0: "(\a. [a]) ` \\<^sub>p \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def Qp.to_R1_car_subset + Qp_val_ring_is_semialgebraic is_semialgebraic_closed by presburger + have "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" + using A0 A 20 by blast + then obtain a where a_def: "a \ \\<^sub>p \ (poly_map 1 [from_Qp_x P0] x) = [a]" + by blast + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using A + by (meson evimage_eq) + then obtain y where y_def: "x = [y] \ y \ carrier Q\<^sub>p" + using A + by (metis Qp.to_R1_to_R Qp.to_R_pow_closed) + have "(poly_map 1 [from_Qp_x P0] x) = [(Qp_ev (from_Qp_x P0) [y])]" + unfolding poly_map_def poly_tuple_eval_def using x_closed + by (smt "20" One_nat_def length_Suc_conv list.size(3) nth_Cons_0 nth_map + poly_tuple_eval_closed poly_tuple_eval_def restrict_apply' Qp.to_R1_to_R y_def zero_less_Suc) + then have "(poly_map 1 [from_Qp_x P0] x) = [P0 \ y]" + by (metis "0" Qp_x_Qp_poly_eval y_def) + then have "[a] = [P0 \ y]" + using a_def + by presburger + then have A1: "a = (\[^](-n)) \ (y \\<^bsub>Q\<^sub>p\<^esub> c)" + using 1[of y] y_def + by blast + have "y \ B\<^bsub>n\<^esub>[c]" + proof- + have B0: "val a = val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n" + using A1 y_def Qp.minus_closed assms times_p_pow_neg_val by blast + have B1: "val a \ 0" + using a_def val_ring_memE by blast + then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n \ 0" + using B0 + by metis + then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) \ n" + using eint_minus_int_pos by blast + then show "y \ B\<^bsub>n\<^esub>[c]" + using c_ballI y_def by blast + qed + then show "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" + using y_def by blast + qed + qed + then show ?thesis + by (meson "0" from_Qp_x_closed) +qed + +lemma ball_is_semialgebraic: + assumes "c \ carrier Q\<^sub>p" + shows "is_semialgebraic 1 (to_R1` B\<^bsub>n\<^esub>[c])" +proof- + obtain P where P_def: "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]) \ to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p) " + using assms balls_as_pullbacks[of c n] by meson + have "is_poly_tuple 1 [P]" + using P_def unfolding is_poly_tuple_def + by (metis (no_types, opaque_lifting) list.inject list.set_cases neq_Nil_conv subset_code(1)) + then show ?thesis + using assms P_def pullback_is_semialg[of 1 "[P]" 1 "((\a. [a]) ` \\<^sub>p) "] + by (metis (mono_tags, lifting) One_nat_def + Qp_val_ring_is_semialgebraic is_semialgebraic_def length_Suc_conv + list.distinct(1) list.size(3)) +qed + +lemma ball_is_univ_semialgebraic: + assumes "c \ carrier Q\<^sub>p" + shows "is_univ_semialgebraic (B\<^bsub>n\<^esub>[c])" + using assms ball_is_semialgebraic c_ball_in_Qp is_univ_semialgebraic_def + by presburger + +abbreviation Qp_to_R1_set where +"Qp_to_R1_set S \ to_R1 ` S" + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Finite Unions and Intersections of Semialgebraic Sets\ + (********************************************************************************************) + (********************************************************************************************) + + +definition are_semialgebraic where +"are_semialgebraic n Xs = (\ x. x \ Xs \ is_semialgebraic n x)" + +lemma are_semialgebraicI: + assumes "\x. x \ Xs \ is_semialgebraic n x " + shows "are_semialgebraic n Xs" + using are_semialgebraic_def assms by blast + +lemma are_semialgebraicE: + assumes "are_semialgebraic n Xs" + assumes "x \ Xs" + shows "is_semialgebraic n x" + using are_semialgebraic_def assms(1) assms(2) by blast + +definition are_univ_semialgebraic where +"are_univ_semialgebraic Xs = (\ x. x \ Xs \ is_univ_semialgebraic x)" + +lemma are_univ_semialgebraicI: + assumes "\x. x \ Xs \ is_univ_semialgebraic x " + shows "are_univ_semialgebraic Xs" + using are_univ_semialgebraic_def assms by blast + +lemma are_univ_semialgebraicE: + assumes "are_univ_semialgebraic Xs" + assumes "x \ Xs" + shows "is_univ_semialgebraic x" + using are_univ_semialgebraic_def assms(1) assms(2) by blast + +lemma are_univ_semialgebraic_semialgebraic: + assumes "are_univ_semialgebraic Xs" + shows "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" + apply(rule are_semialgebraicI) + using are_univ_semialgebraicE assms image_iff is_univ_semialgebraicE + by (metis (no_types, lifting)) + +lemma to_R1_set_union: +"to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" + using image_iff by blast + +lemma to_R1_inter: + assumes "Xs \ {}" + shows "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" +proof + show "to_R1 ` (\ Xs) \ \ (Qp_to_R1_set ` Xs)" + by blast + show "\ (Qp_to_R1_set ` Xs) \ to_R1 ` (\ Xs)" + proof fix x + assume A: "x \ \ (Qp_to_R1_set ` Xs)" + then have 0: "\S. S \ Xs \ x \ (Qp_to_R1_set S)" + by blast + obtain S where "S \ Xs \ x \ (Qp_to_R1_set S)" + using assms 0 + by blast + then obtain b where b_def: "b \ S \ x = [b]" + by blast + have "b \ (\ Xs)" + using "0" b_def by blast + then show "x \ to_R1 ` (\ Xs)" + using b_def by blast + qed +qed + +lemma finite_union_is_semialgebraic: + assumes "finite Xs" + shows "Xs \ semialg_sets n \ is_semialgebraic n (\ Xs)" + apply(rule finite.induct[of Xs]) + apply (simp add: assms) + apply (simp add: empty_is_semialgebraic) + by (metis Sup_insert insert_subset is_semialgebraicI union_is_semialgebraic) + +lemma finite_union_is_semialgebraic': + assumes "finite Xs" + assumes "Xs \ semialg_sets n " + shows "is_semialgebraic n (\ Xs)" + using assms(1) assms(2) finite_union_is_semialgebraic by blast + +lemma(in padic_fields) finite_union_is_semialgebraic'': + assumes "finite S" + assumes "\x. x \ S \ is_semialgebraic m (F x)" + shows "is_semialgebraic m (\ x \ S. F x)" + using assms finite_union_is_semialgebraic[of "F`S" m] unfolding is_semialgebraic_def + by blast + +lemma finite_union_is_univ_semialgebraic': + assumes "finite Xs" + assumes "are_univ_semialgebraic Xs" + shows "is_univ_semialgebraic (\ Xs)" +proof- + have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" + using assms finite_union_is_semialgebraic'[of "((`) (\a. [a]) ` Xs)"] to_R1_set_union[of Xs] + by (metis (no_types, lifting) are_semialgebraicE are_univ_semialgebraic_semialgebraic + finite_imageI is_semialgebraicE subsetI) + then show ?thesis + using is_univ_semialgebraicI by blast +qed + +lemma finite_intersection_is_semialgebraic: + assumes "finite Xs" + shows "Xs \ semialg_sets n \ Xs \{} \ is_semialgebraic n (\ Xs)" + apply(rule finite.induct[of Xs]) + apply (simp add: assms) + apply auto[1] +proof fix A::"((nat \ int) \ (nat \ int)) set list set set" fix a + assume 0: "finite A" + assume 1: "A \ semialg_sets n \ A \ {} \ is_semialgebraic n (\ A) " + assume 2: "insert a A \ semialg_sets n \ insert a A \ {}" + show "is_semialgebraic n (\ (insert a A))" + proof(cases "A = {}") + case True + then have "insert a A = {a}" + by simp + then show ?thesis + by (metis "2" cInf_singleton insert_subset is_semialgebraicI) + next + case False + then have "A \ semialg_sets n \ A \ {}" + using "2" by blast + then have "is_semialgebraic n (\ A) " + using "1" by linarith + then show ?thesis + using 0 1 2 intersection_is_semialg + by (metis Inf_insert insert_subset is_semialgebraicI) + qed +qed + +lemma finite_intersection_is_semialgebraic': + assumes "finite Xs" + assumes "Xs \ semialg_sets n \ Xs \{}" + shows " is_semialgebraic n (\ Xs)" + by (simp add: assms(1) assms(2) finite_intersection_is_semialgebraic) + +lemma finite_intersection_is_semialgebraic'': + assumes "finite Xs" + assumes "are_semialgebraic n Xs \ Xs \{}" + shows " is_semialgebraic n (\ Xs)" + by (meson are_semialgebraicE assms(1) assms(2) + finite_intersection_is_semialgebraic' is_semialgebraicE subsetI) + +lemma finite_intersection_is_univ_semialgebraic: + assumes "finite Xs" + assumes "are_univ_semialgebraic Xs" + assumes "Xs \ {}" + shows "is_univ_semialgebraic (\ Xs)" +proof- + have "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" + using are_univ_semialgebraic_semialgebraic assms(2) by blast + then have "is_semialgebraic 1 (\ (Qp_to_R1_set ` Xs))" + using assms finite_intersection_is_semialgebraic''[of "Qp_to_R1_set ` Xs" 1] + by blast + then have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" + using assms to_R1_inter[of Xs] + by presburger + then show ?thesis + using is_univ_semialgebraicI by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Cartesian Products of Semialgebraic Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma Qp_times_basic_semialg_right: + assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" +proof + show "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ basic_semialg_set (n + m) k a" + proof fix x + assume "x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" + then obtain as bs where as_bs_def: "as \ (basic_semialg_set n k a) \ bs \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ x = as@bs" + using cartesian_product_memE[of x "basic_semialg_set n k a" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p n] + basic_semialg_set_def + by (metis (no_types, lifting) append_take_drop_id basic_semialg_set_memE(1) subsetI) + have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" + using as_bs_def basic_semialg_set_memE(1) cartesian_product_closed' + by blast + have 1: "(Qp_ev a x = Qp_ev a as)" + using as_bs_def poly_eval_cartesian_prod[of as n bs m a] assms basic_semialg_set_memE(1) by blast + obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev a as = (y[^]k))" + using as_bs_def using basic_semialg_set_memE(2)[of as n k a] by blast + show "x \ basic_semialg_set (n + m) k a" + apply(rule basic_semialg_set_memI[of _ _ y]) + apply (simp add: "0") + apply (simp add: y_def) + using "1" y_def by blast + qed + show "basic_semialg_set (n + m) k a \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" + proof fix x + assume A: "x \ basic_semialg_set (n + m) k a" + have A0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" + using A basic_semialg_set_memE(1) by blast + have A1: "set x \ carrier Q\<^sub>p" + using A0 + by (metis (no_types, lifting) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth subsetI) + have A2: "length x = n + m" + using A0 cartesian_power_car_memE + by blast + obtain as where as_def: "as = take n x" + by blast + obtain bs where bs_def: "bs = drop n x" + by blast + have 0: "x = as@bs" + using A as_def bs_def + by (metis append_take_drop_id) + have 1: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + using as_def A2 + apply (simp add: A2 min.absorb2) + by (metis (no_types, lifting) A1 as_def dual_order.trans set_take_subset) + have 2: "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + apply(rule cartesian_power_car_memI) + using bs_def A2 + apply (simp add: A2) + by (metis A1 bs_def order.trans set_drop_subset) + obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)" + using basic_semialg_set_memE A by meson + have 3: "as \ basic_semialg_set n k a" + apply(rule basic_semialg_set_memI[of _ _ y]) + apply (simp add: "1") + using \y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)\ apply blast + using y_def A 1 0 2 assms(1) poly_eval_cartesian_prod by blast + show " x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" + using 3 2 "0" + by (metis (mono_tags, lifting) as_def basic_semialg_set_memE(1) bs_def cartesian_product_memI subsetI) + qed +qed + +lemma Qp_times_basic_semialg_right_is_semialgebraic: + assumes "k > 0" + assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" +proof- + have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" + using Qp_times_basic_semialg_right assms + by presburger + have 1: " a \ carrier (Q\<^sub>p[\\<^bsub>n+m\<^esub>])" + using assms poly_ring_car_mono'(2) by blast + have 2: "is_semialgebraic (n + m) (basic_semialg_set (n + m) k a)" + using assms basic_semialg_is_semialgebraic'[of a "n+m" k "basic_semialg_set (n + m) k a"] + "1" by blast + show ?thesis + using 0 2 + by metis +qed + +lemma Qp_times_basic_semialg_right_is_semialgebraic': + assumes "A \ basic_semialgs n" + shows "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" +proof- + obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" + using assms is_basic_semialg_def + by (metis mem_Collect_eq) + then show ?thesis using + Qp_times_basic_semialg_right_is_semialgebraic[of k P] + using assms(1) by blast +qed + +lemma cartesian_product_memE': + assumes "x \ cartesian_product A B" + obtains a b where "a \ A \ b \ B \ x = a@b" + using assms unfolding cartesian_product_def by blast + +lemma Qp_times_basic_semialg_left: + assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) = basic_semialg_set (n+m) k (shift_vars n m a)" +proof + show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) \ basic_semialg_set (n + m) k (shift_vars n m a)" + proof fix x + assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" + then obtain as bs where as_bs_def: "as \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ bs \ (basic_semialg_set n k a) \ x = as@bs " + using cartesian_product_memE' by blast + have 0: "Qp_ev (shift_vars n m a) x = Qp_ev a bs" + using A as_bs_def assms shift_vars_eval[of a n as m bs ] + by (metis (no_types, lifting) basic_semialg_set_memE(1)) + obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a bs = (y[^]k)" + using as_bs_def basic_semialg_set_memE(2) + by blast + have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" + using A as_bs_def + by (metis (no_types, lifting) add.commute basic_semialg_set_memE(1) cartesian_product_closed') + show "x \ basic_semialg_set (n + m) k (shift_vars n m a)" + apply(rule basic_semialg_set_memI[of _ _ y]) + apply (simp add: "1") + using y_def apply blast + using "0" y_def by blast + qed + show "basic_semialg_set (n + m) k (shift_vars n m a) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) " + proof fix x + assume A: "x \ basic_semialg_set (n + m) k (shift_vars n m a)" + then obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev (shift_vars n m a) x = (y[^]k)" + using assms basic_semialg_set_memE[of x "n + m" k "shift_vars n m a"] + shift_vars_closed[of a m] Qp.cring_axioms + by blast + have "x \ carrier (Q\<^sub>p\<^bsup>m+n\<^esup>)" + using A basic_semialg_set_memE(1) + by (metis add.commute) + then have "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + using cartesian_product_carrier by blast + then obtain as bs where as_bs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (meson cartesian_product_memE') + have "bs \ (basic_semialg_set n k a)" + apply(rule basic_semialg_set_memI[of _ _ y]) + using as_bs_def apply blast + apply (simp add: y_def) + using y_def shift_vars_eval[of a n as m bs ] as_bs_def assms(1) + by metis + then show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" + using as_bs_def unfolding cartesian_product_def + by blast + qed +qed + +lemma Qp_times_basic_semialg_left_is_semialgebraic: + assumes "k > 0" + assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a))" + using basic_semialg_is_semialgebraic'[of a "n+m" k] Qp_times_basic_semialg_left + by (metis assms(1) assms(2) basic_semialg_is_semialgebraic is_basic_semialg_def neq0_conv shift_vars_closed) + +lemma Qp_times_basic_semialg_left_is_semialgebraic': + assumes "A \ basic_semialgs n" + shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) A)" +proof- + obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" + using assms is_basic_semialg_def + by (metis mem_Collect_eq) + then show ?thesis using + Qp_times_basic_semialg_left_is_semialgebraic[of k P] + using assms(1) by blast +qed + +lemma product_of_basic_semialgs_is_semialg: + assumes "k > 0" + assumes "l > 0" + assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b))" +proof- + have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" + using Qp_times_basic_semialg_right assms by presburger + have 1: "cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b) = basic_semialg_set (m + n) l (shift_vars m n b)" + using Qp_times_basic_semialg_left assms by blast + have 2: "(cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b)) = + cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ + cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b)" + proof- + have 0: "basic_semialg_set n k a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using basic_semialg_set_memE(1) by blast + have 1: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + by simp + have 2: "carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by simp + have 3: "basic_semialg_set m l b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using basic_semialg_set_memE(1) by blast + show ?thesis + using 0 1 2 3 cartesian_product_intersection[of "(basic_semialg_set n k a)" Q\<^sub>p n + "(carrier (Q\<^sub>p\<^bsup>m\<^esup>))" m + "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))" "(basic_semialg_set m l b)"] + by (smt Collect_cong inf.absorb_iff1 inf.absorb_iff2) + qed + then show ?thesis + using Qp_times_basic_semialg_left_is_semialgebraic + Qp_times_basic_semialg_right_is_semialgebraic assms + by (metis (no_types, lifting) add.commute intersection_is_semialg) +qed + +lemma product_of_basic_semialgs_is_semialg': + assumes "A \ (basic_semialgs n)" + assumes "B \ (basic_semialgs m)" + shows "is_semialgebraic (n + m) (cartesian_product A B)" +proof- + obtain k a where ka_def: "k > 0 \ a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ A = (basic_semialg_set n k a)" + using assms + by (metis is_basic_semialg_def mem_Collect_eq neq0_conv) + obtain l b where lb_def: "l > 0 \ b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]) \ B = (basic_semialg_set m l b)" + by (metis assms(2) gr_zeroI is_basic_semialg_def mem_Collect_eq) + show ?thesis using ka_def lb_def assms product_of_basic_semialgs_is_semialg + by blast +qed + +lemma car_times_semialg_is_semialg: + assumes "is_semialgebraic m B" + shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) B)" + apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) + using assms is_semialgebraic_def semialg_sets_def apply blast + apply (simp add: carrier_is_semialgebraic cartesian_product_carrier) +proof- + show "\A. A \ basic_semialgs m \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" + proof- + fix A assume A: "A \ basic_semialgs m " + then have " (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = A" + by (metis basic_semialg_set_memE(1) inf_absorb1 is_basic_semialg_def mem_Collect_eq subsetI) + then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" + using add.commute[of n m] assms A + Qp_times_basic_semialg_left_is_semialgebraic' + by (simp add: \n + m = m + n\) + qed + show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ + is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ + C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ + is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C) \ + is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" + proof- fix A C assume A: " A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" + "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" + "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" + " is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" + then have B: "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" + using union_is_semialgebraic by blast + show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" + proof- + have 0: "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A(1) gen_boolean_algebra_subset + by blast + have 1: " C \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A(3) gen_boolean_algebra_subset + by blast + then show ?thesis using 0 A B + using cartesian_product_binary_union_right[of A Q\<^sub>p m C "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))"] + unfolding is_semialgebraic_def semialg_sets_def + by presburger + qed + qed + show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ + is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ + is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" + proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" + "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" + then have "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using gen_boolean_algebra_subset + by blast + then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" + using A cartesian_product_car_complement_right[of A Q\<^sub>p m n] + unfolding is_semialgebraic_def semialg_sets_def + by (metis (mono_tags, lifting) semialg_complement semialg_sets_def) + qed +qed + +lemma basic_semialg_times_semialg_is_semialg: + assumes "A \ basic_semialgs n" + assumes "is_semialgebraic m B" + shows " is_semialgebraic (n + m) (cartesian_product A B)" + apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) + using assms(2) is_semialgebraic_def semialg_sets_def apply blast + using Qp_times_basic_semialg_right_is_semialgebraic' assms(1) apply blast + apply (metis assms(1) basic_semialg_is_semialgebraic inf.absorb1 is_semialgebraic_closed mem_Collect_eq product_of_basic_semialgs_is_semialg') + apply (metis (no_types, lifting) cartesian_product_binary_union_right is_semialgebraicI is_semialgebraic_closed semialg_sets_def union_is_semialgebraic) +proof- + show "\Aa. Aa \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ + is_semialgebraic (n + m) (cartesian_product A Aa) \ is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - Aa))" + proof- fix B assume A: "B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" + "is_semialgebraic (n + m) (cartesian_product A B)" + show "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B))" + using A assms cartesian_product_complement_right[of B Q\<^sub>p m A n] add.commute[of n m] + proof - + have f1: "\n B. \ is_semialgebraic n B \ B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (meson is_semialgebraic_closed) + have "is_basic_semialg n A" + using \A \ {S. is_basic_semialg n S}\ by blast + then have f2: "is_semialgebraic n A" + using padic_fields.basic_semialg_is_semialgebraic padic_fields_axioms by blast + have "B \ semialg_sets m" + using \B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {S. is_basic_semialg m S}\ semialg_sets_def by blast + then have "is_semialgebraic m B" + by (meson padic_fields.is_semialgebraicI padic_fields_axioms) + then show ?thesis + using f2 f1 by (metis (no_types) Qp_times_basic_semialg_right_is_semialgebraic' \A \ {S. is_basic_semialg n S}\ \\B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>); A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \ cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) - cartesian_product A B = cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B)\ \is_semialgebraic (n + m) (cartesian_product A B)\ diff_is_semialgebraic) + qed + qed +qed + +text\Semialgebraic sets are closed under cartesian products\ + +lemma cartesian_product_is_semialgebraic: + assumes "is_semialgebraic n A" + assumes "is_semialgebraic m B" + shows "is_semialgebraic (n + m) (cartesian_product A B)" + apply(rule gen_boolean_algebra.induct[of A "carrier (Q\<^sub>p\<^bsup>n\<^esup>)""basic_semialgs n"]) + using assms is_semialgebraicE semialg_sets_def apply blast + using assms car_times_semialg_is_semialg apply blast + using assms basic_semialg_times_semialg_is_semialg + apply (simp add: Int_absorb2 basic_semialg_is_semialgebraic is_semialgebraic_closed) +proof- + show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ + is_semialgebraic (n + m) (cartesian_product A B) \ + C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ + is_semialgebraic (n + m) (cartesian_product C B) \ is_semialgebraic (n + m) (cartesian_product (A \ C) B)" + proof- fix A C assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" + "is_semialgebraic (n + m) (cartesian_product A B)" + "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" + "is_semialgebraic (n + m) (cartesian_product C B)" + show "is_semialgebraic (n + m) (cartesian_product (A \ C) B)" + using A cartesian_product_binary_union_left[of A Q\<^sub>p n C B] + by (metis (no_types, lifting) gen_boolean_algebra_subset union_is_semialgebraic) + qed + show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ + is_semialgebraic (n + m) (cartesian_product A B) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" + proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" + "is_semialgebraic (n + m) (cartesian_product A B)" + show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" + using assms A cartesian_product_complement_left[of A Q\<^sub>p n B m] + unfolding is_semialgebraic_def semialg_sets_def + by (metis car_times_semialg_is_semialg diff_is_semialgebraic is_semialgebraicE is_semialgebraicI + is_semialgebraic_closed semialg_sets_def) + qed +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\$N^{th}$ Power Residues\ +(**************************************************************************************************) +(**************************************************************************************************) + + +definition nth_root_poly where +"nth_root_poly (n::nat) a = ((X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> n) \\<^bsub>Q\<^sub>p_x\<^esub> (to_poly a)" + +lemma nth_root_poly_closed: + assumes "a \ carrier Q\<^sub>p" + shows "nth_root_poly n a \ carrier Q\<^sub>p_x" + using assms unfolding nth_root_poly_def + by (meson UPQ.P.minus_closed UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_poly_closed) + +lemma nth_root_poly_eval: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + shows "(nth_root_poly n a) \ b = (b[^]n) \\<^bsub>Q\<^sub>p\<^esub> a" + using assms unfolding nth_root_poly_def + using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_X_pow UPQ.to_fun_diff UPQ.to_fun_to_poly UPQ.to_poly_closed by presburger + +text\Hensel's lemma gives us this criterion for the existence of \n\-th roots\ + +lemma nth_root_poly_root: + assumes "(n::nat) > 1" + assumes "a \ \\<^sub>p" + assumes "a \ \" + assumes "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > 2* val ([n]\\)" + shows "(\ b \ \\<^sub>p. ((b[^]n) = a))" +proof- + obtain \ where alpha_def: "\ \ carrier Z\<^sub>p \ \ \ = a" + using assms(2) by blast + have 0: "\ \ carrier Z\<^sub>p" + by (simp add: alpha_def) + have 1: "\ \ \\<^bsub>Z\<^sub>p\<^esub>" + using assms alpha_def inc_of_one + by blast + obtain N where N_def: "N = [n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + by blast + have N_closed: "N \ carrier Z\<^sub>p" + using N_def Zp_nat_mult_closed + by blast + have 2: "\ ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = ([n]\ \)" + proof(induction n) + case 0 + have 00: "[(0::nat)] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = \\<^bsub>Z\<^sub>p\<^esub>" + using Zp_nat_inc_zero by blast + have 01: "[(0::nat)] \\<^bsub>Q\<^sub>p\<^esub> \ = \" + using Qp.nat_inc_zero by blast + then show ?case + using 00 inc_of_nat by blast + next + case (Suc n) + then show ?case + using inc_of_nat by blast + qed + have 3: "val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \) = val (\ \\<^bsub>Q\<^sub>p\<^esub> a)" + using alpha_def Zp.one_closed ring_hom_one[of \ Z\<^sub>p Q\<^sub>p] inc_is_hom Zp.ring_hom_minus[of Q\<^sub>p \ "\\<^bsub>Z\<^sub>p\<^esub>" \ ] + Qp.ring_axioms + unfolding \_def + by (metis Q\<^sub>p_def Zp.minus_closed Zp_def padic_fields.val_of_inc padic_fields_axioms) + have 4: "([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) \ nonzero Z\<^sub>p" + proof- + have 40: "int n \ 0" + using of_nat_0_le_iff by blast + have "nat (int n) = n" + using nat_int by blast + hence "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = [int n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + using 40 unfolding add_pow_def int_pow_def nat_pow_def + proof - + have "(if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) 0 else rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n) = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n" + by (meson of_nat_less_0_iff) + then show "rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n = (let f = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) in if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> f (nat (- int n)) else f (nat (int n)))" + using \nat (int n) = n\ by presburger + qed + thus ?thesis + using Zp_char_0[of n] Zp.not_nonzero_memE Zp_char_0' assms(1) gr_implies_not_zero by blast + qed + then have 5: "val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = val ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))" + using 2 ord_of_inc + by (metis N_closed N_def val_of_inc) + then have 6: "(val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \)) > 2*(val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>))" + using assms 3 by presburger + have "\ b \ carrier Z\<^sub>p. (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" + using Zp_nth_root_lemma[of \ n] assms "0" "1" "6" by blast + then obtain b where b_def: "b \ carrier Z\<^sub>p \ (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" + by blast + then have "\ (b [^]\<^bsub>Z\<^sub>p\<^esub>n) = a" + using alpha_def by blast + then have "(\ b) [^] n = a" + by (metis Qp.nat_inc_zero Q\<^sub>p_def Qp.nat_pow_zero Zp.nat_pow_0 Zp.nat_pow_zero + Zp_nat_inc_zero \_def alpha_def assms(3) b_def frac_inc_of_nat inc_of_one inc_pow not_nonzero_Qp) + then show ?thesis + using b_def by blast +qed + +text\All points sufficiently close to 1 have nth roots\ + +lemma eint_nat_times_2: +"2*(n::nat) = 2*eint n" + using times_eint_simps(1) + by (metis mult.commute mult_2_right of_nat_add) + +lemma P_set_of_one: +"P_set 1 = nonzero Q\<^sub>p" + apply(rule equalityI) apply(rule subsetI) + unfolding P_set_def nonzero_def mem_Collect_eq apply blast + apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq + using Qp.nat_pow_eone by blast + +lemma nth_power_fact: + assumes "(n::nat) \ 1" + shows "\ (m::nat) > 0. \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" +proof(cases "n = 1") + case True + have "\ u \ carrier Q\<^sub>p. ac 1 u = 1 \ val u = 0 \ u \ P_set n" + unfolding True P_set_of_one + by (metis iless_Suc_eq padic_fields.val_ring_memE padic_fields.zero_in_val_ring padic_fields_axioms val_nonzero zero_eint_def) + then show ?thesis by blast +next + case F: False + obtain m where m_def: "m = 1 + nat ( 2*(ord ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))))" + by blast + then have m_pos: "m > 0" + by linarith + have "\ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" + proof + fix u + assume A: "u \ carrier Q\<^sub>p" + show " ac m u = 1 \ val u = 0 \ u \ P_set n" + proof + assume B: "ac m u = 1 \ val u = 0" + then have 0: "val u = val \" + by (smt A ac_def not_nonzero_Qp val_one val_ord zero_eint_def) + have 1: "ac m u = ac m \" + by (metis B Qp.one_nonzero ac_p ac_p_int_pow_factor angular_component_factors_x angular_component_p inc_of_one m_pos p_nonzero) + have 2: "u \ nonzero Q\<^sub>p" + proof- + have "ac m \ = 0" + by (meson ac_def) + then have "u \ \" + by (metis B zero_neq_one) + then show ?thesis + using A not_nonzero_Qp Qp.nonzero_memI by presburger + qed + then have 3: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using m_pos 0 1 ac_ord_prop[of "\" u "0::int" m] + by (metis B Qp.one_nonzero add.right_neutral eint.inject val_ord zero_eint_def) + show "u \ P_set n" + proof(cases "u = \") + case True + then show ?thesis + by (metis P_set_one insert_iff zeroth_P_set) + next + case False + have F0: "u \ \\<^sub>p" + apply(rule val_ring_memI, rule A) + unfolding 0 val_one by auto + have F1: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" + using False 3 by blast + have "ord (\ \ u) \ m" + by (metis A F1 False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) val_ord) + hence F2: "ord (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" + using m_def F1 A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) + int_nat_eq of_nat_1 of_nat_add val_ord[of "\ \ u"] eint_nat_times_2 + by linarith + have "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(val ([n]\ \))" + proof- + have 0: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" + using F2 val_ord[of "\ \ u"] A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(2) by presburger + have "n > 0" + using assms by linarith + hence "eint (ord ([n] \ \)) = val ([n] \ \)" + using val_ord_nat_inc[of n] + by blast + hence "2*ord ([n]\ \) = 2*val ([n]\ \)" + by (metis inc_of_nat times_eint_simps(1)) + thus ?thesis + using 0 val_ord[of "\ \ u"] assms + by presburger + qed + then have "(\ b \ \\<^sub>p. ((b[^]n) = u))" + using m_def False nth_root_poly_root[of n u] F0 assms F by linarith + then have "(\ b \ carrier Q\<^sub>p. ((b[^]n) = u))" + using val_ring_memE by blast + then show "u \ P_set n" + using P_set_def[of n] 2 + by blast + qed + qed + qed + then show ?thesis using m_pos by blast +qed + +definition pow_res where +"pow_res (n::nat) x = {a. a \ carrier Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. (a = x \ (y[^]n)))}" + +lemma nonzero_pow_res: + assumes "x \ nonzero Q\<^sub>p" + shows "pow_res (n::nat) x \ nonzero Q\<^sub>p" +proof + fix a + assume "a \ pow_res n x" + then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (a = x \ (y[^]n))" + using pow_res_def by blast + then show "a \ nonzero Q\<^sub>p" + using assms Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero by blast +qed + +lemma pow_res_of_zero: + shows "pow_res n \ = {\}" + unfolding pow_res_def apply(rule equalityI) + apply(rule subsetI) + unfolding mem_Collect_eq + apply (metis Qp.integral_iff Qp.nat_pow_closed Qp.nonzero_closed Qp.zero_closed insertCI) + apply(rule subsetI) unfolding mem_Collect_eq + by (metis Qp.nat_pow_one Qp.one_nonzero Qp.r_one Qp.zero_closed equals0D insertE) + +lemma equal_pow_resI: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ pow_res n x" + shows "pow_res n x = pow_res n y" +proof + have y_closed: "y \ carrier Q\<^sub>p" + using assms unfolding pow_res_def by blast + obtain c where c_def: "c \ nonzero Q\<^sub>p \ y = x \ (c[^]n)" + using assms pow_res_def by blast + have "((inv c)[^]n) = inv (c[^]n)" + using c_def Qp.field_axioms Qp.nat_pow_of_inv Units_eq_nonzero by blast + then have "y \ ((inv c)[^]n) = x" + using y_closed c_def assms Qp.inv_cancelL(2) Qp.nonzero_closed Qp_nat_pow_nonzero Units_eq_nonzero + by presburger + then have P0: "(inv c) \ nonzero Q\<^sub>p \ x =y \ ((inv c)[^]n) " + using c_def nonzero_inverse_Qp by blast + show "pow_res n x \ pow_res n y" + proof + fix a + assume A: "a \ pow_res n x" + then have "a \ carrier Q\<^sub>p" + by (metis (no_types, lifting) mem_Collect_eq pow_res_def) + obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = x \ (b[^]n)" + using A pow_res_def by blast + then have 0: "b \ nonzero Q\<^sub>p \ a = y \ ((inv c)[^]n) \ (b[^]n)" + using \y \ inv c [^] n = x\ by blast + have "b \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" + proof- + have "(inv c)[^]n \ (b[^]n) = ((inv c) \ b)[^]n" + using c_def b_def assms P0 Qp.nat_pow_distrib Qp.nonzero_closed by presburger + then have " y \ (((inv c)[^]n) \ (b[^]n)) = y \ (((inv c) \ b)[^]n)" + by presburger + then show ?thesis + using y_closed 0 P0 Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) by presburger + qed + then have "((inv c) \ b) \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" + by (metis P0 Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed not_nonzero_Qp) + then show "a \ pow_res n y" using pow_res_def \a \ carrier Q\<^sub>p\ by blast + qed + show "pow_res n y \ pow_res n x" + proof + fix a + assume A: "a \ pow_res n y" + then have 0: "a \ carrier Q\<^sub>p" + by (metis (no_types, lifting) mem_Collect_eq pow_res_def) + obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = y \ (b[^]n)" + using A pow_res_def by blast + then have "a = (x \ (c[^]n)) \ (b[^]n)" + using c_def by blast + then have "a = x \ ((c[^]n) \ (b[^]n))" + by (meson Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) b_def c_def) + then have "a = x \ ((c \ b)[^]n)" + using Qp.nat_pow_distrib Qp.nonzero_closed b_def c_def by presburger + then have "(c \ b) \ nonzero Q\<^sub>p \ a = x \ ((c \ b)[^]n)" + by (metis Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed b_def c_def not_nonzero_Qp) + then show "a \ pow_res n x" + using pow_res_def 0 by blast + qed +qed + +lemma zeroth_pow_res: + assumes "x \ carrier Q\<^sub>p" + shows "pow_res 0 x = {x}" + apply(rule equalityI) + apply(rule subsetI) + unfolding pow_res_def mem_Collect_eq + using assms apply (metis Qp.nat_pow_0 Qp.r_one singletonI) + apply(rule subsetI) + unfolding pow_res_def mem_Collect_eq + using assms by (metis Qp.nat_pow_0 Qp.one_nonzero Qp.r_one equals0D insertE) + +lemma Zp_car_zero_res: assumes "x \ carrier Z\<^sub>p" + shows "x 0 = 0" + using assms unfolding Zp_def + using Zp_def Zp_defs(3) padic_set_zero_res prime by blast + +lemma zeroth_ac: + assumes "x \ carrier Q\<^sub>p" + shows "ac 0 x = 0" + apply(cases "x = \ ") + unfolding ac_def apply presburger + using assms angular_component_closed[of x] Zp_car_zero_res unfolding nonzero_def mem_Collect_eq + by presburger + +lemma nonzero_ac_imp_nonzero: + assumes "x \ carrier Q\<^sub>p" + assumes "ac m x \ 0" + shows "x \ nonzero Q\<^sub>p" + using assms unfolding ac_def nonzero_def mem_Collect_eq + by presburger + +lemma nonzero_ac_val_ord: + assumes "x \ carrier Q\<^sub>p" + assumes "ac m x \ 0" + shows "val x = ord x" + using nonzero_ac_imp_nonzero assms val_ord by blast + +lemma pow_res_equal_ord: + assumes "n > 0" + shows "\m > 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" +proof- + obtain m where m_def_0: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n)" + using assms nth_power_fact[of n] + by (metis less_imp_le_nat less_one linorder_neqE_nat nat_le_linear zero_less_iff_neq_zero) + then have m_def: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ ord u = 0 \ u \ P_set n)" + by (smt nonzero_ac_val_ord zero_eint_def) + have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" + proof + fix x + show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" + proof fix y + show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" + proof + assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y " + then have 0: "ac m (x \
y) = 1" + using ac_inv[of y m] ac_mult + by (metis ac_inv'''(1) ac_mult' m_def nonzero_inverse_Qp) + have 1: "ord (x \
y) = 0" + using A ord_fract by presburger + have 2: "(x \
y) \ nonzero Q\<^sub>p" + using A + by (metis Qp.nonzero_closed Qp.nonzero_mult_closed local.fract_cancel_right nonzero_inverse_Qp not_nonzero_Qp zero_fract) + have 3: "(x \
y) \ P_set n" + using m_def 0 1 2 nonzero_def + by (smt Qp.nonzero_closed) + then obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^]n) = (x \
y)" + using P_set_def + by blast + then have "(x \
y) \ y = (b[^]n) \ y" + by presburger + then have "x = (b[^]n) \ y" + using A b_def + by (metis Qp.nonzero_closed local.fract_cancel_left) + then have "x = y \(b[^]n)" + using A b_def + by (metis Qp.nonzero_closed local.fract_cancel_right) + then have "x \ pow_res n y" + unfolding pow_res_def using A b_def + by (metis (mono_tags, lifting) "2" Qp.nat_pow_0 Qp.nonzero_closed Qp_nonzero_nat_pow mem_Collect_eq not_gr_zero) + then show "pow_res n x = pow_res n y" + using A equal_pow_resI[of x y n] unfolding nonzero_def + by (metis (mono_tags, lifting) A Qp.nonzero_closed equal_pow_resI) + qed + qed + qed + then show ?thesis using m_def by blast +qed + +lemma pow_res_equal: + assumes "n > 0" + shows "\m> 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = (ord y mod n) \ pow_res n x = pow_res n y" +proof- + obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" + using assms pow_res_equal_ord + by meson + have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" + proof fix x + show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" + proof fix y + show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" + proof + assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n" + show "pow_res n x = pow_res n y" + proof- + have A0: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p" + using A by blast + have A1: "ac m x = ac m y" + using A by blast + have A2: "ord x = ord y mod int n" + using A by blast + obtain k where k_def: "k = ord x" + by blast + obtain l where l_def: "ord y = ord x + (l:: int)*(int n)" + using assms A2 + by (metis A k_def mod_eqE mod_mod_trivial mult_of_nat_commute) + have m_def': "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" + using m_def + by blast + have 0: "ord (y \ (\[^](- l*n))) = ord x" + proof- + have 0: "ord (y \ (\[^](- l*n))) = ord y + (ord (\[^](- l*n)))" + using ord_mult p_nonzero A0 Qp_int_pow_nonzero + by blast + have 1: "ord (\[^](- l*n)) = - l*n" + using ord_p_pow_int[of "-l*n"] + by blast + then have "ord (y \ (\[^](- l*n))) = ord y - l*n" + using 0 + by linarith + then show ?thesis + using k_def l_def by linarith + qed + have 1: "ac m (y \ (\[^](- l*n))) = ac m y" + using assms ac_p_int_pow_factor_right[of ] m_def A Qp.nonzero_closed by presburger + have 2: "y \ (\[^](- l*n)) \ nonzero Q\<^sub>p" + using A0 Qp_int_pow_nonzero[of \ "- l*n"] Qp.cring_axioms nonzero_def cring.cring_simprules(5) + fract_cancel_left not_nonzero_Qp p_intpow_inv'' p_nonzero zero_fract Qp.integral_iff + Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_memI Qp.nonzero_mult_closed + minus_mult_commute mult_minus_right p_intpow_closed(1) p_intpow_closed(2) + by presburger + then have 3: "pow_res n (y \ (\[^](- l*n))) = pow_res n x" + using 2 A0 m_def'[of "y \ (\[^](- l*n))" x] "0" "1" A1 + by linarith + have 4: "(y \ (\[^](- l*n))) = (y \ ((\[^]- l)[^]n))" + using Qp_int_nat_pow_pow[of \ "-l" n] p_nonzero + by presburger + have "y \ (\[^](- l*n))\ pow_res n y " + using "2" "4" Qp_int_pow_nonzero nonzero_def p_nonzero + unfolding pow_res_def nonzero_def + proof - + assume a1: "\x n. x \ {a \ carrier Q\<^sub>p. a \ \} \ x [^] (n::int) \ {a \ carrier Q\<^sub>p. a \ \}" + assume a2: "\ \ {a \ carrier Q\<^sub>p. a \ \}" + assume a3: "y \ \ [^] (- l * int n) \ {a \ carrier Q\<^sub>p. a \ \}" + have f4: "\ [^] (- 1 * l) \ {r \ carrier Q\<^sub>p. r \ \}" + using a2 a1 by presburger + have f5: "- l = - 1 * l" + by linarith + then have f6: "y \ \ [^] (- 1 * l * int n) = y \ (\ [^] (- 1 * l)) [^] n" + using \y \ \ [^] (- l * int n) = y \ (\ [^] - l) [^] n\ by presburger + then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. r \ \}" + using f5 a3 by presburger + then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" + using f4 by blast + then have "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" + using f6 f5 by presburger + then show "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra\{r \ carrier Q\<^sub>p. r \ \}. r = y \ ra [^] n}" + by meson + qed + then have "pow_res n (y \ (\[^](- l*n))) = pow_res n y" + using equal_pow_resI[of "(y \ (\[^](- l*n)))" y n] "2" A0 assms + Qp.nonzero_mult_closed p_intpow_closed(2) + by (metis (mono_tags, opaque_lifting) "3" A Qp.nonzero_closed equal_pow_resI) + then show ?thesis using 3 by blast + qed + qed + qed + qed + then show ?thesis + using m_def + by blast +qed + +definition pow_res_classes where +"pow_res_classes n = {S. \x \ nonzero Q\<^sub>p. S = pow_res n x }" + +lemma pow_res_semialg_def: + assumes "x \ nonzero Q\<^sub>p" + assumes "n \ 1" + shows "\P \ carrier Q\<^sub>p_x. pow_res n x = (univ_basic_semialg_set n P) - {\}" +proof- + have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (inv x) \ a = (y[^]n)}" + proof + show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" + proof + fix a + assume A: "a \ pow_res n x" + then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))" + unfolding pow_res_def + by blast + then obtain y where y_def: "y \ nonzero Q\<^sub>p \a = x \ (y[^]n)" + by blast + then have "y \ nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" + proof - + show ?thesis + by (metis (no_types, opaque_lifting) Qp.m_assoc Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed + \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ y [^] n)\ assms(1) local.fract_cancel_right nonzero_inverse_Qp y_def) + qed + then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" + using assms \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))\ + by blast + qed + + show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)} \ pow_res n x" + proof + fix a + assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" + show "a \ pow_res n x" + proof- + have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" + using A by blast + then obtain y where y_def: "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" + by blast + then have "y\nonzero Q\<^sub>p \ a = x \(y[^]n)" + by (metis Qp.l_one Qp.m_assoc Qp.nonzero_closed Qp.not_nonzero_memI + \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = y [^] n)\ assms(1) field_inv(2) inv_in_frac(1)) + then show ?thesis + by (metis (mono_tags, lifting) \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))\ mem_Collect_eq pow_res_def) + qed + qed + qed + obtain P where P_def: "P = up_ring.monom Q\<^sub>p_x (inv x) 1" + by blast + have P_closed: "P \ carrier Q\<^sub>p_x" + using P_def Qp.nonzero_closed Qp.nonzero_memE(2) UPQ.is_UP_monomE(1) UPQ.is_UP_monomI assms(1) inv_in_frac(1) by presburger + have P_eval: "\a. a \ carrier Q\<^sub>p \ (P \ a) = (inv x) \ a" + using P_def to_fun_monom[of ] + by (metis Qp.nat_pow_eone Qp.nonzero_closed Qp.not_nonzero_memI assms(1) inv_in_frac(1)) + have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" + proof + show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + proof fix a + assume "a \ pow_res n x" + then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" + using 0 + by blast + then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + using P_eval + by (metis (mono_tags, lifting) mem_Collect_eq) + qed + show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ pow_res n x" + proof fix a + assume "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" + by blast + then have "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" + using P_eval \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ + by blast + then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" + using \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast + then show "a \ pow_res n x" + using 0 + by blast + qed + qed + have 1: "univ_basic_semialg_set n P - {\} = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" + proof + show "univ_basic_semialg_set n P - {\} \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + proof + fix a + assume A: "a \ univ_basic_semialg_set n P - {\}" + then have A0: "a \ carrier Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" + unfolding univ_basic_semialg_set_def by blast + then have A0': "a \ nonzero Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" + using A + by (metis DiffD2 not_nonzero_Qp singletonI) + then obtain y where y_def: "y\carrier Q\<^sub>p \ P \ a = (y[^]n)" + by blast + have A1: "(P \ a) \ \" + using P_eval A0' Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) assms(1) inv_in_frac(1) inv_in_frac(2) by presburger + have A2: "y \ nonzero Q\<^sub>p" + proof- + have A20: "(y[^]n) \\" + using A1 y_def + by blast + have "y \ \" + apply(rule ccontr) using A20 assms + by (metis Qp.nat_pow_eone Qp.semiring_axioms Qp.zero_closed le_zero_eq semiring.nat_pow_zero) + then show ?thesis + using y_def A1 not_nonzero_Qp Qp.not_nonzero_memE by blast + qed + then have "y \ nonzero Q\<^sub>p \ P \ a = (y[^]n)" using y_def + by blast + then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + using A0 nonzero_def + by blast + qed + show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ univ_basic_semialg_set n P - {\}" + proof + fix a + assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" + then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" + by blast + then have "y \\ \ y\ carrier Q\<^sub>p \ P \ a = (y[^]n)" + by (metis (mono_tags, opaque_lifting) Qp.nonzero_closed Qp.not_nonzero_memI) + then have "a \\" + using P_eval + by (metis Qp.m_comm Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.zero_closed assms(1) inv_in_frac(1) zero_fract) + then show "a \ univ_basic_semialg_set n P - {\}" + unfolding univ_basic_semialg_set_def + using A \y \ \ \ y \ carrier Q\<^sub>p \ P \ a = (y[^]n)\ + by blast + qed + qed + show ?thesis using 0 1 + by (metis (no_types, lifting) P_closed) +qed + +lemma pow_res_is_univ_semialgebraic: + assumes "x \ carrier Q\<^sub>p" + shows "is_univ_semialgebraic (pow_res n x)" +proof(cases "n = 0") + case True + have T0: "pow_res n x = {x}" + unfolding True using assms + by (simp add: assms zeroth_pow_res) + have "[x] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using assms Qp.to_R1_closed by blast + hence "is_semialgebraic 1 {[x]}" + using is_algebraic_imp_is_semialg singleton_is_algebraic by blast + thus ?thesis unfolding T0 using assms + by (simp add: \x \ carrier Q\<^sub>p\ finite_is_univ_semialgebraic) +next + case False + show ?thesis + proof(cases "x = \") + case True + then show ?thesis using finite_is_univ_semialgebraic False pow_res_of_zero + by (metis Qp.zero_closed empty_subsetI finite.emptyI finite.insertI insert_subset) + next + case F: False + then show ?thesis + using False pow_res_semialg_def[of x n] diff_is_univ_semialgebraic[of _ "{\}"] finite_is_univ_semialgebraic[of "{\}"] + by (metis Qp.zero_closed assms empty_subsetI finite.emptyI finite.insertI insert_subset less_one less_or_eq_imp_le linorder_neqE_nat not_nonzero_Qp univ_basic_semialg_set_is_univ_semialgebraic) + qed +qed + +lemma pow_res_is_semialg: + assumes "x \ carrier Q\<^sub>p" + shows "is_semialgebraic 1 (to_R1 ` (pow_res n x))" + using assms pow_res_is_univ_semialgebraic is_univ_semialgebraicE + by blast + +lemma pow_res_refl: + assumes "x \ carrier Q\<^sub>p" + shows "x \ pow_res n x" +proof- + have "x = x \ (\ [^]n)" + using assms Qp.nat_pow_one Qp.r_one by presburger + thus ?thesis + using assms unfolding pow_res_def mem_Collect_eq + using Qp.one_nonzero by blast +qed + +lemma equal_pow_resE: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "n > 0" + assumes "pow_res n a = pow_res n b" + shows "\ s \ P_set n. a = b \ s" +proof- + have "a \ pow_res n b" + using assms pow_res_refl by blast + then obtain y where y_def: " y \ nonzero Q\<^sub>p \ a = b \ y[^]n" + unfolding pow_res_def by blast + thus ?thesis unfolding P_set_def + using Qp.nonzero_closed Qp_nat_pow_nonzero by blast +qed + +lemma pow_res_one: + assumes "x \ nonzero Q\<^sub>p" + shows "pow_res 1 x = nonzero Q\<^sub>p" +proof show "pow_res 1 x \ nonzero Q\<^sub>p" + using assms nonzero_pow_res[of x 1] by blast + show "nonzero Q\<^sub>p \ pow_res 1 x" + proof fix y assume A: "y \ nonzero Q\<^sub>p" + then have 0: "\ \ nonzero Q\<^sub>p \ y = x \ ((inv x)\ y)[^](1::nat)" + using assms Qp.m_comm Qp.nat_pow_eone Qp.nonzero_closed Qp.nonzero_mult_closed + Qp.one_nonzero local.fract_cancel_right nonzero_inverse_Qp by presburger + have 1: "(inv x)\ y \ nonzero Q\<^sub>p" + using A assms by (metis Qp.Units_m_closed Units_eq_nonzero nonzero_inverse_Qp) + then show "y \ pow_res 1 x" + unfolding pow_res_def using 0 1 A Qp.nonzero_closed by blast + qed +qed + + +lemma pow_res_zero: + assumes "n > 0" + shows "pow_res n \ = {\}" +proof + show "pow_res n \ \ {\}" + unfolding pow_res_def + using Qp.l_null Qp.nat_pow_closed Qp.nonzero_closed by blast + show "{\} \ pow_res n \" + using assms unfolding pow_res_def + using Qp.l_null Qp.one_closed Qp.one_nonzero empty_subsetI insert_subset by blast +qed + + +lemma equal_pow_resI': + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ P_set n" + assumes "a = b \ c" + assumes "n > 0" + shows "pow_res n a = pow_res n b" +proof- + obtain y where y_def: "c = y[^]n \ y \ carrier Q\<^sub>p" + using assms unfolding P_set_def by blast + have c_nonzero: "c \ nonzero Q\<^sub>p" + using P_set_nonzero'(1) assms(3) by blast + have y_nonzero: "y \ nonzero Q\<^sub>p" + using y_def c_nonzero Qp_nonzero_nat_pow assms(5) by blast + have 0: "a \ pow_res n b" + using assms y_nonzero y_def unfolding pow_res_def + by blast + show ?thesis + apply(cases "b = \") + using pow_res_zero Qp.l_null Qp.nonzero_closed assms(4) c_nonzero apply presburger + by (metis "0" assms(1) assms(2) assms(5) not_nonzero_Qp equal_pow_resI) +qed + +lemma equal_pow_resI'': + assumes "n > 0" + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "a \ inv b \ P_set n" + shows "pow_res n a = pow_res n b" + using assms equal_pow_resI'[of a b "a \ inv b" n] Qp.nonzero_closed local.fract_cancel_right + by blast + +lemma equal_pow_resI''': + assumes "n > 0" + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "c \ nonzero Q\<^sub>p" + assumes "pow_res n (c \ a) = pow_res n (c \ b)" + shows "pow_res n a = pow_res n b" +proof- + have 0: "c \a \ nonzero Q\<^sub>p" + by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) assms(4)) + have 1: "c \b \ nonzero Q\<^sub>p" + by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(3) assms(4)) + have "c\a \ pow_res n (c\b)" + proof(cases "n = 1") + case True + then show ?thesis + using assms 0 1 pow_res_one[of "c\b"] by blast + next + case False + then have "n \ 2" + using assms(1) by linarith + then show ?thesis + using assms 0 1 pow_res_refl[of "c\a" n] unfolding nonzero_def + by blast + qed + then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (c \ a) = (c \ b)\y[^]n" + using assms unfolding pow_res_def by blast + then have "a = b\y[^]n" + using assms + by (metis Qp.m_assoc Qp.m_lcancel Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero) + then show ?thesis + by (metis P_set_memI Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) assms(3) equal_pow_resI' y_def) +qed + +lemma equal_pow_resI'''': + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "a = b \ u" + assumes "u \ P_set n" + shows "pow_res n a = pow_res n b" +proof(cases "a = \") + case True + then have "b = \" + using assms unfolding P_set_def + by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.not_nonzero_memI mem_Collect_eq) + then show ?thesis using pow_res_zero + using True by blast +next + case False + then have 0: "a \ nonzero Q\<^sub>p" + using Qp.not_nonzero_memE assms(2) by blast + have 1: "b \ nonzero Q\<^sub>p" + using 0 assms unfolding P_set_def + by (metis (no_types, lifting) Qp.integral_iff Qp.nonzero_closed mem_Collect_eq not_nonzero_Qp) + have 2: "a \ (inv b)\ P_set n" + using assms 0 1 + by (metis P_set_nonzero'(2) Qp.inv_cancelR(1) Qp.m_comm Qp.nonzero_memE(2) Units_eq_nonzero inv_in_frac(1)) + then show ?thesis using equal_pow_resI'' + by (meson "0" "1" assms(1) equal_pow_resI) +qed + + +lemma Zp_Units_ord_zero: + assumes "a \ Units Z\<^sub>p" + shows "ord_Zp a = 0" +proof- + have "inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p" + apply(rule Zp.nonzero_memI, rule Zp.Units_inv_closed, rule assms) + using assms Zp.Units_inverse in_Units_imp_not_zero by blast + then have "ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub> a) = ord_Zp a + ord_Zp (inv \<^bsub>Z\<^sub>p\<^esub> a)" + using assms ord_Zp_mult Zp.Units_nonzero zero_not_one + by (metis Zp.zero_not_one) + then show ?thesis + by (smt Zp.Units_closed Zp.Units_r_inv Zp.integral_iff Zp.nonzero_closed \inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p\ assms ord_Zp_one ord_pos) +qed + +lemma pow_res_nth_pow: + assumes "a \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "pow_res n (a[^]n) = pow_res n \" +proof + show "pow_res n (a [^] n) \ pow_res n \" + proof fix x assume A: "x \ pow_res n (a [^] n)" + then show "x \ pow_res n \" + by (metis P_set_memI Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) + Qp.nonzero_pow_nonzero Qp.one_closed assms(1) assms(2) equal_pow_resI') + qed + show "pow_res n \ \ pow_res n (a [^] n)" + proof fix x assume A: "x \ pow_res n \" + then obtain y where y_def: "y \ nonzero Q\<^sub>p \ x = \\y[^]n" + unfolding pow_res_def by blast + then have 0: "x = y[^]n" + using Qp.l_one Qp.nonzero_closed by blast + have "y[^]n = a[^]n \ (inv a \ y)[^]n" + proof- + have "a[^]n \ (inv a \ y)[^]n = a[^]n \ (inv a)[^]n \ y[^]n" + using Qp.Units_inv_closed Qp.m_assoc Qp.nat_pow_closed Qp.nat_pow_distrib Qp.nonzero_closed Units_eq_nonzero assms(1) y_def by presburger + then show ?thesis + by (metis Qp.Units_inv_inv Qp.inv_cancelR(1) Qp.nat_pow_distrib Qp.nonzero_closed Qp.nonzero_mult_closed Units_eq_nonzero assms(1) nonzero_inverse_Qp y_def) + qed + then show "x \ pow_res n (a [^] n)" + using y_def A assms unfolding pow_res_def mem_Collect_eq + by (metis "0" Qp.integral Qp.m_closed Qp.nonzero_closed Qp.not_nonzero_memI inv_in_frac(1) inv_in_frac(2) not_nonzero_Qp) + qed +qed + +lemma pow_res_of_p_pow: + assumes "n > 0" + shows "pow_res n (\[^]((l::int)*n)) = pow_res n \" +proof- + have 0: "\[^]((l::int)*n) = (\[^]l)[^]n" + using Qp_p_int_nat_pow_pow by blast + have "\[^]((l::int)*n) \ P_set n" + using P_set_memI[of _ "\[^]l"] + by (metis "0" Qp.not_nonzero_memI Qp_int_pow_nonzero p_intpow_closed(1) p_nonzero) + thus ?thesis + using "0" assms p_intpow_closed(2) pow_res_nth_pow by presburger +qed + +lemma pow_res_nonzero: + assumes "n > 0" + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "pow_res n a = pow_res n b" + shows "b \ nonzero Q\<^sub>p" + using assms nonzero_pow_res[of a n] pow_res_zero[of n] + by (metis insert_subset not_nonzero_Qp) + +lemma pow_res_mult: + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + assumes "pow_res n a = pow_res n c" + assumes "pow_res n b = pow_res n d" + shows "pow_res n (a \ b) = pow_res n (c \ d)" +proof(cases "a \ nonzero Q\<^sub>p") + case True + then have "c \ nonzero Q\<^sub>p" + using assms pow_res_nonzero by blast + then obtain \ where alpha_def: "\ \ nonzero Q\<^sub>p \ a = c \ \[^]n" + using assms True pow_res_refl[of a n] unfolding assms unfolding pow_res_def + by blast + show ?thesis + proof(cases "b \ nonzero Q\<^sub>p") + case T: True + then have "d \ nonzero Q\<^sub>p" + using assms pow_res_nonzero by blast + then obtain \ where beta_def: "\ \ nonzero Q\<^sub>p \ b = d \ \[^]n" + using T pow_res_refl[of b n] unfolding assms unfolding pow_res_def + using assms by blast + then have "a \ b = (c \ d) \ (\[^]n \ \[^] n)" + using Qp.m_assoc Qp.m_lcomm Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero alpha_def assms(3) assms(4) assms(5) by presburger + then have 0: "a \ b = (c \ d) \ ((\ \ \)[^] n)" + by (metis Qp.nat_pow_distrib Qp.nonzero_closed alpha_def beta_def) + show ?thesis + apply(intro equal_pow_resI'[of _ _ "(\ \ \)[^] n"] Qp.ring_simprules assms + P_set_memI[of _ "\ \ \"] Qp.nat_pow_closed nonzero_memE 0 Qp_nat_pow_nonzero + ) + using alpha_def beta_def apply auto + apply(intro nonzero_memI Qp.nonzero_mult_closed) + using alpha_def beta_def nonzero_memE apply auto + by (meson Qp.integral_iff) + next + case False + then have "b = \" + by (meson assms not_nonzero_Qp) + then have "d = \" + using assms by (metis False not_nonzero_Qp pow_res_nonzero) + then show ?thesis + using Qp.r_null \b = \\ assms by presburger + qed +next + case False + then have "a = \" + by (meson assms not_nonzero_Qp) + then have "c = \" + using assms by (metis False not_nonzero_Qp pow_res_nonzero) + then show ?thesis + using Qp.r_null \a = \\ assms Qp.l_null by presburger + qed + +lemma pow_res_p_pow_factor: + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + shows "pow_res n a = pow_res n (\[^]((l::int)*n) \ a)" +proof(cases "a = \") + case True + then show ?thesis + using Qp.r_null p_intpow_closed(1) by presburger +next + case False + then show ?thesis using assms pow_res_of_p_pow + using Qp.m_comm Qp.one_closed Qp.r_one p_intpow_closed(1) pow_res_mult by presburger +qed + +lemma pow_res_classes_finite: + assumes "n \ 1" + shows "finite (pow_res_classes n)" +proof(cases "n = 1") + case True + have "pow_res_classes n = {(nonzero Q\<^sub>p)}" + using True pow_res_one unfolding pow_res_classes_def + using Qp.one_nonzero by blast + then show ?thesis by auto +next + case False + then have n_bound: "n \ 2" + using assms by linarith + obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" + using assms False pow_res_equal_ord n_bound + by (metis gr_zeroI le_numeral_extra(2)) + obtain f where f_def: "f = (\ \ \. (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)))" + by blast + have 0: "\x. x \ nonzero Q\<^sub>p \ pow_res n x = f (ac m x) (ord x)" + proof- fix x assume A: "x \ nonzero Q\<^sub>p" + obtain \ where eta_def: "\ = ac m x" + by blast + obtain \ where nu_def: "\ = ord x" + by blast + have "\y \pow_res n x. ac m y = ac m x \ ord y = ord x" + using pow_res_refl A assms neq0_conv Qp.nonzero_closed by blast + hence "pow_res n x \ (pow_res_classes n) \ (\ y \ (pow_res n x). ac m y = \ \ ord y = \)" + unfolding nu_def eta_def using assms unfolding pow_res_classes_def + using A by blast + then have 0: "(\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" + by blast + have "f \ \ = (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" + using f_def by blast + then have 1: "f \ \ \ (pow_res_classes n) \ ((\ y \ (f \ \). ac m y = \ \ ord y = \))" + using 0 someI_ex[of "\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)"] + unfolding f_def by blast + then obtain y where y_def: "y \ (f \ \) \ ac m y = ac m x \ ord y = ord x" + unfolding nu_def eta_def by blast + obtain a where a_def: "a \ nonzero Q\<^sub>p \ (f \ \) = pow_res n a" + using 1 unfolding pow_res_classes_def by blast + then have 2: "y \ pow_res n a" + using y_def by blast + have 3: "y \ nonzero Q\<^sub>p" + using y_def nonzero_pow_res[of a n] a_def by blast + then have 4: "pow_res n y = pow_res n a" + using 3 y_def a_def equal_pow_resI[of y a n] n_bound Qp.nonzero_closed + by (metis equal_pow_resI) + have 5: "pow_res n y = f \ \" + using 4 a_def by blast + then show "pow_res n x = f (ac m x) (ord x)" + unfolding eta_def nu_def + using "3" A m_def y_def by blast + qed + obtain N where N_def: "N > 0 \ (\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n)" + using n_bound nth_power_fact assms by blast + have 1: "\x. x \ nonzero Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. ord y \ 0 \ ord y < n \ pow_res n x = pow_res n y)" + proof- fix x assume x_def: "x \ nonzero Q\<^sub>p" + then obtain k where k_def: "k = ord x mod n" + by blast + then obtain l where l_def: "ord x = (int n)*l + k" + using cancel_div_mod_rules(2)[of n "ord x"0] unfolding k_def + by (metis group_add_class.add.right_cancel) + have "x = (\[^](ord x)) \ \ (angular_component x)" + using x_def angular_component_factors_x by blast + then have "x = (\[^](n*l + k)) \ \ (angular_component x)" + unfolding l_def by blast + hence "x = \[^](int n*l) \ (\[^] k) \ \ (angular_component x)" + by (metis p_intpow_add) + hence 0: "x = (\[^]l)[^]n \ (\[^] k) \ \ (angular_component x)" + using p_pow_factor[of n l k] \x = \ [^] (int n * l + k) \ \ (angular_component x)\ by presburger + have 1: "\ (angular_component x) \ carrier Q\<^sub>p" + using x_def angular_component_closed inc_closed by blast + hence 2: "x = (\[^]l)[^]n \ ((\[^] k) \ \ (angular_component x))" + using 0 by (metis Qp.m_assoc Qp.nat_pow_closed p_intpow_closed(1)) + obtain a where a_def: "a = (\[^] k) \ \ (angular_component x)" + by blast + have 30: "angular_component x \ Units Z\<^sub>p" + using angular_component_unit x_def by blast + then have 3: "\ (angular_component x) \ Units Q\<^sub>p" + by (metis Units_eq_nonzero Zp.Units_closed in_Units_imp_not_zero inc_of_nonzero not_nonzero_Qp) + have 4: "\ (angular_component x) \ nonzero Q\<^sub>p" + using 3 Units_nonzero_Qp by blast + have a_nonzero: "a \ nonzero Q\<^sub>p" + unfolding a_def 4 + by (meson "3" Qp.UnitsI(1) Qp.Units_m_closed Units_nonzero_Qp p_intpow_closed(1) p_intpow_inv) + have 5: "x = a \(\[^]l)[^]n" + using 2 a_nonzero unfolding a_def + using Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed p_intpow_closed(1) by presburger + hence "x \ pow_res n a" + unfolding pow_res_def + using Qp.nonzero_closed Qp_int_pow_nonzero p_nonzero x_def by blast + hence 6:"pow_res n a = pow_res n x" + using x_def a_def equal_pow_resI[of x a n] a_nonzero n_bound Qp.nonzero_closed equal_pow_resI + by blast + have 7: "ord (\ (angular_component x)) = 0" + proof- + have "ord_Zp (angular_component x) = 0" using 30 Zp_Units_ord_zero by blast + then have "val_Zp (angular_component x) = 0" + using "30" unit_imp_val_Zp0 by blast + then have "val (\ (angular_component x)) = 0" + by (metis angular_component_closed val_of_inc x_def) + then show ?thesis using angular_component_closed x_def + by (metis "30" Zp.Units_closed \ord_Zp (angular_component x) = 0\ in_Units_imp_not_zero not_nonzero_Qp ord_of_inc) + qed + have 8: "ord a = k" + unfolding a_def using 3 4 7 ord_mult[of "\ [^] k" "\ (angular_component x)"] ord_p_pow_int[of k] + p_pow_nonzero + using Qp_int_pow_nonzero p_nonzero by presburger + have 9: "k < n" + unfolding k_def + using assms by auto + show " \y\nonzero Q\<^sub>p. 0 \ ord y \ ord y < int n \ pow_res n x = pow_res n y" + by (metis "6" "8" a_nonzero assms k_def less_imp_of_nat_less of_nat_eq_0_iff pos_mod_conj rel_simps(45) zero_less_iff_neq_zero) + qed + have 2: "\x. x \ (pow_res_classes n) \ \ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" + proof- fix a assume A: "a \ (pow_res_classes n)" + then obtain x where x_def: "x \ nonzero Q\<^sub>p \ a = pow_res n x" + unfolding pow_res_classes_def by blast + then obtain x' where x'_def: "x' \ nonzero Q\<^sub>p \ ord x' \ 0 \ ord x' < n \ pow_res n x' = a" + using 1[of x] unfolding x_def by blast + hence 20: "f (ac m x') (ord x') = a" + using 0 by blast + have 21: "ac m x' \ Units (Zp_res_ring m)" + using x'_def ac_units m_def by presburger + then have 22: "ac m x' \ Units (Zp_res_ring m) \ (ord x') \ ({0.. a = f (ac m x') (ord x')" + using x'_def 20 atLeastLessThan_iff by blast + then show "\ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. a = f \ \" by blast + qed + obtain F where F_def: "F = (\ps. f (fst ps) (snd ps))" + by blast + have 3: "\x. x \ (pow_res_classes n) \ \ ps \ Units (Zp_res_ring m) \ {0.. pow_res_classes n" + obtain \ \ where eta_nu_def: " \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" + using 2 A by blast + then have "F (\, \) = x" + unfolding F_def by (metis fst_conv snd_conv) + then show " \ ps \ Units (Zp_res_ring m) \ {0.. F ` (Units (Zp_res_ring m) \ {0.. {0.. pow_res_classes n" + shows "is_univ_semialgebraic S" +proof- + obtain x where x_def: "x\nonzero Q\<^sub>p \ S = pow_res n x" + using assms unfolding pow_res_classes_def by blast + then show ?thesis using pow_res_is_univ_semialgebraic + using Qp.nonzero_closed by blast +qed + +lemma pow_res_classes_semialg: + assumes "S \ pow_res_classes n" + shows "is_semialgebraic 1 (to_R1` S)" + using pow_res_classes_univ_semialg + assms(1) is_univ_semialgebraicE by blast + +definition nth_pow_wits where +"nth_pow_wits n = (\ S. (SOME x. x \ (S \ \\<^sub>p)))` (pow_res_classes n)" + +lemma nth_pow_wits_finite: + assumes "n > 0" + shows "finite (nth_pow_wits n)" +proof- + have "n \ 1" + by (simp add: assms leI) + thus ?thesis + unfolding nth_pow_wits_def using assms pow_res_classes_finite[of n] by blast +qed + +lemma nth_pow_wits_covers: + assumes "n > 0" + assumes "x \ nonzero Q\<^sub>p" + shows "\y \ (nth_pow_wits n). y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" +proof- + have PP: "(pow_res n x) \ pow_res_classes n" + unfolding pow_res_classes_def using assms by blast + obtain k where k_def: "val x = eint k" + using assms val_ord by blast + obtain N::int where N_def: "N = (if k < 0 then -k else k)" by blast + then have N_nonneg: "N \ 0" + unfolding N_def + by presburger + have 0: "int n \ 1" + using assms by linarith + have "N*(int n) + k \ 0" + proof(cases "k<0") + case True then have "N = -k" unfolding N_def + by presburger + then have "N*n + k = k*(1- int n)" + using distrib_left[of k 1 "-int n"] mult_cancel_left2 mult_minus_left + by (metis add.inverse_inverse diff_minus_eq_add minus_mult_minus neg_equal_iff_equal uminus_add_conv_diff) + then show ?thesis using 0 True zero_less_mult_iff[of k "1 - int n"] + proof - + have "0 \ N * (int n - 1)" + by (meson "0" N_nonneg diff_ge_0_iff_ge zero_le_mult_iff) + then show ?thesis + by (metis (no_types) \N = - k\ add.commute distrib_left minus_add_cancel mult_minus1_right uminus_add_conv_diff) + qed + next + case False + then have "N = k" unfolding N_def + by presburger + then show ?thesis using 0 False + by (metis N_nonneg add_increasing2 mult_nonneg_nonneg of_nat_0_le_iff) + qed + then have 1: "ord (\[^](N*n)\x) \ 0" + using ord_mult k_def val_ord assms + by (metis Qp_int_pow_nonzero eint.inject ord_p_pow_int p_nonzero) + have 2: "\[^](N*n)\x \ pow_res n x" + proof- + have "\[^](N*n) = (\[^]N)[^]n" + using Qp_p_int_nat_pow_pow by blast + then have "\[^]N \ nonzero Q\<^sub>p \ \[^](N*n)\x = x \ (\[^]N)[^]n" + by (metis Qp.m_comm Qp.nonzero_closed Qp_int_pow_nonzero assms(2) p_nonzero) + then show ?thesis unfolding pow_res_def + by (metis (mono_tags, lifting) Qp.m_closed Qp.nonzero_closed assms(2) mem_Collect_eq p_intpow_closed(1)) + qed + have 3: "\[^](N*n)\x \ \\<^sub>p" + using 1 assms + by (metis Q\<^sub>p_def Qp.nonzero_mult_closed Qp_int_pow_nonzero Z\<^sub>p_def val_ring_ord_criterion \_def p_nonzero padic_fields.zero_in_val_ring padic_fields_axioms) + have 4: "x \ pow_res n (\[^](N*n)\x)" + using 2 equal_pow_resI[of x "\[^](N*n)\x" n] pow_res_refl[of "\[^](N*n)\x" n] assms + Qp.nonzero_mult_closed p_intpow_closed(2) pow_res_refl Qp.nonzero_closed by metis + have 5: "\[^](N*n)\x \ (pow_res n x \ \\<^sub>p)" + using 2 3 by blast + have 6: "(SOME z. z \ (pow_res n x) \ \\<^sub>p) \ pow_res n x \ \\<^sub>p" using 5 + by (meson someI) + obtain y where y_def: "y = (SOME z. z \ (pow_res n x) \ \\<^sub>p)" + by blast + then have A: "y \ pow_res n x" + using "6" by blast + then have "pow_res n x = pow_res n y" + using equal_pow_resI[of x y n] assms y_def Qp.nonzero_closed nonzero_pow_res by blast + then have 7: "x \ pow_res n y" + using pow_res_refl[of x n] assms unfolding nonzero_def by blast + have 8: "y \ nonzero Q\<^sub>p " + using y_def PP 6 A nonzero_pow_res[of x n] assms + by blast + have 9: "y \ \\<^sub>p" + using y_def "6" by blast + have "y\(\S. SOME x. x \ S \ \\<^sub>p) ` pow_res_classes n \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" + using y_def PP 6 7 8 9 A nonzero_pow_res[of x n] assms + by blast + then show ?thesis unfolding nth_pow_wits_def by blast +qed + +lemma nth_pow_wits_closed: + assumes "n > 0" + assumes "x \ nth_pow_wits n" + shows "x \ carrier Q\<^sub>p" "x \ \\<^sub>p" "x \ nonzero Q\<^sub>p" "\ y \ pow_res_classes n. y = pow_res n x" +proof- + obtain c where c_def: "c \ pow_res_classes n \ x = (SOME x. x \ (c \ \\<^sub>p))" + by (metis (no_types, lifting) assms(2) image_iff nth_pow_wits_def) + then obtain y where y_def: "y \ nonzero Q\<^sub>p \ c = pow_res n y" + unfolding pow_res_classes_def by blast + then obtain a where a_def: "a \ (nth_pow_wits n) \ a \ nonzero Q\<^sub>p \ a \ \\<^sub>p \ y \ pow_res n a" + using nth_pow_wits_covers[of n y] assms(1) by blast + have 00: "pow_res n a = c" + using equal_pow_resI[of a y n] y_def assms a_def unfolding nonzero_def by blast + then have P :"a \ c \ \\<^sub>p" + using pow_res_refl[of a n] assms a_def unfolding 00 nonzero_def by blast + then show 0: "x \ \\<^sub>p" using c_def + by (metis Collect_mem_eq Int_Collect tfl_some) + then show "x \ carrier Q\<^sub>p" + using val_ring_memE by blast + have 1: "c \ nonzero Q\<^sub>p" + using c_def nonzero_pow_res[of y n] unfolding pow_res_classes_def + using assms(1) y_def by blast + have "(SOME x. x \ (c \ \\<^sub>p)) \ (c \ \\<^sub>p)" + using P tfl_some + by (smt Int_def someI_ex) + then have 2: "x \ c" + using c_def by blast + thus "x \ nonzero Q\<^sub>p" + using 1 by blast + show "\ y \ pow_res_classes n. y = pow_res n x" + using 00 2 c_def P a_def equal_pow_resI[of a x n] 0 val_ring_memE assms(1) by blast +qed + +lemma finite_extensional_funcset: + assumes "finite A" + assumes "finite (B::'b set)" + shows "finite (A \\<^sub>E B)" + using finite_PiE[of A "\_. B"] assms by blast + +lemma nth_pow_wits_exists: + assumes "m > 0" + assumes "c \ pow_res_classes m" + shows "\x. x \ c \ \\<^sub>p" +proof- + obtain x where x_def: "x \ nonzero Q\<^sub>p \ pow_res m x = c" + using assms unfolding pow_res_classes_def by blast + obtain y where y_def: "y \ (nth_pow_wits m) \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res m y" + using nth_pow_wits_covers assms x_def + by blast + have 0: "pow_res m x = pow_res m y" + using x_def y_def equal_pow_resI Qp.nonzero_closed assms(1) by blast + then have 1: "y \ pow_res m x" + using pow_res_refl[of y m ] y_def assms unfolding nonzero_def by blast + thus ?thesis using x_def y_def assms + by blast +qed + +lemma pow_res_classes_mem_eq: + assumes "m > 0" + assumes "a \ pow_res_classes m" + assumes "x \ a" + shows "a = pow_res m x" +proof- + obtain y where y_def: "y \ nonzero Q\<^sub>p \ a = pow_res m y" + using assms unfolding pow_res_classes_def by blast + then show ?thesis using assms equal_pow_resI[of y x m] + by (meson Qp.nonzero_closed nonzero_pow_res equal_pow_resI subsetD) +qed + +lemma nth_pow_wits_neq_pow_res: + assumes "m > 0" + assumes "x \ nth_pow_wits m" + assumes "y \ nth_pow_wits m" + assumes "x \ y" + shows "pow_res m x \ pow_res m y" +proof- + obtain a where a_def: "a \ pow_res_classes m \ x = (\ S. (SOME x. x \ (S \ \\<^sub>p))) a" + using assms unfolding nth_pow_wits_def by blast + obtain b where b_def: "b \ pow_res_classes m \ y = (\ S. (SOME x. x \ (S \ \\<^sub>p))) b" + using assms unfolding nth_pow_wits_def by blast + have a_neq_b: "a \ b" + using assms a_def b_def by blast + have 0: "x \ a \ \\<^sub>p" + using a_def nth_pow_wits_exists[of m a] assms + by (meson someI_ex) + have 1: "y \ b \ \\<^sub>p" + using b_def nth_pow_wits_exists[of m b] assms + by (meson someI_ex) + have 2: "pow_res m x = a" + using a_def pow_res_classes_mem_eq[of m a x] assms 0 + by blast + have 3: "pow_res m y = b" + using b_def pow_res_classes_mem_eq[of m b y] assms 1 + by blast + show ?thesis + by (simp add: "2" "3" a_neq_b) +qed + +lemma nth_pow_wits_disjoint_pow_res: + assumes "m > 0" + assumes "x \ nth_pow_wits m" + assumes "y \ nth_pow_wits m" + assumes "x \ y" + shows "pow_res m x \ pow_res m y = {}" + using assms nth_pow_wits_neq_pow_res disjoint_iff_not_equal + by (metis (no_types, opaque_lifting) nth_pow_wits_closed(4) pow_res_classes_mem_eq) + +lemma nth_power_fact': + assumes "0 < (n::nat)" + shows "\m>0. \u\carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" + using nth_power_fact[of n] assms + by (metis less_one less_or_eq_imp_le linorder_neqE_nat neq0_conv) + +lemma equal_pow_res_criterion: + assumes "N > 0" + assumes "n > 0" + assumes "\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "a = b \ (\ \ c)" + assumes "val c \ N" + shows "pow_res n a = pow_res n b" +proof(cases "b = \") + case True + then have "a = \" + using assms Qp.add.m_closed Qp.l_null Qp.one_closed by presburger + then show ?thesis using True + by blast +next + case False + then have F0: "a \
b = \ \ c" + by (metis Qp.Units_one_closed Qp.add.m_closed Qp.inv_cancelR(2) Qp.one_closed Qp.unit_factor assms(4) assms(5) assms(6) assms(7) field_inv(2) inv_in_frac(1)) + have "0 < eint N" + using assms by (metis eint_ord_simps(2) of_nat_0_less_iff zero_eint_def) + hence F1: "val \ < val c" + using assms less_le_trans[of 0 N "val c"] unfolding val_one + by blast + hence F2: " val \ = val (\ \ c)" + using assms val_one one_nonzero Qp.add.m_comm Qp.one_closed val_ultrametric_noteq by metis + have "val \ + eint (int N) \ val (\ \ (\ \ c))" + proof- + have "val (\ \ (\ \ c)) = val c" + using Qp.add.inv_closed Qp.minus_eq Qp.minus_sum Qp.one_closed Qp.r_neg2 assms(6) val_minus by presburger + thus ?thesis + unfolding val_one using assms F1 by (metis add.left_neutral) + qed + hence F3: "ac N \ = ac N (\ \ c)" + using F2 F1 assms ac_val[of \ "\ \ c" N] + by (metis Qp.add.m_closed Qp.one_closed val_nonzero) + have F4: "\ \ c \ P_set n" + using assms F1 F2 F3 val_one ac_one + by (metis Qp.add.m_closed Qp.one_closed Qp.one_nonzero ac_inv'' ac_inv'''(1) ac_one') + then show ?thesis + using assms(2) assms(4) assms(5) assms(7) equal_pow_resI' by blast +qed + + + +lemma pow_res_nat_pow: + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "pow_res n a = pow_res n b" + shows "pow_res n (a[^](k::nat)) = pow_res n (b[^]k)" + apply(induction k) + using assms apply (metis Group.nat_pow_0) + using assms pow_res_mult by (smt Qp.nat_pow_Suc2 Qp.nat_pow_closed) + +lemma pow_res_mult': + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + assumes "e \ carrier Q\<^sub>p" + assumes "f \ carrier Q\<^sub>p" + assumes "pow_res n a = pow_res n d" + assumes "pow_res n b = pow_res n e" + assumes "pow_res n c = pow_res n f" + shows "pow_res n (a \ b \ c) = pow_res n (d \ e \ f)" +proof- + have "pow_res n (a \ b) = pow_res n (d \ e)" + using pow_res_mult assms by meson + then show ?thesis using pow_res_mult assms + by (meson Qp.m_closed) +qed + +lemma pow_res_disjoint: + assumes "n > 0" + assumes "a \ nonzero Q\<^sub>p" + assumes "a \ pow_res n \" + shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" + using assms unfolding pow_res_def + using Qp.l_one Qp.nonzero_closed by blast + +lemma pow_res_disjoint': + assumes "n > 0" + assumes "a \ nonzero Q\<^sub>p" + assumes "pow_res n a \ pow_res n \" + shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" +using assms pow_res_disjoint pow_res_refl + by (metis pow_res_nth_pow) + +lemma pow_res_one_imp_nth_pow: + assumes "n > 0" + assumes "a \ pow_res n \" + shows "\y \ nonzero Q\<^sub>p. a = y[^]n" + using assms unfolding pow_res_def + using Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed by blast + +lemma pow_res_eq: + assumes "n > 0" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ pow_res n a" + shows "pow_res n b = pow_res n a" +proof(cases "a = \") + case True + then show ?thesis using assms by (metis pow_res_zero singletonD) +next + case False + then have a_nonzero: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast + show ?thesis + proof(cases "n = 1") + case True + then show ?thesis using a_nonzero assms + using pow_res_one Q\<^sub>p_def Zp_def padic_fields_axioms by blast + next + case False + then have "n \ 2" + using assms(1) by linarith + then show ?thesis using False a_nonzero assms Qp.nonzero_closed nonzero_pow_res equal_pow_resI + by blast + qed +qed + +lemma pow_res_classes_n_eq_one: +"pow_res_classes 1 = {nonzero Q\<^sub>p}" + unfolding pow_res_classes_def using pow_res_one Qp.one_nonzero by blast + +lemma nth_pow_wits_closed': + assumes "n > 0" + assumes "x \ nth_pow_wits n" + shows "x \ \\<^sub>p \ x \ nonzero Q\<^sub>p" using nth_pow_wits_closed + assms by blast + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Semialgebraic Sets Defined by Congruences\ +(**************************************************************************************************) +(**************************************************************************************************) + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\$p$-adic ord Congruence Sets\ + (********************************************************************************************) + (********************************************************************************************) + +lemma carrier_is_univ_semialgebraic: +"is_univ_semialgebraic (carrier Q\<^sub>p)" + apply(rule is_univ_semialgebraicI) + using Qp.to_R1_carrier carrier_is_semialgebraic + by presburger + +lemma nonzero_is_univ_semialgebraic: +"is_univ_semialgebraic (nonzero Q\<^sub>p)" +proof- + have "nonzero Q\<^sub>p = carrier Q\<^sub>p - {\}" + unfolding nonzero_def by blast + then show ?thesis using diff_is_univ_semialgebraic[of "carrier Q\<^sub>p" "{\}"] + by (metis Diff_empty Diff_insert0 carrier_is_univ_semialgebraic empty_subsetI + finite.emptyI finite.insertI finite_is_univ_semialgebraic insert_subset) +qed + +definition ord_congruence_set where +"ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x mod n = a}" + +lemma ord_congruence_set_nonzero: +"ord_congruence_set n a \ nonzero Q\<^sub>p" + by (metis (no_types, lifting) mem_Collect_eq ord_congruence_set_def subsetI) + +lemma ord_congruence_set_closed: +"ord_congruence_set n a \ carrier Q\<^sub>p" + using nonzero_def ord_congruence_set_nonzero + unfolding nonzero_def + by (meson Qp.nonzero_closed ord_congruence_set_nonzero subset_iff) + +lemma ord_congruence_set_memE: + assumes "x \ ord_congruence_set n a" + shows "x \ nonzero Q\<^sub>p" + "ord x mod n = a" + using assms ord_congruence_set_nonzero apply blast + by (metis (mono_tags, lifting) assms mem_Collect_eq ord_congruence_set_def) + +lemma ord_congruence_set_memI: + assumes "x \ nonzero Q\<^sub>p" + assumes "ord x mod n = a" + shows "x \ ord_congruence_set n a" + using assms + by (metis (mono_tags, lifting) mem_Collect_eq ord_congruence_set_def) + +text\ + We want to prove that ord\_congruence\_set is a finite union of semialgebraic sets, + hence is also semialgebraic. +\ + +lemma pow_res_ord_cong: + assumes "x \ carrier Q\<^sub>p" + assumes "x \ ord_congruence_set n a" + shows "pow_res n x \ ord_congruence_set n a" +proof fix y + assume A: "y \ pow_res n x" + show "y \ ord_congruence_set (int n) a" + proof- + obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]n)" + using A pow_res_def[of n x] by blast + have 0: "x \ nonzero Q\<^sub>p" + using assms(2) ord_congruence_set_memE(1) + by blast + have 1: "y \ nonzero Q\<^sub>p" + using A + by (metis "0" Qp.integral Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero a_def not_nonzero_Qp) + have 2: "ord y = ord x + n* ord a" + using a_def 0 1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_mult + by presburger + show ?thesis + apply(rule ord_congruence_set_memI) + using assms ord_congruence_set_memE 2 1 + apply blast + using "2" assms(2) ord_congruence_set_memE(2) + by presburger + qed +qed + +lemma pow_res_classes_are_univ_semialgebraic: + shows "are_univ_semialgebraic (pow_res_classes n)" + apply(rule are_univ_semialgebraicI) + using pow_res_classes_univ_semialg by blast + +lemma ord_congruence_set_univ_semialg: + assumes "n \ 0" + shows "is_univ_semialgebraic (ord_congruence_set n a)" +proof(cases "n = 0") + case True + have T0: "ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x = a}" + unfolding ord_congruence_set_def True by presburger + have T1: "{x \ nonzero Q\<^sub>p. ord x = a} = {x \ nonzero Q\<^sub>p. val x = a}" + apply(rule equalityI'') + using val_ord apply blast + using val_ord + by (metis eint.inject) + have T2: "{x \ nonzero Q\<^sub>p. val x = a} = {x \ carrier Q\<^sub>p. val x = a}" + apply(rule equalityI'') + using Qp.nonzero_closed apply blast + by (metis iless_Suc_eq val_nonzero val_val_ring_prod zero_in_val_ring) + show ?thesis unfolding T0 T1 T2 using univ_val_eq_set_is_univ_semialgebraic by blast +next + case False + obtain F where F_def: "F = {S \ (pow_res_classes (nat n)). S \(ord_congruence_set n a) }" + by blast + have 0: "F \ pow_res_classes (nat n)" + using F_def by blast + have 1: "finite F" + using 0 False nat_mono[of 1 n] nat_numeral[] pow_res_classes_finite[of "nat n"] rev_finite_subset + by (smt assms nat_one_as_int) + have 2: "are_univ_semialgebraic F" + apply(rule are_univ_semialgebraicI) using 0 pow_res_classes_are_univ_semialgebraic + by (metis (mono_tags) are_univ_semialgebraicE are_univ_semialgebraic_def assms nat_mono nat_numeral subset_iff) + have 3: "\ F = (ord_congruence_set n a)" + proof + show "\ F \ ord_congruence_set n a" + using F_def + by blast + show "ord_congruence_set n a \ \ F" + proof fix x + assume A: "x \ ord_congruence_set n a" + have x_nonzero: "x \ nonzero Q\<^sub>p" + using A ord_congruence_set_memE(1) by blast + have 0: "pow_res (nat n) x \ F" + using A pow_res_classes_def F_def + by (smt nonzero_def assms mem_Collect_eq nat_0_le ord_congruence_set_memE(1) pow_res_ord_cong) + have 1: "x \ pow_res (nat n) x" using False x_nonzero assms pow_res_refl[of x "nat n"] + using Qp.nonzero_closed by blast + show "x \ \ F" + using 0 1 + by blast + qed + qed + then show ?thesis + using "1" "2" finite_union_is_univ_semialgebraic' + by fastforce +qed + +lemma ord_congruence_set_is_semialg: + assumes "n \ 0" + shows "is_semialgebraic 1 (Qp_to_R1_set (ord_congruence_set n a))" + using assms is_univ_semialgebraicE ord_congruence_set_univ_semialg + by blast + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Congruence Sets for the order of the Evaluation of a Polynomial\ + (********************************************************************************************) + (********************************************************************************************) + + +lemma poly_map_singleton: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "poly_map n [f] x = [(Qp_ev f x)]" + unfolding poly_map_def poly_tuple_eval_def + using assms + by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') + +definition poly_cong_set where +"poly_cong_set n f m a = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). (Qp_ev f x) \ \ \ (ord (Qp_ev f x) mod m = a)}" + +lemma poly_cong_set_as_pullback: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "poly_cong_set n f m a = poly_map n [f] \\<^bsub>n\<^esub>(Qp_to_R1_set (ord_congruence_set m a))" +proof + show "poly_cong_set n f m a \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" + proof fix x + assume A: "x \ poly_cong_set n f m a" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (metis (no_types, lifting) mem_Collect_eq poly_cong_set_def) + have 1: "(Qp_ev f x) \ \ " + by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) + have 2: "(ord (Qp_ev f x) mod m = a)" + by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) + have 3: "(Qp_ev f x) \ (ord_congruence_set m a)" + using "0" "1" "2" eval_at_point_closed assms not_nonzero_Qp ord_congruence_set_memI + by metis + show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" + proof- + have 00: "poly_map n [f] x = [(Qp_ev f x)]" + using "0" assms poly_map_singleton by blast + have 01: "[eval_at_point Q\<^sub>p x f] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using "0" assms eval_at_point_closed Qp.to_R1_closed by blast + hence 02: "poly_map n [f] x \ (\a. [a]) ` ord_congruence_set m a" + using 3 "00" by blast + then show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" + using 0 unfolding evimage_def + by blast + qed + qed + show "poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a + \ poly_cong_set n f m a" + proof fix x + assume A: "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` (ord_congruence_set m a))" + have 0: "((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using ord_congruence_set_closed Qp.to_R1_carrier by blast + have "is_poly_tuple n [f]" + using assms unfolding is_poly_tuple_def + by (simp add: assms) + then have 1:"poly_map n [f] \\<^bsub>n\<^esub>((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using 0 A assms One_nat_def + by (metis extensional_vimage_closed) + then have 2: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A unfolding evimage_def by blast + then have 3: "poly_map n [f] x \ ((\a. [a]) ` ord_congruence_set m a)" + using A assms 0 One_nat_def + by blast + have "poly_map n [f] x = [(Qp_ev f x)]" + using "2" assms poly_map_singleton by blast + then have "Qp_ev f x \ ord_congruence_set m a" + using 3 + by (metis (mono_tags, lifting) image_iff list.inject) + then show "x \ poly_cong_set n f m a" + unfolding poly_cong_set_def + by (metis (mono_tags, lifting) "2" Qp.nonzero_memE(2) + mem_Collect_eq ord_congruence_set_memE(1) ord_congruence_set_memE(2)) + qed +qed + +lemma singleton_poly_tuple: +"is_poly_tuple n [f] \ f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + unfolding is_poly_tuple_def + by (metis (no_types, lifting) list.distinct(1) list.set_cases list.set_intros(1) set_ConsD subset_code(1)) + +lemma poly_cong_set_is_semialgebraic: + assumes "m \ 0" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialgebraic n (poly_cong_set n f m a)" +proof- + have 0: "(\a. [a]) ` ord_congruence_set m a \ semialg_sets 1" + using assms + ord_congruence_set_is_semialg[of m a] + unfolding is_semialgebraic_def + by blast + have 1: "length [f] = 1" + by simp + hence " poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a \ semialg_sets n" + using 0 singleton_poly_tuple[of n f] zero_neq_one assms + pullback_is_semialg[of n "[f]" 1 "(\a. [a]) ` ord_congruence_set m a"] + unfolding is_semialgebraic_def + by blast + thus ?thesis using assms poly_cong_set_as_pullback[of f n m a] + unfolding is_semialgebraic_def + by presburger +qed + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Congruence Sets for Angular Components\ + (********************************************************************************************) + (********************************************************************************************) + + +text\If a set is a union of \n\-th power residues, then it is semialgebraic.\ + +lemma pow_res_union_imp_semialg: + assumes "n \ 1" + assumes "S \ nonzero Q\<^sub>p" + assumes "\x. x \ S \ pow_res n x \ S" + shows "is_univ_semialgebraic S" +proof- + obtain F where F_def: "F = {T. T \ pow_res_classes n \ T \ S}" + by blast + have 0: "F \ pow_res_classes n" + using F_def by blast + have 1: "finite F" + using 0 pow_res_classes_finite[of n] assms(1) finite_subset + by auto + have 2: "are_univ_semialgebraic F" + using 0 + by (meson are_univ_semialgebraicE are_univ_semialgebraicI assms(1) + pow_res_classes_are_univ_semialgebraic padic_fields_axioms subsetD) + have 3: "S = \ F" + proof + show "S \ \ F" + proof fix x + assume A: "x \ S" + then have "pow_res n x \ S" + using assms(3) by blast + then have "pow_res n x \ F" + using A assms(2) F_def pow_res_classes_def + by (smt mem_Collect_eq subsetD) + then have "pow_res n x \ \ F" + by blast + then show "x \ \ F" + using A assms(1) assms(2) pow_res_refl[of x n] unfolding nonzero_def by blast + qed + show "\ F \ S" + using F_def + by blast + qed + show ?thesis + using 1 2 3 finite_union_is_univ_semialgebraic' + by blast +qed + +definition ac_cong_set1 where +"ac_cong_set1 n y = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = ac n y}" + +lemma ac_cong_set1_is_univ_semialg: + assumes "n > 0" + assumes "b \ nonzero Q\<^sub>p" + assumes "b \ \\<^sub>p" + shows "is_univ_semialgebraic (ac_cong_set1 n b)" +proof(cases "n = 1 \ p = 2") + case True + have "(ac_cong_set1 n b) = nonzero Q\<^sub>p" + proof + have 0: "Units (Zp_res_ring n) = {1}" + proof show "Units (Zp_res_ring n) \ {1}" + proof fix x assume A: "x \ Units (Zp_res_ring n)" + have 0: "carrier (Zp_res_ring n) = {0..(int 2) - 1}" + using True + by (metis assms(1) int_ops(3) p_residues power_one_right residues.res_carrier_eq) + have 1: "carrier (Zp_res_ring n) = {0..(1::int)}" + proof- have "int 2 - 1 = (1::int)" + by linarith + then show ?thesis + using 0 + by presburger + qed + have 15: "{0..(1::int)} = {0, (1::int)}" + using atLeastAtMostPlus1_int_conv [of 0 "0::int"] + by (smt atLeastAtMost_singleton insert_commute) + have 2: "carrier (Zp_res_ring n) = {0,(1::int)}" + using "1" "15" + by blast + have 3: "0 \ Units (Zp_res_ring n)" + using True zero_not_in_residue_units by blast + have "x \ carrier (Zp_res_ring n)" + using A unfolding Units_def by blast + then have "x = 1" using A 2 3 + by (metis "1" atLeastAtMost_iff atLeastatMost_empty + atLeastatMost_empty_iff2 linorder_neqE_linordered_idom mod_by_1 mod_pos_pos_trivial ) + then show "x \ {1}" + by simp + qed + show "{1} \ Units (Zp_res_ring n)" + by (meson assms(1) empty_subsetI insert_subset residue_1_unit(1)) + qed + show "ac_cong_set1 n b \ nonzero Q\<^sub>p" + by (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) + show "nonzero Q\<^sub>p \ ac_cong_set1 n b" + proof fix x + assume A: "x \ nonzero Q\<^sub>p" + then have P0: "ac n x = 1" + using 0 ac_units assms(1) by blast + have P1: "ac n b = 1" + using assms 0 ac_units assms(1) by blast + then have "ac n x = ac n b" + using P0 by metis + then show " x \ ac_cong_set1 n b" + unfolding ac_cong_set1_def using A + proof - + have "x \ {r \ carrier Q\<^sub>p. r \ \}" + by (metis (no_types) \x \ nonzero Q\<^sub>p\ nonzero_def ) + then show "x \ {r \ carrier Q\<^sub>p. r \ \ \ ac n r = ac n b}" + using \ac n x = ac n b\ by force + qed + qed + qed + then show "is_univ_semialgebraic (ac_cong_set1 n b)" + by (simp add: nonzero_is_univ_semialgebraic) +next + case F: False + have F0: "2 \ card (Units (Zp_res_ring n))" + proof(cases "n = 1") + case True + then have "field (Zp_res_ring n)" + using p_res_ring_1_field by blast + then have F00: "Units (Zp_res_ring n) = carrier (Zp_res_ring n) - {\\<^bsub>Zp_res_ring n\<^esub>}" + using field.field_Units by blast + have F01: "\\<^bsub>Zp_res_ring n\<^esub> \ carrier (Zp_res_ring n)" + using assms(1) cring.cring_simprules(2) padic_integers.R_cring padic_integers_axioms by blast + have F02: "card (carrier (Zp_res_ring n)) = p \ finite (carrier (Zp_res_ring n))" + by (smt F01 True nat_eq_iff2 p_res_ring_zero p_residue_ring_car_memE(1) power_one_right residue_ring_card) + have F03: "\\<^bsub>residue_ring (p ^ n)\<^esub> \ carrier (residue_ring (p ^ n)) " + using F01 by blast + have F04: "int (card (carrier (residue_ring (p ^ n)))) \ int (card {\\<^bsub>residue_ring (p ^ n)\<^esub>}) " + by (smt F02 F03 nat_int of_nat_0_le_iff of_nat_1 of_nat_power p_res_ring_0 p_res_ring_zero + p_residue_ring_car_memE(1) power_increasing power_one_right residue_ring_card) + have "card (carrier (residue_ring (p ^ n))) - 1 = p - 1" + using F02 prime + by (metis Totient.of_nat_eq_1_iff True less_imp_le_nat less_one nat_int nat_less_eq_zless + of_nat_1 of_nat_diff of_nat_zero_less_power_iff p_residues pos_int_cases + power_0 power_one_right residue_ring_card residues.m_gt_one zero_le_one) + hence F05: "card (carrier (residue_ring (p ^ n)) - {\\<^bsub>residue_ring (p ^ n)\<^esub>}) = p - 1" + using F02 F03 F04 card_Diff_singleton_if[of "(carrier (Zp_res_ring n))" "\\<^bsub>residue_ring (p^n)\<^esub>"] + True int_ops(6)[of "card (carrier (residue_ring (p ^ n)))" "card {\\<^bsub>residue_ring (p ^ n)\<^esub>}"] + p_res_ring_zero p_residue_ring_car_memE(1) + by (metis) + hence F06: "card (Units (Zp_res_ring n)) = p -1" + using True F02 F01 F00 + by (metis p_res_ring_zero) + have F04: "p - 1 \2 " + using F prime + by (meson True linorder_cases not_less prime_ge_2_int zle_diff1_eq) + then show ?thesis + using F03 F06 + by linarith + next + case False + then show ?thesis + by (metis assms(1) less_imp_le_nat mod2_gr_0 mod_less nat_le_linear nat_neq_iff residue_units_card_geq_2) + qed + show ?thesis + apply(rule pow_res_union_imp_semialg[of "card (Units (Zp_res_ring n))"]) + using F0 assms apply linarith + apply (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) +proof- + fix x + assume AA: "x \ ac_cong_set1 n b" + show "pow_res (card (Units (Zp_res_ring n))) x \ ac_cong_set1 n b" + proof + fix y + assume A: "y \ pow_res (card (Units (Zp_res_ring n))) x" + show "y \ ac_cong_set1 n b" + proof- + obtain k where k_def: "k = card (Units (Zp_res_ring n))" + by blast + have "k \2" + using assms k_def F F0 by blast + then obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]k)" + using k_def A pow_res_def[of k x] + by blast + have 0: "x \ nonzero Q\<^sub>p" + using AA ac_cong_set1_def + by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) + have 1: "y \ nonzero Q\<^sub>p" + by (metis "0" Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero \\thesis. (\a. a \ nonzero Q\<^sub>p \ y = x \ a [^] k \ thesis) \ thesis\) + have "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)" + using a_def 0 1 Qp_nat_pow_nonzero ac_mult' + by blast + then have 2: "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" + proof- + have "ac n (a[^]k) = ac n a [^]\<^bsub>Zp_res_ring n\<^esub> k" + using a_def assms(1) ac_nat_pow'[of a n k] + by linarith + then show ?thesis + using \ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)\ + by presburger + qed + then have "ac n y = ac n x" + proof- + have "(ac n a) \ Units (Zp_res_ring n)" + by (metis (mono_tags, opaque_lifting) a_def ac_units assms(1) ) + then have "(ac n a)^k mod (p^n) = 1" + using k_def a_def ac_nat_pow ac_nat_pow' assms(1) residue_units_nilpotent + using neq0_conv by presburger + then have 00: "(ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k = 1" + by (metis a_def ac_nat_pow ac_nat_pow' mod_by_1 power_0 + zero_neq_one) + have "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x \\<^bsub>Zp_res_ring n\<^esub> \\<^bsub>Zp_res_ring n\<^esub>" + using 00 assms(1) p_res_ring_one by presburger + hence "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x" + by (metis "0" Qp.nonzero_closed Qp.one_nonzero Qp.r_one ac_mult' ac_one' assms(1)) + then show ?thesis + using 2 "0" 00 + by linarith + qed + then show ?thesis + using "1" AA nonzero_def + ac_cong_set1_def[of n b] mem_Collect_eq + by smt + qed + qed +qed +qed + +definition ac_cong_set where +"ac_cong_set n k = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" + +lemma ac_cong_set_is_univ_semialg: + assumes "n >0 " + assumes "k \ Units (Zp_res_ring n)" + shows "is_univ_semialgebraic (ac_cong_set n k)" +proof- + have "k \ carrier (Zp_res_ring n)" + using assms(2) Units_def[of "Zp_res_ring n"] + by blast + then have k_n: "([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) n = k" + using assms + by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) + obtain b where b_def: "b = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + by blast + have 0: "k mod p \ 0" + using assms residue_UnitsE[of n k] + by (metis le_eq_less_or_eq le_refl less_one nat_le_linear p_residues power_0 + power_one_right residues.mod_in_res_units residues_def zero_less_one + zero_neq_one zero_not_in_residue_units zero_power) + then have "val_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = 0" + using val_Zp_p_int_unit by blast + then have 1: "val b = 0" + by (metis Zp_int_inc_closed b_def val_of_inc) + have 2: "b \ \\<^sub>p" + using b_def Zp_int_mult_closed + by blast + have "ord_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0" + using 0 ord_Zp_p_int_unit by blast + have "ac_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using "0" Zp_int_inc_closed ac_Zp_of_Unit ord_Zp_p_int_unit \val_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0\ + by blast + then have "(angular_component b) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using b_def 1 2 angular_component_ord_zero[of b] + by (metis Qp.int_inc_zero Qp.one_closed val_ring_memE Zp.int_inc_zero Zp.one_closed + Zp.one_nonzero Zp_int_inc_closed angular_component_of_inclusion inc_closed inc_of_int + inc_of_one inc_to_Zp local.val_zero not_nonzero_Qp val_ineq val_one zero_in_val_ring) + then have "ac n b = k" + using ac_def[of n b] k_n + by (metis Qp_char_0_int Zp_defs(1) ac_def b_def inc_of_int inc_of_one) + then have 3: "(ac_cong_set n k) = (ac_cong_set1 n b)" + unfolding ac_cong_set_def ac_cong_set1_def + by meson + have 4: "b \ nonzero Q\<^sub>p" + using 1 2 val_nonzero + by (metis Qp.one_closed val_ring_memE Zp_def \_def local.one_neq_zero + not_nonzero_Qp padic_fields.val_ring_memE padic_fields_axioms val_ineq val_one) + then show ?thesis + using 1 2 3 assms ac_cong_set1_is_univ_semialg[of n b] val_nonzero[of b 1] + by presburger +qed + +definition val_ring_constant_ac_set where +"val_ring_constant_ac_set n k = {a \ \\<^sub>p. val a = 0 \ ac n a = k}" + +lemma val_nonzero': + assumes "a \ carrier Q\<^sub>p" + assumes "val a = eint k" + shows "a \ nonzero Q\<^sub>p" + using val_nonzero[of a "k + 1"] + by (metis Suc_ile_eq assms(1) assms(2) eint_ord_code(3) val_nonzero) + +lemma val_ord': + assumes "a \ carrier Q\<^sub>p" + assumes "a \\" + shows "val a = ord a" + by (meson assms(1) assms(2) not_nonzero_Qp val_ord) + +lemma val_ring_constant_ac_set_is_univ_semialgebraic: + assumes "n > 0" + assumes "k \ 0" + shows "is_univ_semialgebraic (val_ring_constant_ac_set n k)" +proof(cases "val_ring_constant_ac_set n k = {}") + case True + then show ?thesis + by (metis equals0D order_refl pow_res_union_imp_semialg subsetI) +next + case False + then obtain b where b_def: "b \ val_ring_constant_ac_set n k" + by blast + have 0: "val_ring_constant_ac_set n k = q_ball n k 0 \" + proof + show "val_ring_constant_ac_set n k \ q_ball n k 0 \" + proof fix x assume A: "x \ val_ring_constant_ac_set n k" then + show "x \ q_ball n k 0 \" + proof- + have 0: "x \ \\<^sub>p \ val x = 0 \ ac n x = k" + using A + unfolding val_ring_constant_ac_set_def + by blast + then have x_car: "x \ carrier Q\<^sub>p" + using val_ring_memE + by blast + then have 00: "x = x \ \" + using Qp.ring_simprules by metis + then have 1: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" + using 0 + by presburger + have 2: "val (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" + using 0 00 + by metis + have 3: "x \ nonzero Q\<^sub>p" + proof(rule ccontr) + assume " x \ nonzero Q\<^sub>p " + then have "x = \" + using Qp.nonzero_memI x_car by blast + then show False + using 0 val_zero + by (metis ac_def assms(2)) + qed + have 4: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" + proof(rule ccontr) + assume "ord (x \ \) \ 0" + then have "val (x \ \) \ 0" + by (metis "00" "3" Qp.one_closed equal_val_imp_equal_ord(1) ord_one val_one) + then show False + using "2" + by blast + qed + show ?thesis + using 0 1 4 + unfolding q_ball_def + using x_car by blast + qed + qed + show "q_ball n k 0 \ \ val_ring_constant_ac_set n k" + proof fix x + assume A: "x \ q_ball n k 0 \" + then have 0: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" + using q_ballE'(1) by blast + have 1: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" + using q_ball_def A + by blast + have 2: "x \ carrier Q\<^sub>p" + using A q_ball_def by blast + have 3: "ord x = 0" + using 2 1 ring.ring_simprules[of Q\<^sub>p] + by (metis Qp.ring_axioms) + have 4: "ac n x = k" + using 0 2 1 cring.axioms(1)[of Q\<^sub>p] ring.ring_simprules[of Q\<^sub>p] + by (metis Qp.ring_axioms) + have 5: "x \ \\<^sub>p" + using Qp_val_ringI[of x] 2 3 val_ord val_nonzero' + by (metis Qp.integral_iff val_ring_memE Zp.nonzero_closed angular_component_closed + angular_component_ord_zero image_eqI local.numer_denom_facts(1) local.numer_denom_facts(2) + local.numer_denom_facts(4) not_nonzero_Qp) + have 6: "x \ \" + using 4 assms ac_def[of n x] + by meson + have 7: "val x = 0" + using 6 3 2 assms val_ord' zero_eint_def by presburger + show " x \ val_ring_constant_ac_set n k" + unfolding val_ring_constant_ac_set_def + using 7 6 5 4 + by blast + qed + qed + obtain b where b_def: "b \ q_ball n k (0::int) \" + using "0" b_def by blast + have 1: "b \ carrier Q\<^sub>p \ ac n b = k" + using b_def unfolding q_ball_def + by (metis (mono_tags, lifting) "0" b_def mem_Collect_eq val_ring_constant_ac_set_def) + then have 2: "b \ nonzero Q\<^sub>p" + using 1 assms + by (metis ac_def not_nonzero_Qp) + have "q_ball n k 0 \ = B\<^bsub>0 + int n\<^esub>[b]" + using 1 b_def nonzero_def [of Q\<^sub>p] assms 0 2 c_ball_q_ball[of b n k "\" b 0] + by (meson Qp.cring_axioms cring.cring_simprules(2)) + then have "is_univ_semialgebraic (q_ball n k (0::int) \) " + using 1 ball_is_univ_semialgebraic[of b "0 + int n"] + by metis + then show ?thesis + using 0 by presburger +qed + +definition val_ring_constant_ac_sets where +"val_ring_constant_ac_sets n = val_ring_constant_ac_set n ` (Units (Zp_res_ring n))" + +lemma val_ring_constant_ac_sets_are_univ_semialgebraic: + assumes "n > 0" + shows "are_univ_semialgebraic (val_ring_constant_ac_sets n)" +proof(rule are_univ_semialgebraicI) + have 0: "\ coprime 0 p" + using coprime_0_right_iff[of p] coprime_commute[of p 0] coprime_int_iff[of "nat p" 0] + nat_dvd_1_iff_1 prime_gt_1_nat zdvd1_eq + by (metis not_prime_unit prime) + have "(0::int) \(Units (Zp_res_ring n))" + apply(rule ccontr) + using 0 assms residues.cring[of "p ^ n"] unfolding residues_def + by (smt less_one not_gr_zero power_le_imp_le_exp power_less_imp_less_exp residue_UnitsE) + fix x + assume A: "x \ val_ring_constant_ac_sets n" + then obtain k where k_def: "x = val_ring_constant_ac_set n k \ k \ Units (Zp_res_ring n)" + by (metis image_iff val_ring_constant_ac_sets_def) + then show "is_univ_semialgebraic x" + using assms + by (metis \0 \ Units (Zp_res_ring n)\ val_ring_constant_ac_set_is_univ_semialgebraic) +qed + +definition ac_cong_set3 where +"ac_cong_set3 n = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = ac n b) \ as = [a, b] }" + +definition ac_cong_set2 where +"ac_cong_set2 n k = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k \ as = [a, b] }" + +lemma ac_cong_set2_cartesian_product: + assumes "k \ Units (Zp_res_ring n)" + assumes "n > 0" + shows "ac_cong_set2 n k = cartesian_product (to_R1` (ac_cong_set n k)) (to_R1` (val_ring_constant_ac_set n k))" +proof + show "ac_cong_set2 n k \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" + proof fix x + assume A: "x \ ac_cong_set2 n k" + show "x \ (cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k))" + unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def + apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 1]) + apply (metis (mono_tags, lifting) mem_Collect_eq subsetI Qp.to_R1_car_subset) + apply (metis (no_types, lifting) val_ring_memE mem_Collect_eq subsetI Qp.to_R1_car_subset) + proof- + obtain a b where ab_def: "x = [a,b] \ a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k" + using A + unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def + by blast + have 0: "take 1 x = [a]" + by (simp add: ab_def) + have 1: "drop 1 x = [b]" + by (simp add: ab_def) + have 2: "a \ {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" + using ab_def nonzero_def + by (smt mem_Collect_eq) + have 3: "b \ {a \ \\<^sub>p. val a = 0 \ ac n a = k}" + using ab_def + by blast + show "take 1 x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" + using 0 2 by blast + show "drop 1 x \ (\a. [a]) ` {a \ \\<^sub>p. val a = 0 \ ac n a = k}" + using 1 3 by blast + qed + qed + show "cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k) \ ac_cong_set2 n k" + proof fix x + have 0: "(\a. [a]) ` ac_cong_set n k \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using assms + by (metis (no_types, lifting) ac_cong_set_def mem_Collect_eq subsetI Qp.to_R1_car_subset) + have 1: "((\a. [a]) ` val_ring_constant_ac_set n k) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + by (smt val_ring_memE mem_Collect_eq subsetI Qp.to_R1_carrier Qp.to_R1_subset val_ring_constant_ac_set_def) + assume A: "x \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" + then have "length x = 2" + using 0 1 A cartesian_product_closed[of "((\a. [a]) ` ac_cong_set n k)" Q\<^sub>p 1 "((\a. [a]) ` val_ring_constant_ac_set n k)" 1] + by (metis (no_types, lifting) cartesian_power_car_memE one_add_one subset_iff) + then obtain a b where ab_def: "take 1 x = [a] \ drop 1 x = [b]" + by (metis One_nat_def add_diff_cancel_left' drop0 drop_Cons_numeral numerals(1) pair_id plus_1_eq_Suc take0 take_Cons_numeral) + have 2: " a \ (ac_cong_set n k) \ b \ (val_ring_constant_ac_set n k)" + proof- + have P0: "take 1 x \ (\a. [a]) ` ac_cong_set n k" + using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] + by blast + have P1: "drop 1 x \ (\a. [a]) ` val_ring_constant_ac_set n k" + using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] + by blast + have P2: "[a] \ (\a. [a]) ` ac_cong_set n k" + using P0 ab_def + by metis + have P3: "[b] \ (\a. [a]) ` val_ring_constant_ac_set n k" + using P1 ab_def by metis + show ?thesis + using P2 P3 + by blast + qed + have 3: "a \ nonzero Q\<^sub>p" + using 2 assms nonzero_def [of Q\<^sub>p] ac_cong_set_def[of n k] + by blast + have 4: "x = [a,b]" + by (metis (no_types, lifting) \length x = 2\ ab_def less_numeral_extra(1) nth_Cons_0 nth_take nth_via_drop pair_id) + then have "\a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ ac n a = k \ ac n b = k \ x = [a, b]" + using 2 3 ab_def unfolding val_ring_constant_ac_set_def ac_cong_set_def + by blast + then show "x \ ac_cong_set2 n k" + unfolding ac_cong_set2_def val_ring_constant_ac_set_def ac_cong_set_def + by blast + qed +qed + +lemma ac_cong_set2_is_semialg: + assumes "k \ Units (Zp_res_ring n)" + assumes "n > 0" + shows "is_semialgebraic 2 (ac_cong_set2 n k)" + using ac_cong_set_is_univ_semialg ac_cong_set2_cartesian_product[of k n] + cartesian_product_is_semialgebraic[of 1 "((\a. [a]) ` ac_cong_set n k)" 1 " ((\a. [a]) ` val_ring_constant_ac_set n k)"] + by (metis assms(1) assms(2) is_univ_semialgebraicE less_one less_or_eq_imp_le nat_neq_iff + one_add_one val_ring_constant_ac_set_is_univ_semialgebraic zero_not_in_residue_units) + +lemma ac_cong_set3_as_union: + assumes "n > 0" + shows "ac_cong_set3 n = \ (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" +proof + show "ac_cong_set3 n \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" + proof fix x assume A: "x \ ac_cong_set3 n" + then have 0: "x \ (ac_cong_set2 n (ac n (x!0)))" + unfolding ac_cong_set2_def ac_cong_set3_def + by (smt mem_Collect_eq nth_Cons_0) + have 1: "(ac n (x!0)) \ Units (Zp_res_ring n)" + using A unfolding ac_cong_set3_def + by (smt ac_units assms mem_Collect_eq nth_Cons_0) + then show "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" + using 0 + by blast + qed + show "\ (ac_cong_set2 n ` Units (Zp_res_ring n)) \ ac_cong_set3 n" + proof fix x assume A: "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" + obtain k where k_def: "x \ (ac_cong_set2 n k) \ k \ (Units (Zp_res_ring n))" + using A by blast + have 0: "k mod p \ 0" + using k_def One_nat_def Suc_le_eq assms less_numeral_extra(1) + power_one_right residues.m_gt_one residues.mod_in_res_units + by (metis p_residues residue_UnitsE zero_not_in_residue_units) + obtain b where b_def: "b = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + by blast + have "k \0" + using 0 mod_0 + by blast + then have 1: "b \ nonzero Z\<^sub>p" + using 0 b_def int_unit + by (metis Zp.Units_nonzero Zp.zero_not_one) + have 10: "ord_Zp b = 0" using 0 1 + using b_def ord_Zp_p_int_unit by blast + have 2: "\ b \ nonzero Q\<^sub>p" using k_def + using "1" inc_of_nonzero by blast + have 3: "angular_component (\ b) = ac_Zp b" + using "1" angular_component_of_inclusion + by blast + have 4: "ac_Zp b = b" + using 1 10 + by (metis "3" Zp.r_one ac_Zp_factors' angular_component_closed inc_of_nonzero int_pow_0 mult_comm ord_Zp_def) + have 5: "ac_Zp b n = k" + proof- + have "k \ carrier (Zp_res_ring n)" + using k_def unfolding Units_def by blast + then show ?thesis + using b_def k_def 4 Zp_int_inc_res mod_pos_pos_trivial + by (metis p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) + qed + then have "ac n (\ b) = k" + using 10 1 2 3 4 unfolding ac_def + using Qp.not_nonzero_memI by metis + then show "x \ ac_cong_set3 n" + unfolding ac_cong_set3_def + using k_def unfolding ac_cong_set2_def + by (smt mem_Collect_eq) + qed +qed + +lemma ac_cong_set3_is_semialgebraic: + assumes "n > 0" + shows "is_semialgebraic 2 (ac_cong_set3 n)" +proof- + have 0: "finite (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" + using assms residues.finite_Units[of "p^n"] unfolding residues_def + using p_residues residues.finite_Units by blast + have 1: "are_semialgebraic 2 (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" + apply(rule are_semialgebraicI) + using ac_cong_set2_is_semialg assms by blast + show ?thesis + using 0 1 ac_cong_set3_as_union + by (metis (no_types, lifting) are_semialgebraicE assms finite_union_is_semialgebraic' is_semialgebraicE subsetI) +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Permutations of indices of semialgebraic sets\ +(**************************************************************************************************) +(**************************************************************************************************) + + +lemma fun_inv_permute: + assumes "\ permutes {.. permutes {.. \ (fun_inv \) = id" + "(fun_inv \) \ \ = id" + using assms unfolding fun_inv_def + using permutes_inv apply blast + using assms permutes_inv_o(1) apply blast + using assms permutes_inv_o(2) by blast + +lemma poly_tuple_pullback_eq_poly_map_vimage: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "S \ carrir (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "poly_map n fs \\<^bsub>n\<^esub> S = poly_tuple_pullback n S fs" + unfolding poly_map_def poly_tuple_pullback_def evimage_def restrict_def + using assms + by (smt vimage_inter_cong) + +lemma permutation_is_semialgebraic: + assumes "is_semialgebraic n S" + assumes "\ permutes {.. ` S)" +proof- + have "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms gen_boolean_algebra_subset is_semialgebraic_def semialg_sets_def + by blast + then have "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" + using Qp.cring_axioms assms pullback_by_permutation_of_poly_list'[of \ n S] unfolding poly_map_def + by blast + then have 0: "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" + using poly_tuple_pullback_def + by blast + have 1: "(fun_inv \) permutes {..) (pvar_list Q\<^sub>p n))"] + permutation_of_poly_list_is_poly_list[of n "(pvar_list Q\<^sub>p n)" "fun_inv \"] + pvar_list_is_poly_tuple[of n] assms poly_tuple_pullback_eq_poly_map_vimage + by (metis "0" \S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ is_semialgebraic_def length_permute_list pvar_list_length) +qed + +lemma permute_list_closed: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms cartesian_power_car_memE length_permute_list apply blast + using assms cartesian_power_car_memE'' permute_list_set by blast + +lemma permute_list_closed': + assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (metis assms(2) cartesian_power_car_memE length_permute_list) + using assms cartesian_power_car_memE'[of "permute_list \ a" Q\<^sub>p n] + by (metis cartesian_power_car_memE in_set_conv_nth length_permute_list set_permute_list subsetI) + +lemma permute_list_compose_inv: + assumes "\ permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "permute_list \ (permute_list (fun_inv \) a) = a" + "permute_list (fun_inv \) (permute_list \ a) = a" + using assms apply (metis cartesian_power_car_memE fun_inv_permute(3) permute_list_compose permute_list_id) + using assms by (metis cartesian_power_car_memE fun_inv_permute(2) fun_inv_permute(1) permute_list_compose permute_list_id) + +lemma permutation_is_semialgebraic_imp_is_semialgebraic: + assumes "is_semialgebraic n (permute_list \ ` S)" + assumes "\ permutes {..) ` (permute_list \ ` S) = S" + proof- + have 0: "(permute_list \ ` S) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms unfolding is_semialgebraic_def semialg_sets_def + using gen_boolean_algebra_subset by blast + have 1: "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + proof fix x assume "x \ S" then show "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using 0 assms + by (meson image_subset_iff permute_list_closed') + qed + show ?thesis + proof show "permute_list (fun_inv \) ` permute_list \ ` S \ S" + using 0 assms permute_list_compose_inv[of \] "1" image_iff image_subset_iff subsetD + by smt + show "S \ permute_list (fun_inv \) ` permute_list \ ` S" + using 0 assms permute_list_compose_inv[of \] + by (smt "1" image_iff subset_eq) + qed + qed + then show ?thesis using permutation_is_semialgebraic + by (metis assms(1) assms(2) fun_inv_permute(1)) +qed + +lemma split_cartesian_product_is_semialgebraic: + assumes "i \ n" + assumes "is_semialgebraic n A" + assumes "is_semialgebraic m B" + shows "is_semialgebraic (n + m) (split_cartesian_product n m i A B)" + using assms cartesian_product_is_semialgebraic scp_permutes[of i n m] + permutation_is_semialgebraic[of "n + m" "cartesian_product A B" "(scp_permutation n m i)"] + unfolding split_cartesian_product_def + by blast + +definition reverse_val_relation_set where +"reverse_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)}" + +lemma Qp_2_car_memE: + assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + shows "x = [x!0, x!1]" +proof- + have "length x = 2" + using assms cartesian_power_car_memE by blast + then show ?thesis + using pair_id by blast +qed + +definition flip where +"flip = (\i::nat. (if i = 0 then 1 else (if i = 1 then 0 else i)))" + +lemma flip_permutes: +"flip permutes {0,1}" + unfolding permutes_def flip_def + by (smt mem_simps(1)) + +lemma flip_eval: +"flip 0 = 1" +"flip 1 = 0" + unfolding flip_def + by auto + +lemma flip_x: + assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + shows "permute_list flip x = [x!1, x!0]" +proof- + have 0: "x = [x!0, x!1]" + using assms Qp_2_car_memE by blast + have 1: "length (permute_list flip x) = length [x!1, x!0]" + using 0 unfolding permute_list_def + by (metis length_Cons length_map map_nth) + have 2: "\i. i < 2 \ permute_list flip x ! i = [x!1, x!0] ! i" + proof- fix i::nat assume A: "i < 2" + show "permute_list flip x ! i = [x!1, x!0] ! i" + using 0 unfolding permute_list_def + by (smt flip_eval(1) flip_eval(2) length_Cons length_greater_0_conv list.simps(8) map_upt_Suc numeral_nat(7) upt_rec) + qed + have "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" + proof- + have 0: "length x = 2" + using assms cartesian_power_car_memE by blast + show "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" using 2 unfolding 0 + by blast + qed + thus ?thesis using 1 + by (metis length_permute_list nth_equalityI) +qed + +lemma permute_with_flip_closed: + assumes "x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" + shows "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" + apply(rule permute_list_closed) + using assms apply blast +proof- + have "{0::nat, 1} = {..<2::nat}" + by auto + thus "flip permutes {..<2}" + using flip_permutes + by auto +qed + +lemma reverse_val_relation_set_semialg: +"is_semialgebraic 2 reverse_val_relation_set" +proof- + have 1: "reverse_val_relation_set = permute_list flip ` val_relation_set" + apply(rule equalityI') + proof- + show " \x. x \ reverse_val_relation_set \ x \ permute_list flip ` val_relation_set" + proof- fix x assume A: "x \ reverse_val_relation_set" + have 0: "permute_list flip x = [x ! 1, x ! 0]" + using flip_x[of x] A unfolding reverse_val_relation_set_def + by blast + have 1: "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + apply(rule permute_with_flip_closed) using A unfolding reverse_val_relation_set_def by blast + have 2: "permute_list flip x \ val_relation_set" + using 1 A unfolding 0 reverse_val_relation_set_def val_relation_set_def mem_Collect_eq + by (metis Qp_2_car_memE list_hd list_tl) + show "x \ permute_list flip ` val_relation_set" + using flip_x[of x] A unfolding reverse_val_relation_set_def val_relation_set_def mem_Collect_eq + by (metis (no_types, lifting) "1" "2" Qp_2_car_memE flip_x image_eqI list_tl nth_Cons_0 val_relation_set_def) + qed + show "\x. x \ permute_list flip ` val_relation_set \ x \ reverse_val_relation_set" + proof- fix x assume a: " x \ permute_list flip ` val_relation_set" + then obtain y where y_def: "y \ val_relation_set \x = permute_list flip y" + by blast + have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using y_def basic_semialg_set_memE(1) val_relation_semialg by blast + have y_length: " length y = 2" + using y_def basic_semialg_set_memE val_relation_semialg + by (metis cartesian_power_car_memE) + obtain a b where ab_def: "y = [a,b]" + using y_length pair_id by blast + have 0: "a = y!0" + using ab_def + by (metis nth_Cons_0) + have 1: "b = y!1" + using ab_def + by (metis cancel_comm_monoid_add_class.diff_cancel eq_numeral_extra(2) nth_Cons') + have a_closed: "a \ carrier Q\<^sub>p" + using 0 y_closed unfolding 0 + by (meson cartesian_power_car_memE' rel_simps(75) zero_order(5)) + have b_closed: "b \ carrier Q\<^sub>p" + proof- + have "1 < (2::nat)" by linarith + thus ?thesis + using y_closed unfolding 1 + by (meson cartesian_power_car_memE') + qed + have 2: "x = [b, a]" using flip_x[of y] y_def y_closed unfolding ab_def unfolding 0 1 + using \y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ permute_list flip y = [y ! 1, y ! 0]\ y_closed y_def by presburger + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using y_def unfolding val_relation_set_def using permute_with_flip_closed[of y] + by blast + show " x \ reverse_val_relation_set" + using x_closed y_def + unfolding val_relation_set_def reverse_val_relation_set_def mem_Collect_eq 2 0 1 + by (metis Qp_2_car_memE list_hd list_tl) + qed + qed + show ?thesis unfolding 1 + apply(rule permutation_is_semialgebraic) + using val_relation_is_semialgebraic apply blast + using flip_permutes + by (metis Suc_1 insert_commute lessThan_0 lessThan_Suc numeral_nat(7)) +qed + +definition strict_val_relation_set where +"strict_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) < val (as ! 1)}" + +definition val_diag where +"val_diag = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)}" + +lemma val_diag_semialg: +"is_semialgebraic 2 val_diag" +proof- + have "val_diag = val_relation_set \reverse_val_relation_set" + apply(rule equalityI') + apply(rule IntI) + unfolding val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq + apply simp + apply simp + apply(erule IntE) unfolding mem_Collect_eq + using basic_trans_rules(24) by blast + then show ?thesis using intersection_is_semialg + by (simp add: reverse_val_relation_set_semialg val_relation_is_semialgebraic) +qed + +lemma strict_val_relation_set_is_semialg: +"is_semialgebraic 2 strict_val_relation_set" +proof- + have 0: "strict_val_relation_set = reverse_val_relation_set - val_diag" + apply(rule equalityI') + apply(rule DiffI) + unfolding strict_val_relation_set_def val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq + using order_le_less apply blast + proof + show "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1) \ x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) = val (x ! 1) \ False" + using order_less_le by blast + show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)} - {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)} \ + x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1)" + apply(erule DiffE) unfolding mem_Collect_eq using order_le_less by blast + qed + show ?thesis unfolding 0 + apply(rule diff_is_semialgebraic ) + using reverse_val_relation_set_semialg apply blast + using val_diag_semialg by blast +qed + +lemma singleton_length: + "length [a] = 1" + by auto + +lemma take_closed': + assumes "m > 0" + assumes "x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" + shows "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + apply(rule take_closed[of m "m+l"]) + apply simp using assms by blast + +lemma triple_val_ineq_set_semialg: + shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}" +proof- + have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" + proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + proof(rule equalityI') + show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!0, x!1]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!0 #[x!1]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] + by presburger + qed + obtain b where b_def: "b = [x!2]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ reverse_val_relation_set" + proof- + have 0: "a = take 2 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_take[of 2 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " + apply(cases "i = 0") + apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) + by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (simp add: a_length) + unfolding 0 using A unfolding mem_Collect_eq + using cartesian_power_car_memE' by fastforce + show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] + by (metis in_set_conv_nth le_numeral_extra(4) less_one numeral_le_iff numeral_less_iff padic_fields.singleton_length padic_fields_axioms semiring_norm(173) semiring_norm(79) Qp.to_R_to_R1) + have 2: "x = a@b" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of a b] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (a @ b) ! i" + apply(cases "i = 0") + apply (metis a_def append.simps(2) nth_Cons_0) + apply(cases "(i:: nat) = 1") + apply (simp add: a_def) + + proof- assume a: "i \0" "i \ 1" + then have "i = 2" + using A1 by presburger + thus ?thesis + by (metis a_length b_def nth_append_length) + qed + qed + have 3: "a = take 2 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_take[of 2 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " + apply(cases "i = 0") + apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) + by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) + qed + show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) + apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) + apply blast + using 3 a_closed apply blast + proof- + have "drop 2 x = b" + unfolding 2 unfolding 3 using 0 + by simp + then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using b_closed by blast + qed + qed + show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" + proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" + using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by metis + have a_length: "length a = 2" + using ab_def unfolding reverse_val_relation_set_def + using cartesian_power_car_memE by blast + have "(0::nat)< 2" by presburger + hence 0: "x!0 = a!0" + unfolding ab_def using a_length + by (metis append.simps(2) nth_Cons_0 pair_id) + have "(1::nat)< 2" by presburger + hence 1: "x!1 = a!1" + unfolding ab_def using a_length + by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed + unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq + by (metis Suc_1 eval_nat_numeral(3) plus_1_eq_Suc semiring_norm(164)) + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" + using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) + qed + have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" + proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" + proof(rule equalityI') + show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" + proof- + fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!1, x!2]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!1 #[x!2]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] + by presburger + qed + obtain b where b_def: "b = [x!0]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ reverse_val_relation_set" + proof- + have 0: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq + by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) + show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] + by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) + have 2: "x = b@a" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of b a] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (b @ a) ! i" + apply(cases "i = 0") + apply (metis append.simps(2) b_def nth_Cons_0) + apply(cases "(i:: nat) = (1::nat)") + using append.simps a_def nth_Cons + apply (metis b_length nth_append_length) + apply(cases "(i:: nat) = (2::nat)") + using A unfolding 0 + apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) + proof- assume A0: "i \0" "i \ 1" "i \2" + then have "i \ 3" by presburger + then show "x ! i = (b @ a) ! i" + using A unfolding 0 by presburger + qed + qed + have 3: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) + apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) + using reverse_val_relation_set_def apply blast + using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] + using a_closed unfolding 3 by blast + qed + show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" + proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " + then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" + using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] + by metis + have a_length: "length a = 2" + using ab_def unfolding reverse_val_relation_set_def + using cartesian_power_car_memE by blast + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have b_length: "length b = 1" + by (simp add: b'_def) + have x_id: "x = b'#a" + unfolding ab_def b'_def by auto + have "(1::nat)< 2" by presburger + hence 0: "x!1 = a!0" + unfolding ab_def b'_def using a_length + by (metis b'_def b_length nth_append_length pair_id) + have 00: "2 = Suc 1" + by auto + have 1: "x!2 = a!1" + using a_length nth_Cons[of b' a "2::nat"] + unfolding x_id 00 + by (meson nth_Cons_Suc) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def + unfolding reverse_val_relation_set_def mem_Collect_eq + using eval_nat_numeral(3) semiring_norm(175) by presburger + + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" + using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) + qed + have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}= + {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" + by blast + show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast +qed + +lemma triple_val_ineq_set_semialg': + shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}" +proof- + have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" + proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + proof(rule equalityI') + show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!0, x!1]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!0 #[x!1]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] + by presburger + qed + obtain b where b_def: "b = [x!2]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ reverse_val_relation_set" + proof- + have 0: "a = take 2 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_take[of 2 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " + apply(cases "i = 0") + apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) + by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using a_def 0 A unfolding mem_Collect_eq + by (metis le_numeral_extra(4) take_closed numeral_le_iff semiring_norm(173) semiring_norm(72)) + show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] + by (metis in_set_conv_nth le_numeral_extra(4) less_one numeral_le_iff numeral_less_iff padic_fields.singleton_length padic_fields_axioms semiring_norm(173) semiring_norm(79) Qp.to_R_to_R1) + have 2: "x = a@b" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of a b] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (a @ b) ! i" + apply(cases "i = 0") + apply (metis a_def append.simps(2) nth_Cons_0) + apply(cases "(i:: nat) = 1") + apply (simp add: a_def) + proof- assume a: "i \0" "i \ 1" + then have "i = 2" + using A1 by presburger + thus ?thesis + by (metis a_length b_def nth_append_length) + qed + qed + have 3: "a = take 2 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_take[of 2 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " + apply(cases "i = 0") + apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) + by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) + qed + show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) + apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) + apply blast + using 3 a_closed apply blast + proof- + have "drop 2 x = b" + unfolding 2 unfolding 3 using 0 + by simp + then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + using b_closed by blast + qed + qed + show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" + proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" + using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by metis + have a_length: "length a = 2" + using ab_def unfolding reverse_val_relation_set_def + using cartesian_power_car_memE by blast + have "(0::nat)< 2" by presburger + hence 0: "x!0 = a!0" + unfolding ab_def using a_length + by (metis append.simps(2) nth_Cons_0 pair_id) + have "(1::nat)< 2" by presburger + hence 1: "x!1 = a!1" + unfolding ab_def using a_length + by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed + unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq + by (metis Suc_1 eval_nat_numeral(3) plus_1_eq_Suc semiring_norm(164)) + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" + using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) + qed + have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" + proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" + proof(rule equalityI') + show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" + proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!1, x!2]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!1 #[x!2]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] + by presburger + qed + obtain b where b_def: "b = [x!0]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ strict_val_relation_set" + proof- + have 0: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq + by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) + show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] + by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) + have 2: "x = b@a" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of b a] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (b @ a) ! i" + apply(cases "i = 0") + apply (metis append.simps(2) b_def nth_Cons_0) + apply(cases "(i:: nat) = (1::nat)") + using append.simps a_def nth_Cons + apply (metis b_length nth_append_length) + apply(cases "(i:: nat) = (2::nat)") + using A unfolding 0 + apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) + proof- assume A0: "i \0" "i \ 1" "i \2" + then have "i \ 3" by presburger + then show "x ! i = (b @ a) ! i" + using A unfolding 0 by presburger + qed + qed + have 3: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) + apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) + using strict_val_relation_set_def apply blast + using take_closed[of 1 3 x Q\<^sub>p] A unfolding mem_Collect_eq apply auto[1] + using a_closed unfolding 3 by blast + qed + show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" + proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " + then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" + using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] + by metis + have a_length: "length a = 2" + using ab_def unfolding strict_val_relation_set_def + using cartesian_power_car_memE by blast + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have b_length: "length b = 1" + by (simp add: b'_def) + have x_id: "x = b'#a" + unfolding ab_def b'_def by auto + have "(1::nat)< 2" by presburger + hence 0: "x!1 = a!0" + unfolding ab_def b'_def using a_length + by (metis b'_def b_length nth_append_length pair_id) + have 00: "2 = Suc 1" + by auto + have 1: "x!2 = a!1" + using a_length nth_Cons[of b' a "2::nat"] + unfolding x_id 00 + by (meson nth_Cons_Suc) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def + unfolding strict_val_relation_set_def mem_Collect_eq + using eval_nat_numeral(3) semiring_norm(175) by presburger + + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" + using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) + qed + have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}= + {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" + by blast + show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast +qed + +lemma triple_val_ineq_set_semialg'': + shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" +proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" + proof(rule equalityI') + show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" + proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!1, x!2]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!1 #[x!2]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] + by presburger + qed + obtain b where b_def: "b = [x!0]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ strict_val_relation_set" + proof- + have 0: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq + by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) + show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] + by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) + have 2: "x = b@a" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of b a] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (b @ a) ! i" + apply(cases "i = 0") + apply (metis append.simps(2) b_def nth_Cons_0) + apply(cases "(i:: nat) = (1::nat)") + using append.simps a_def nth_Cons + apply (metis b_length nth_append_length) + apply(cases "(i:: nat) = (2::nat)") + using A unfolding 0 + apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) + proof- assume A0: "i \0" "i \ 1" "i \2" + then have "i \ 3" by presburger + then show "x ! i = (b @ a) ! i" + using A unfolding 0 by presburger + qed + qed + have 3: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) + apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) + using strict_val_relation_set_def apply blast + using take_closed[of 1 3 x] A unfolding mem_Collect_eq + using one_le_numeral apply blast + using a_closed unfolding 3 by blast + qed + show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" + proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " + then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" + using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] + by metis + have a_length: "length a = 2" + using ab_def unfolding strict_val_relation_set_def + using cartesian_power_car_memE by blast + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have b_length: "length b = 1" + by (simp add: b'_def) + have x_id: "x = b'#a" + unfolding ab_def b'_def by auto + have "(1::nat)< 2" by presburger + hence 0: "x!1 = a!0" + unfolding ab_def b'_def using a_length + by (metis b'_def b_length nth_append_length pair_id) + have 00: "2 = Suc 1" + by auto + have 1: "x!2 = a!1" + using a_length nth_Cons[of b' a "2::nat"] + unfolding x_id 00 + by (meson nth_Cons_Suc) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def + unfolding strict_val_relation_set_def mem_Collect_eq + using eval_nat_numeral(3) semiring_norm(175) by presburger + + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" + using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) +qed + +lemma triple_val_ineq_set_semialg''': + shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" +proof- + have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" + proof(rule equalityI') + show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" + proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" + then have 0: "length x = 3" unfolding mem_Collect_eq + using cartesian_power_car_memE by blast + obtain a where a_def: "a = [x!1, x!2]" + by blast + have a_length: "length a = 2" + proof- + have "a = x!1 #[x!2]" + unfolding a_def + by blast + thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] + by presburger + qed + obtain b where b_def: "b = [x!0]" + by blast + have b_length: "length b = 1" + unfolding b_def singleton_length by auto + have a_closed: "a \ reverse_val_relation_set" + proof- + have 0: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq + by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) + show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq + by (metis Qp_2_car_memE list_tl nth_Cons_0) + qed + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + unfolding b_length apply blast + apply(rule subsetI) + unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] + by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) + have 2: "x = b@a" + apply(rule nth_equalityI) + using 0 unfolding a_length b_length length_append[of b a] apply presburger + proof- fix i assume A: "i < length x" + then have A1: "i < 3" + unfolding 0 by blast + show "x ! i = (b @ a) ! i" + apply(cases "i = 0") + apply (metis append.simps(2) b_def nth_Cons_0) + apply(cases "(i:: nat) = (1::nat)") + using append.simps a_def nth_Cons + apply (metis b_length nth_append_length) + apply(cases "(i:: nat) = (2::nat)") + using A unfolding 0 + apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) + proof- assume A0: "i \0" "i \ 1" "i \2" + then have "i \ 3" by presburger + then show "x ! i = (b @ a) ! i" + using A unfolding 0 by presburger + qed + qed + have 3: "a = drop 1 x" + apply(rule nth_equalityI) + unfolding a_length 0 length_drop[of 1 x] + apply linarith + proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" + apply(cases "i = 0") + unfolding a_def using nth_drop[of 1 x i] + apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) + apply(cases "i = 1") + using nth_drop[of 1 x i] unfolding 0 + apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) + using a by presburger + qed + show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" + apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) + apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) + using reverse_val_relation_set_def apply blast + using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] + using a_closed unfolding 3 by blast + qed + show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" + proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " + then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" + using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] + by metis + have a_length: "length a = 2" + using ab_def unfolding reverse_val_relation_set_def + using cartesian_power_car_memE by blast + obtain b' where b'_def: "b = [b']" + using ab_def cartesian_power_car_memE + by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) + have b'_closed: "b' \ carrier Q\<^sub>p" + using b'_def ab_def cartesian_power_car_memE + by (metis Qp.R1_memE' list_hd) + have b_length: "length b = 1" + by (simp add: b'_def) + have x_id: "x = b'#a" + unfolding ab_def b'_def by auto + have "(1::nat)< 2" by presburger + hence 0: "x!1 = a!0" + unfolding ab_def b'_def using a_length + by (metis b'_def b_length nth_append_length pair_id) + have 00: "2 = Suc 1" + by auto + have 1: "x!2 = a!1" + using a_length nth_Cons[of b' a "2::nat"] + unfolding x_id 00 + by (meson nth_Cons_Suc) + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" + unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def + unfolding reverse_val_relation_set_def mem_Collect_eq + using eval_nat_numeral(3) semiring_norm(175) by presburger + + show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" + using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast + qed + qed + show ?thesis unfolding 0 + using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] + by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Semialgebraic Functions\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + The most natural way to define a semialgebraic function $f: \mathbb{Q}_p^n \to \mathbb{Q}_p$ is a + function whose graph is a semialgebraic subset of $\mathbb{Q}_p^{n+1}$. However, the definition + given here is slightly different, and devised by Denef in \cite{denef1986} in order to prove + Macintyre's theorem. As Denef notes, we can use Macintyre's theorem to deduce that the given + definition perfectly aligns with the intuitive one. +\ + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Defining Semialgebraic Functions\ + (********************************************************************************************) + (********************************************************************************************) + + +text\Apply a function f to the tuple consisting of the first n indices, leaving the remaining indices +unchanged\ + +definition partial_image where +"partial_image m f xs = (f (take m xs))#(drop m xs)" + +definition partial_pullback where +"partial_pullback m f l S = (partial_image m f) \\<^bsub>m+l\<^esub> S " + +lemma partial_pullback_memE: + assumes "as \ partial_pullback m f l S" + shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "partial_image m f as \ S" + using assms apply (metis evimage_eq partial_pullback_def) + using assms unfolding partial_pullback_def + by blast + +lemma partial_pullback_closed: +"partial_pullback m f l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + using partial_pullback_memE(1) by blast + +lemma partial_pullback_memI: + assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" + assumes "(f (take m as))#(drop m as) \ S" + shows "as \ partial_pullback m f k S" + using assms unfolding partial_pullback_def partial_image_def evimage_def + by blast + +lemma partial_image_eq: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + assumes "x = as @ bs" + shows "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 partial_image_def + by blast +qed + +lemma partial_pullback_memE': + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + assumes "x = as @ bs" + assumes "x \ partial_pullback n f k S" + shows "(f as)#bs \ S" + using partial_pullback_memE[of x n f k S] partial_image_def[of n f x] + by (metis assms(1) assms(2) assms(3) assms(4) partial_image_eq) + +text\Partial pullbacks have the same algebraic properties as pullbacks\ + +lemma partial_pullback_intersect: +"partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" + unfolding partial_pullback_def + by simp + +lemma partial_pullback_union: +"partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" + unfolding partial_pullback_def + by simp + +lemma cartesian_power_drop: + assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + shows "drop n x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms cartesian_power_car_memE + apply (metis add_diff_cancel_left' length_drop) + using assms cartesian_power_car_memE'' + by (metis order.trans set_drop_subset) + +lemma partial_pullback_complement: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (partial_pullback m f l S) " + apply(rule equalityI) + using partial_pullback_def[of m f l "(carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S)"] + partial_pullback_def[of m f l S] + apply (smt Diff_iff evimage_Diff partial_pullback_memE(1) subsetI) +proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - partial_pullback m f l S" + show " x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) " + apply(rule partial_pullback_memI) + using A + apply blast + proof + have 00: "Suc l = l + 1" + by auto + have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + by (meson A DiffD1 cartesian_power_drop) + have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A by (meson DiffD1 le_add1 take_closed) + have "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>l+1\<^esup>) " + using assms 0 1 00 cartesian_power_cons[of "drop m x" Q\<^sub>p l "f (take m x)"] + by blast + thus "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) " + using 00 by metis + show "f (take m x) # drop m x \ S" + using A unfolding partial_pullback_def partial_image_def + by blast + qed +qed + +lemma partial_pullback_carrier: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + apply(rule equalityI) + using partial_pullback_memE(1) apply blast +proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + show "x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>))" + apply(rule partial_pullback_memI) + using A cartesian_power_drop[of x m l] assms + apply blast +proof- + have "f (take m x) \ carrier Q\<^sub>p" + using A assms take_closed[of m "m+l" x Q\<^sub>p] + by (meson Pi_mem le_add1) + then show "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)" + using cartesian_power_drop[of x m l] + by (metis A add.commute cartesian_power_cons plus_1_eq_Suc) +qed +qed + +text\Definition 1.4 from Denef\ + +definition is_semialg_function where +"is_semialg_function m f = ((f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p) \ + (\l \ 0. \S \ semialg_sets (1 + l). is_semialgebraic (m + l) (partial_pullback m f l S)))" + +lemma is_semialg_function_closed: + assumes "is_semialg_function m f" + shows "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using is_semialg_function_def assms by blast + +lemma is_semialg_functionE: + assumes "is_semialg_function m f" + assumes "is_semialgebraic (1 + k) S" + shows " is_semialgebraic (m + k) (partial_pullback m f k S)" + using is_semialg_function_def assms + by (meson is_semialgebraicE le0) + +lemma is_semialg_functionI: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" + shows "is_semialg_function m f" + using assms unfolding is_semialg_function_def + by blast + +text\Semialgebraicity for functions can be verified on basic semialgebraic sets \ + +lemma is_semialg_functionI': + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" + shows "is_semialg_function m f" + apply(rule is_semialg_functionI) + using assms(1) apply blast +proof- + show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" + proof- fix k S assume A: "S \ semialg_sets (1 + k)" + show "is_semialgebraic (m + k) (partial_pullback m f k S)" + apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" "basic_semialgs (1 + k)"]) + using A unfolding semialg_sets_def + apply blast + using partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger + apply (metis assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg partial_pullback_carrier partial_pullback_intersect plus_1_eq_Suc) + using partial_pullback_union union_is_semialgebraic apply presburger + using assms(1) complement_is_semialg partial_pullback_complement plus_1_eq_Suc by presburger + qed +qed + +text\Graphs of semialgebraic functions are semialgebraic\ +abbreviation graph where +"graph \ fun_graph Q\<^sub>p" + +lemma graph_memE: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "x \ graph m f" + shows "f (take m x) = x!m" + "x = (take m x)@[f (take m x)]" + "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" +proof- + obtain a where a_def: "a\carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ x = a @ [f a]" + using assms + unfolding fun_graph_def + by blast + then have 0: "a = take m x" + by (metis append_eq_conv_conj cartesian_power_car_memE) + then show "f (take m x) = x!m" + by (metis a_def cartesian_power_car_memE nth_append_length) + show "x = (take m x)@[f (take m x)]" + using "0" a_def + by blast + show "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using "0" a_def by blast +qed + +lemma graph_memI: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "f (take m x) = x!m" + assumes "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" + shows "x \ graph m f" +proof- + have 0: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + apply(rule take_closed[of _ "m + 1"]) + apply simp + using assms(3) by blast + have "x = (take m x)@[x!m]" + by (metis \take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)\ add.commute + assms(3) cartesian_power_car_memE length_append_singleton lessI + nth_equalityI nth_take plus_1_eq_Suc take_Suc_conv_app_nth) + then have "x = (take m x)@[f (take m x)]" + using assms(2) + by presburger + then show ?thesis + using assms 0 + unfolding fun_graph_def + by blast +qed + +lemma graph_mem_closed: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "x \ graph m f" + shows "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" +proof(rule cartesian_power_car_memI') + show "length x = m + 1" + using assms graph_memE[of f m x] + by (smt Groups.add_ac(2) cartesian_power_car_memE fun_graph_def length_append_singleton mem_Collect_eq plus_1_eq_Suc) + show "\i. i < m + 1 \ x ! i \ carrier Q\<^sub>p" + proof- fix i assume A: "i < m + 1" + then show "x ! i \ carrier Q\<^sub>p" + proof(cases "i = m") + case True + then show ?thesis using graph_memE[of f m x] + by (metis PiE assms(1) assms(2)) + next + case False + then show ?thesis using graph_memE[of f m x] + by (metis \i < m + 1\ add.commute assms(1) assms(2) cartesian_power_car_memE' less_SucE nth_take plus_1_eq_Suc) + qed + qed +qed + +lemma graph_closed: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + shows "graph m f \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" + using assms graph_mem_closed + by blast + +text\The \m\-dimensional diagonal set is semialgebraic\ + +notation diagonal ("\ ") + +lemma diag_is_algebraic: + shows "is_algebraic Q\<^sub>p (n + n) (\ n)" + using Qp.cring_axioms diagonal_is_algebraic + by blast + +lemma diag_is_semialgebraic: + shows "is_semialgebraic (n + n) (\ n)" + using diag_is_algebraic is_algebraic_imp_is_semialg + by blast + +text\Transposition permutations\ + +definition transpose where +"transpose i j = (Fun.swap i j id)" + +lemma transpose_permutes: + assumes "i< n" + assumes "j < n" + shows "transpose i j permutes {..x. x \ {.. Fun.swap i j id x = x" + using assms by (auto simp: Transposition.transpose_def) + show "\y. \!x. Fun.swap i j id x = y" + proof fix y show "\!x. Fun.swap i j id x = y" + using swap_id_eq[of i j y] + by (metis eq_id_iff swap_apply(1) swap_apply(2) swap_id_eq swap_self) +qed +qed + +lemma transpose_alt_def: +"transpose a b x = (if x = a then b else if x = b then a else x)" + using swap_id_eq + by (simp add: transpose_def) + +definition last_to_first where +"last_to_first n = (\i. if i = (n-1) then 0 else if i < n-1 then i + 1 else i)" + +definition first_to_last where +"first_to_last n = fun_inv (last_to_first n)" + +lemma last_to_first_permutes: + assumes "(n::nat) > 0" + shows "last_to_first n permutes {..x. x \ {.. last_to_first n x = x" + proof fix x show " x \ {.. last_to_first n x = x" + proof assume A: "x \ {.. x < n" + by blast then have "x \ n" by linarith + then show "last_to_first n x = x" + unfolding last_to_first_def using assms + by auto + qed + qed + show "\y. \!x. last_to_first n x = y" + proof fix y + show "\!x. last_to_first n x = y" + proof(cases "y = 0") + case True + then have 0: "last_to_first n (n-1) = y" + using last_to_first_def + by (simp add: last_to_first_def) + have 1: "\x. last_to_first n x = y \ x = n-1" + unfolding last_to_first_def using True + by (metis add_gr_0 less_numeral_extra(1) not_gr_zero) + show ?thesis + using 0 1 + by blast + next + case False + then show ?thesis + proof(cases "y < n") + case True + then have 0: "last_to_first n (y-1) = y" + using False True + unfolding last_to_first_def + using add.commute by auto + have 1: "\x. last_to_first n x = y \ x =(y-1)" + unfolding last_to_first_def + using True False + by auto + show ?thesis using 0 1 by blast + next + case F: False + then have 0: "y \ n" + using not_less by blast + then have 1: "last_to_first n y = y" + by (simp add: \\x. x \ {.. last_to_first n x = x\) + have 2: "\x. last_to_first n x = y \ x =y" + using 0 unfolding last_to_first_def + using False by presburger + then show ?thesis using 1 2 by blast + qed + qed + qed +qed + +definition graph_swap where +"graph_swap n f = permute_list ((first_to_last (n+1))) ` (graph n f)" + +lemma last_to_first_eq: + assumes "length as = n" + shows "permute_list (last_to_first (n+1)) (a#as) = (as@[a])" +proof- + have 0: "\i. i < (n+1) \ permute_list (last_to_first (n + 1)) (a # as) ! i = (as@[a]) ! i" + proof- + fix i assume A: "i < n+1" + show "permute_list (last_to_first (n + 1)) (a # as) ! i = (as @ [a]) ! i" + proof(cases "i = n") + case True + have 0: "(as @ [a]) ! i = a" + by (metis True assms nth_append_length) + have 1: "length (a#as) = n + 1" + by (simp add: assms) + have 2: "i < length (a # as)" + using "1" A by linarith + have 3: "last_to_first (n + 1) permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "a \ carrier Q\<^sub>p" + shows "permute_list (first_to_last (n+1)) (as@[a]) = (a#as)" +proof- + have "length as = n" + using assms(1) cartesian_power_car_memE by blast + then show ?thesis + using last_to_first_eq last_to_first_permutes[of n] + permute_list_compose_inv(2)[of "(last_to_first (n + 1))" n "a # as"] + unfolding first_to_last_def + by (metis add_gr_0 assms(1) assms(2) cartesian_power_append last_to_first_permutes + less_one permute_list_closed' permute_list_compose_inv(2)) +qed + +lemma graph_swapI: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "(f as)#as \ graph_swap n f" +proof- + have 0: "as@[f as] \ graph n f" + using assms using graph_memI[of f n] fun_graph_def + by blast + have 1: "f as \ carrier Q\<^sub>p" + using assms + by blast + then show ?thesis + using assms 0 first_to_last_eq[of as "n" "f as"] + unfolding graph_swap_def + by (metis image_eqI) +qed + +lemma graph_swapE: + assumes "x \ graph_swap n f" + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "hd x = f (tl x)" +proof- + obtain y where y_def: "y \ graph n f \ x = permute_list (first_to_last (n+1)) y" + using assms graph_swap_def + by (smt image_def mem_Collect_eq) + then have "take n y \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms(2) graph_memE(3) + by blast + then show "hd x = f (tl x)" + by (metis (no_types, lifting) add.commute assms(2) cartesian_power_car_memE' + first_to_last_eq graph_memE(1) graph_memE(2) graph_mem_closed lessI list.sel(1) + list.sel(3) plus_1_eq_Suc y_def) +qed + +text\Semialgebraic functions have semialgebraic graphs\ + +lemma graph_as_partial_pullback: + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "partial_pullback n f 1 (\ 1) = graph n f" +proof + show "partial_pullback n f 1 (\ 1) \ graph n f" + proof fix x assume A: "x \ partial_pullback n f 1 (\ 1)" + then have 0: "f (take n x) # drop n x \ \ 1" + by (metis local.partial_image_def partial_pullback_memE(2)) + then have 1: "length (f (take n x) # drop n x) = 2" + using diagonal_def + by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq one_add_one) + then obtain b where b_def: "[b] = drop n x" + by (metis list.inject pair_id) + then have "[f (take n x), b] \ \ 1" + using "0" + by presburger + then have "b = f (take n x)" + using 0 + by (smt One_nat_def Qp.cring_axioms diagonal_def drop0 drop_Suc_Cons list.inject mem_Collect_eq take_Suc_Cons) + then have "x = (take n x)@[f (take n x)]" + by (metis append_take_drop_id b_def) + then show "x \ graph n f" using graph_memI[of f n x] + by (metis (no_types, lifting) A \b = f (take n x)\ + assms b_def nth_via_drop partial_pullback_memE(1)) + qed + show "graph n f \ partial_pullback n f 1 (\ 1)" + proof fix x + assume A: "x \ graph n f " + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" + using assms graph_mem_closed by blast + have "x = (take n x) @ [f (take n x)]" + using A graph_memE(2)[of f n x] assms + by blast + then have "partial_image n f x = [f (take n x), f (take n x)]" + by (metis append_take_drop_id local.partial_image_def same_append_eq) + then have "partial_image n f x \ \ 1" + using assms 0 diagonal_def[of 1] Qp.cring_axioms diagonalI[of "partial_image n f x"] + by (metis (no_types, lifting) A append_Cons append_eq_conv_conj + cartesian_power_car_memE cartesian_power_car_memE' graph_memE(1) + less_add_one self_append_conv2 Qp.to_R1_closed) + then show "x \ partial_pullback n f 1 (\ 1)" + unfolding partial_pullback_def using 0 + by blast + qed +qed + +lemma semialg_graph: + assumes "is_semialg_function n f" + shows "is_semialgebraic (n + 1) (graph n f)" + using assms graph_as_partial_pullback[of f n] unfolding is_semialg_function_def + by (metis diag_is_semialgebraic is_semialgebraicE less_imp_le_nat less_numeral_extra(1)) + +text\Functions induced by polynomials are semialgebraic\ + +definition var_list_segment where +"var_list_segment i j = map (\i. pvar Q\<^sub>p i) [i..< j]" + +lemma var_list_segment_length: + assumes "i \ j" + shows "length (var_list_segment i j) = j - i" + using assms var_list_segment_def + by fastforce + +lemma var_list_segment_entry: + assumes "k < j - i" + assumes "i \ j" + shows "var_list_segment i j ! k = pvar Q\<^sub>p (i + k)" + using assms var_list_segment_length + unfolding var_list_segment_def + using nth_map_upt by blast + +lemma var_list_segment_is_poly_tuple: + assumes "i \j" + assumes "j \ n" + shows "is_poly_tuple n (var_list_segment i j)" + apply(rule Qp_is_poly_tupleI) + using assms var_list_segment_entry var_list_segment_length Qp.cring_axioms pvar_closed[of _ n] + by (metis (no_types, opaque_lifting) add.commute add_lessD1 diff_add_inverse le_Suc_ex + less_diff_conv) + +lemma map_by_var_list_segment: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "j \ n" + assumes "i \ j" + shows "poly_map n (var_list_segment i j) as = list_segment i j as" + apply(rule nth_equalityI ) + unfolding poly_map_def var_list_segment_def list_segment_def restrict_def poly_tuple_eval_def + apply (metis (full_types) assms(1) length_map) + using assms eval_pvar[of _ n as] Qp.cring_axioms length_map add.commute + length_upt less_diff_conv less_imp_add_positive nth_map nth_upt + trans_less_add2 + by (smt le_add_diff_inverse2) + +lemma map_by_var_list_segment_to_length: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "i \ n" + shows "poly_map n (var_list_segment i n) as = drop i as" + apply(rule nth_equalityI ) + apply (metis Qp_poly_mapE' assms(1) assms(2) cartesian_power_car_memE length_drop var_list_segment_length) + using assms map_by_var_list_segment[of as n n i] list_segment_drop[of i as] cartesian_power_car_memE[of as Q\<^sub>p n] + map_nth[of ] nth_drop nth_map[of _ "[i..p)" ] nth_map[of _ "map (pvar Q\<^sub>p) [i..p as"] + unfolding poly_map_def poly_tuple_eval_def var_list_segment_def restrict_def list_segment_def + by (smt add.commute add_eq_self_zero drop_map drop_upt le_Suc_ex le_refl) + +lemma map_tail_by_var_list_segment: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "a \ carrier Q\<^sub>p" + assumes "i < n" + shows "poly_map (n+1) (var_list_segment 1 (n+1)) (a#as) = as" +proof- + have 0: "(a#as) \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" + using assms + by (meson cartesian_power_cons) + have 1: "length as = n" + using assms cartesian_power_car_memE + by blast + have 2: "drop 1 (a # as) = as" + using 0 1 using list_segment_drop[of 1 "a#as"] + by (metis One_nat_def drop0 drop_Suc_Cons ) + have "1 \n + 1" by auto + then show ?thesis + using 0 2 map_by_var_list_segment_to_length[of "a#as" "n+1" 1] + by presburger +qed + +lemma Qp_poly_tuple_Cons: + assumes "is_poly_tuple n fs" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" + assumes "k \n" + shows "is_poly_tuple n (f#fs)" + using is_poly_tuple_Cons[of n fs f] poly_ring_car_mono[of k n] assms + by blast + +lemma poly_map_Cons: + assumes "is_poly_tuple n fs" + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "poly_map n (f#fs) a = (Qp_ev f a)#poly_map n fs a" + using assms poly_map_cons by blast + +lemma poly_map_append': + assumes "is_poly_tuple n fs" + assumes "is_poly_tuple n gs" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "poly_map n (fs@gs) a = poly_map n fs a @ poly_map n gs a" + using assms(3) poly_map_append by blast + +lemma partial_pullback_by_poly: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + shows "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n+k) S (f# (var_list_segment n (n+k)))" +proof + show "partial_pullback n (Qp_ev f) k S \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" + proof fix x assume A: " x \ partial_pullback n (Qp_ev f) k S" + then obtain as bs where as_bs_def: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ x = as @ bs" + using partial_pullback_memE(1)[of x n "(Qp_ev f)" k S] cartesian_power_decomp + by metis + then have 0: "(Qp_ev f as#bs) \ S" + using A partial_pullback_memE' + by blast + have 1: "Qp_ev f as = Qp_ev f (as@bs)" + using assms as_bs_def poly_eval_cartesian_prod[of as n bs k f] + Qp.cring_axioms [of ] + by metis + then have 2: "((Qp_ev f x) #bs) \ S" + using "0" as_bs_def + by presburger + have 3: "bs = list_segment n (n+k) x" + using as_bs_def list_segment_drop[of n x] + by (metis (no_types, lifting) add_cancel_right_right add_diff_cancel_left' + append_eq_append_conv append_take_drop_id cartesian_power_car_memE + length_0_conv length_append length_map length_upt linorder_neqE_nat + list_segment_def not_add_less1) + have 4: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" + using Qp_poly_tuple_Cons + var_list_segment_is_poly_tuple + by (metis add.commute assms(1) dual_order.refl le_add2) + have 5: "f \ carrier (Q\<^sub>p [\\<^bsub>n + k\<^esub>])" + using poly_ring_car_mono[of n "n + k"] assms le_add1 by blast + have 6: "is_poly_tuple (n + k) (var_list_segment n (n + k))" + by (simp add: var_list_segment_is_poly_tuple) + have 7: "x \ carrier (Q\<^sub>p\<^bsup>n + k\<^esup>)" + using as_bs_def cartesian_power_concat(1) by blast + hence 8: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#poly_map (n+k) (var_list_segment n (n + k)) x" + using 5 6 7 A poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 4 + unfolding partial_pullback_def evimage_def + by blast + hence 6: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#bs" + using 3 "7" le_add1 le_refl map_by_var_list_segment by presburger + show " x \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" + unfolding poly_tuple_pullback_def using 6 + by (metis "2" "7" IntI poly_map_apply vimage_eq) + qed + show "poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k)) \ partial_pullback n (Qp_ev f) k S" + proof fix x + assume A: "x \ poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" + have 0: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" + using Qp_poly_tuple_Cons assms(1) le_add1 var_list_segment_is_poly_tuple + by blast + have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + using A unfolding poly_tuple_pullback_def + by blast + have 2: "poly_map (n+k) (f # var_list_segment n (n + k)) x \ S" + using 1 assms A unfolding poly_map_def poly_tuple_pullback_def restrict_def + by (metis (no_types, opaque_lifting) Int_commute add.commute evimage_def evimage_eq) + have 3: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#(drop n x)" + using poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 1 assms(1) map_by_var_list_segment_to_length + le_add1 poly_map_cons by presburger + have 4: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f (take n x))#(drop n x)" + using assms 1 3 eval_at_points_higher_pow[of f n "n + k" "x"] le_add1 + by (metis nat_le_iff_add) + show "x \ partial_pullback n (Qp_ev f) k S" + apply(rule partial_pullback_memI) + using 1 apply blast + using 2 3 4 by metis + qed +qed + +lemma poly_is_semialg: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "is_semialg_function n (Qp_ev f)" +proof(rule is_semialg_functionI) + show "Qp_ev f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + using assms + by (meson Pi_I eval_at_point_closed) + show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" + proof- fix k::nat fix S + assume A: "S \ semialg_sets (1 + k)" + have 0: "is_poly_tuple (n + k) (f # var_list_segment n (n + k))" + by (metis add.commute assms le_add2 order_refl Qp_poly_tuple_Cons + var_list_segment_is_poly_tuple) + have 1: "length (f # var_list_segment n (n + k)) = k + 1" + by (metis add.commute add_diff_cancel_left' le_add1 length_Cons + plus_1_eq_Suc var_list_segment_length) + have 2: "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" + using A assms partial_pullback_by_poly[of f n S k] + unfolding semialg_sets_def + using gen_boolean_algebra_subset + by blast + then show "is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" + using add.commute[of 1 k] 0 1 assms(1) + pullback_is_semialg[of "n+k" "(f # var_list_segment n (n + k))" "k+1" S] + by (metis A is_semialgebraicI is_semialgebraic_closed poly_tuple_pullback_eq_poly_map_vimage) + qed +qed + +text\Families of polynomials defined by semialgebraic coefficient functions\ + +lemma semialg_function_on_carrier: + assumes "is_semialg_function n f" + assumes "restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + shows "is_semialg_function n g" +proof(rule is_semialg_functionI) + have 0: "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + using assms(1) is_semialg_function_closed + by blast + show "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" then show " g x \ carrier Q\<^sub>p" + using assms(2) 0 + by (metis (no_types, lifting) PiE restrict_Pi_cancel) + qed + show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n g k S)" + proof- fix k S + assume A: "S \ semialg_sets (1 + k)" + have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" + using A assms(1) is_semialg_functionE is_semialgebraicI + by blast + have 2: "(partial_pullback n f k S) = (partial_pullback n g k S)" + unfolding partial_pullback_def partial_image_def evimage_def + proof + show "(\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + proof fix x assume "x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) " + have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms + by (meson \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ + 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. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + using assms \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ + by blast + qed + show "(\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + proof fix x assume A: "x \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + using A by blast + qed + qed + then show "is_semialgebraic (n + k) (partial_pullback n g k S)" + using 1 by auto + qed +qed + +lemma semialg_function_on_carrier': + assumes "is_semialg_function n f" + assumes "\a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = g a" + shows "is_semialg_function n g" + using assms semialg_function_on_carrier unfolding restrict_def + by (meson restrict_ext semialg_function_on_carrier) + +lemma constant_function_is_semialg: + assumes "n > 0" + assumes "x \ carrier Q\<^sub>p" + assumes "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = x" + shows "is_semialg_function n f" +proof(rule semialg_function_on_carrier[of _ "Qp_ev (Qp_to_IP x)"]) + show "is_semialg_function n (Qp_ev (Qp_to_IP x))" + using assms poly_is_semialg[of "(Qp_to_IP x)"] Qp_to_IP_car + by blast + have 0: "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = Qp_ev (Qp_to_IP x) a" + using eval_at_point_const assms + by blast + then show "restrict (Qp_ev (Qp_to_IP x)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + by (metis (no_types, lifting) restrict_ext) +qed + +lemma cartesian_product_singleton_factor_projection_is_semialg: + assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "is_semialgebraic (m+n) (cartesian_product A {b})" + shows "is_semialgebraic m A" +proof- + obtain f where f_def: "f = map (pvar Q\<^sub>p) [0..p" "[0.. ((nat \ int) \ (nat \ int)) set) list) = map (\i::nat. Qp.indexed_const (b ! i)) [(0::nat)..i. i \ set [0::nat..< n] \ b!i \ carrier Q\<^sub>p" + using assms(2) cartesian_power_car_memE'[of b Q\<^sub>p n] by blast + hence 1: "\i. i \ set [0::nat..< n] \ Qp.indexed_const (b ! i) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" + using assms Qp_to_IP_car by blast + show ?thesis + unfolding is_poly_tuple_def g_def + apply(rule subsetI) + using set_map[of "\i. Qp.indexed_const (b ! i)" "[0..x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ poly_tuple_eval (f@g) x = x@b" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + have 30: "poly_tuple_eval f x = x" + proof- + have 300: "length (poly_tuple_eval f x) = length x" + unfolding poly_tuple_eval_def using cartesian_power_car_memE + by (metis "4" A length_map) + have "\i. i < length x \ poly_tuple_eval f x ! i = x ! i" + unfolding f_def poly_tuple_eval_def using nth_map + by (metis "4" A add_cancel_right_left cartesian_power_car_memE eval_pvar f_def length_map nth_upt) + thus ?thesis using 300 + by (metis nth_equalityI) + qed + have 31: "poly_tuple_eval g x = b" + proof- + have 310: "length (poly_tuple_eval g x) = length b" + unfolding poly_tuple_eval_def g_def using cartesian_power_car_memE + by (metis assms(2) length_map map_nth) + have 311: "length b = n" using assms cartesian_power_car_memE by blast + hence "\i. i < n \ poly_tuple_eval g x ! i = b ! i" proof- fix i assume "i < n" + thus "poly_tuple_eval g x ! i = b ! i" + unfolding g_def poly_tuple_eval_def using eval_at_point_const[of "b!i" x m] 310 nth_map + by (metis "311" A assms(2) cartesian_power_car_memE' length_map map_nth) + qed + thus ?thesis using 311 310 nth_equalityI + by (metis list_eq_iff_nth_eq) + qed + have 32: "poly_tuple_eval (f @ g) x = poly_map m (f@g) x" + unfolding poly_map_def restrict_def using A + by (simp add: A) + have 33: "poly_tuple_eval f x = poly_map m f x" + unfolding poly_map_def restrict_def using A + by (simp add: A) + have 34: "poly_tuple_eval g x = poly_map m g x" + unfolding poly_map_def restrict_def using A + by (simp add: A) + show "poly_tuple_eval (f @ g) x = x @ b" + using assms 1 2 30 31 poly_map_append[of x m f g] A unfolding 32 33 34 + by (simp add: A \b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\) + qed + have 4: "A = (poly_tuple_eval (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b}))" + proof + show "A \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" + proof(rule subsetI) fix x assume A: "x \ A" + then have 0: "poly_tuple_eval (f@g) x = x@b" + using 3 assms by blast + then show " x \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" + using A cartesian_product_memE + by (smt Un_upper1 assms(1) assms(2) cartesian_product_memI' evimageI2 in_mono insert_is_Un mk_disjoint_insert singletonI) + qed + show "poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" + proof(rule subsetI) fix x assume A: "x \ (poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b})" + then have "poly_tuple_eval (f @ g) x \ cartesian_product A {b}" + by blast + then have "x@b \ cartesian_product A {b}" + using A 3 by (metis evimage_eq) + then show "x \ A" + using A + by (metis append_same_eq cartesian_product_memE' singletonD) + qed + qed + have 5: "A = poly_map m (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b})" + proof + show "A \ poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" + unfolding poly_map_def evimage_def restrict_def using 4 + by (smt IntI assms(1) evimageD in_mono subsetI vimageI) + show "poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" + unfolding poly_map_def evimage_def restrict_def using 4 + by (smt Int_iff evimageI2 subsetI vimage_eq) + qed + have 6: "length (f @ g) = m + n" + unfolding f_def g_def by (metis index_list_length length_append length_map map_nth) + show ?thesis using 2 5 6 assms pullback_is_semialg[of m "f@g" "m+n" "cartesian_product A {b}"] + by (metis is_semialgebraicE zero_eq_add_iff_both_eq_0) +qed + +lemma cartesian_product_factor_projection_is_semialg: + assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "B \ {}" + assumes "is_semialgebraic (m+n) (cartesian_product A B)" + shows "is_semialgebraic m A" +proof- + obtain b where b_def: "b \ B" + using assms by blast + have "is_semialgebraic n {b}" + using assms b_def is_algebraic_imp_is_semialg singleton_is_algebraic by blast + hence 0: "is_semialgebraic (m+n) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b})" + using car_times_semialg_is_semialg assms(4) by blast + have "(cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b}) \ (cartesian_product A B) + = (cartesian_product A {b})" + using assms b_def cartesian_product_intersection[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m "{b}" n A B] + by (metis (no_types, lifting) Int_absorb1 Int_empty_left Int_insert_left_if1 \is_semialgebraic n {b}\ is_semialgebraic_closed set_eq_subset) + hence "is_semialgebraic (m+n) (cartesian_product A {b})" + using assms 0 intersection_is_semialg by metis + thus ?thesis using assms cartesian_product_singleton_factor_projection_is_semialg + by (meson \is_semialgebraic n {b}\ insert_subset is_semialgebraic_closed) +qed + +lemma partial_pullback_cartesian_product: + assumes "\ \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + shows "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) = partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) " +proof + show "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" + proof fix x assume A: "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + then obtain y t where yt_def: "x = y@[t] \ y \ partial_pullback m \ 0 S \ t \ carrier Q\<^sub>p" + by (metis cartesian_product_memE' Qp.to_R1_to_R Qp.to_R_pow_closed) + then have "[\ y] \ S" + using partial_pullback_memE unfolding partial_image_def + by (metis (no_types, lifting) add.right_neutral append.right_neutral cartesian_power_drop le_zero_eq take_closed partial_pullback_memE' take_eq_Nil) + then have 0: "[\ y]@[t] \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + using cartesian_product_memI' yt_def + by (metis assms(2) carrier_is_semialgebraic is_semialgebraic_closed Qp.to_R1_closed) + have 1: " x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" + using A yt_def + by (metis add.right_neutral cartesian_power_append partial_pullback_memE(1)) + show "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" + apply(rule partial_pullback_memI) + using "1" apply blast + using yt_def 0 + by (smt Cons_eq_appendI add.right_neutral local.partial_image_def partial_image_eq partial_pullback_memE(1) self_append_conv2 Qp.to_R1_closed) + qed + show "partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + proof(rule subsetI) fix x assume A: "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" + using assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] + by blast + have 1: "\ (take m x) # drop m x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + using A assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] + unfolding partial_image_def + by blast + have 2: "\ (take m (take m x)) # drop m (take m x) = [\ (take m x)]" + using 0 1 + by (metis add.commute add.right_neutral append.right_neutral append_take_drop_id take0 take_drop) + show "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + apply(rule cartesian_product_memI[of _ Q\<^sub>p m _ 1]) + apply (metis add_cancel_right_right partial_pullback_closed) + apply blast + apply(rule partial_pullback_memI[of _ m 0 \ S]) using 0 + apply (metis Nat.add_0_right le_iff_add take_closed) + using 2 apply (metis (no_types, lifting) "1" add.commute add.right_neutral assms(2) cartesian_product_memE(1) list.inject plus_1_eq_Suc take_Suc_Cons take_drop) + using 0 cartesian_power_drop by blast + qed +qed + +lemma cartesian_product_swap: + assumes "A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "is_semialgebraic (m+n) (cartesian_product A B)" + shows "is_semialgebraic (m+n) (cartesian_product B A)" +proof- + obtain f where f_def: "f = (\i. (if i < m then n + i else (if i < m+n then i - m else i)))" + by blast + have 0: "\i. i \ {.. f i \ {n..i. i \ {m.. f i \ {..i. i \ {.. f i \ {..x. x \ {.. f x = x" + unfolding f_def by simp + show "\y. \!x. f x = y" + proof fix y + show "\!x. f x = y" + proof(cases "y < n") + case True + have T0: "f (y+m) = y" + unfolding f_def using True + by simp + have "\i. f i = y \ i \ {m..i. f i = y \ i = y+m" + using T0 unfolding f_def by auto + thus ?thesis using T0 by blast + next + case False + show ?thesis + proof(cases "y \ {n..i. f i = y \ i \ {..i. f i = y \ i = y- n" + using f_def by force + then show ?thesis using T0 by blast + next + case F: False + then show ?thesis using 0 1 2 unfolding f_def + using False add_diff_inverse_nat lessThan_iff by auto + qed + qed + qed + qed + have "permute_list f ` (cartesian_product A B) = (cartesian_product B A)" + proof + show "permute_list f ` cartesian_product A B \ cartesian_product B A" + proof fix x assume A: " x \ permute_list f ` cartesian_product A B" + then obtain a b where ab_def: "a \ A \b \ B \ x = permute_list f (a@b)" + by (metis (mono_tags, lifting) cartesian_product_memE' image_iff) + have 0: "x = permute_list f (a@b)" + using ab_def by blast + have 1: "length a = n" + using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast + have 2: "length b = m" + using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast + have 3: "length x = m + n" + using 1 2 0 f_permutes by simp + have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" + unfolding 0 using permute_list_nth + by (metis "0" "3" f_permutes length_permute_list trans_less_add1) + hence 5: "\i. i < m \ x ! i = b!i" + unfolding f_def using 1 2 + by (metis "4" f_def nth_append_length_plus) + have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" + unfolding 0 using f_def permute_list_nth f_permutes + by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 + ordered_cancel_comm_monoid_diff_class.diff_add) + have 7: "x = b@a" + proof(rule nth_equalityI) + show "length x = length (b @ a)" + using 1 2 3 by simp + show "\i. i < length x \ x ! i = (b @ a) ! i" + unfolding 3 using 1 2 4 5 + by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) + qed + show "x \ cartesian_product B A" unfolding 7 using ab_def unfolding cartesian_product_def by blast + qed + show "cartesian_product B A \ permute_list f ` cartesian_product A B" + proof fix y assume A: "y \ cartesian_product B A" + then obtain b a where ab_def: "b \ B \ a \ A \ y = b@a" + using cartesian_product_memE' by blast + obtain x where 0: "x = permute_list f (a@b)" + by blast + have 1: "length a = n" + using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast + have 2: "length b = m" + using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast + have 3: "length x = m + n" + using 1 2 0 f_permutes by simp + have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" + unfolding 0 using permute_list_nth + by (metis "0" "3" f_permutes length_permute_list trans_less_add1) + hence 5: "\i. i < m \ x ! i = b!i" + unfolding f_def using 1 2 + by (metis "4" f_def nth_append_length_plus) + have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" + unfolding 0 using f_def permute_list_nth f_permutes + by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 + ordered_cancel_comm_monoid_diff_class.diff_add) + have 7: "x = b@a" + proof(rule nth_equalityI) + show "length x = length (b @ a)" + using 1 2 3 by simp + show "\i. i < length x \ x ! i = (b @ a) ! i" + unfolding 3 using 1 2 4 5 + by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) + qed + show "y \ permute_list f ` cartesian_product A B" + using ab_def 7 cartesian_product_memI'[of _ Q\<^sub>p] unfolding 0 + by (metis assms(1) assms(2) image_eqI) + qed + qed + thus ?thesis using assms f_permutes permutation_is_semialgebraic + by metis +qed + +lemma Qp_zero_subset_is_semialg: + assumes "S \ carrier (Q\<^sub>p\<^bsup>0\<^esup>)" + shows "is_semialgebraic 0 S" +proof(cases "S = {}") + case True + then show ?thesis + by (simp add: empty_is_semialgebraic) +next + case False + then have "S = carrier (Q\<^sub>p\<^bsup>0\<^esup>)" + using assms unfolding Qp_zero_carrier by blast + then show ?thesis + by (simp add: carrier_is_semialgebraic) +qed + +lemma cartesian_product_empty_list: +"cartesian_product A {[]} = A" +"cartesian_product {[]} A = A" +proof + show "cartesian_product A {[]} \ A" + apply(rule subsetI) + unfolding cartesian_product_def + by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) + show "A \ cartesian_product A {[]}" + apply(rule subsetI) + unfolding cartesian_product_def + by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) + show "cartesian_product {[]} A = A" + proof + show "cartesian_product {[]} A \ A" + apply(rule subsetI) + unfolding cartesian_product_def + by (smt append_self_conv2 bex_empty insert_compr mem_Collect_eq) + show "A \ cartesian_product {[]} A" + apply(rule subsetI) + unfolding cartesian_product_def + by blast + qed +qed + +lemma cartesian_product_singleton_factor_projection_is_semialg': + assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "is_semialgebraic (m+n) (cartesian_product A {b})" + shows "is_semialgebraic m A" +proof(cases "n > 0") + case True + show ?thesis + proof(cases "m > 0") + case T: True + then show ?thesis + using assms True cartesian_product_singleton_factor_projection_is_semialg by blast + next + case False + then show ?thesis using Qp_zero_subset_is_semialg assms by blast + qed +next + case False + then have F0: "b = []" + using assms Qp_zero_carrier by blast + have "cartesian_product A {b} = A" + unfolding F0 + by (simp add: cartesian_product_empty_list(1)) + then show ?thesis using assms False + by (metis add.right_neutral gr0I) +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection \More on graphs of functions\ +(**************************************************************************************************) +(**************************************************************************************************) + + +text\This section lays the groundwork for showing that semialgebraic functions are closed under + various algebraic operations\ + +text\The take and drop functions on lists are polynomial maps\ + +lemma function_restriction: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ S" + assumes "n \ k" + shows "(g \ (take n)) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ S" +proof fix x + assume "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + then have "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms(2) take_closed + by blast + then show "(g \ take n) x \ S" + using assms comp_apply + by (metis Pi_iff comp_def) +qed + +lemma partial_pullback_restriction: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "n < k" + shows "partial_pullback k (g \ take n) m S = + split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" +proof(rule equalityI) + show "partial_pullback k (g \ take n) m S \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" + proof fix x assume A: "x \ partial_pullback k (g \ take n) m S" + obtain as bs where asbs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using partial_pullback_memE[of x k "g \ take n" m S] A cartesian_power_decomp[of x Q\<^sub>p k m] + by metis + have 0: "((g \ (take n)) as)#bs \ S" + using asbs_def partial_pullback_memE'[of as k bs m x] A + by blast + have 1: "(g (take n as))#bs \ S" + using 0 + by (metis comp_apply) + have 2: "take n as @ bs \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" + by (meson asbs_def assms(2) cartesian_power_concat(1) less_imp_le_nat take_closed) + have 3: "(take n as)@bs \ (partial_pullback n g m S)" + using 1 2 partial_pullback_memI[of "(take n as)@bs" n m g S] + by (metis (mono_tags, opaque_lifting) asbs_def assms(2) local.partial_image_def nat_less_le + partial_image_eq subsetD subset_refl take_closed) + have 4: "drop n as \ (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" + using asbs_def assms(2) drop_closed + by blast + show " x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" + using split_cartesian_product_memI[of "take n as" bs + "partial_pullback n g m S" "drop n as" + "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p "n + m" "k - n" n ] 4 + by (metis (no_types, lifting) "3" append.assoc append_take_drop_id + asbs_def assms(2) cartesian_power_car_memE less_imp_le_nat partial_pullback_memE(1) + subsetI take_closed) + qed + show "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ partial_pullback k (g \ take n) m S" + proof fix x assume A: "x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" + show "x \ partial_pullback k (g \ take n) m S" + proof(rule partial_pullback_memI) + have 0: "(partial_pullback n g m S) \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" + using partial_pullback_closed by blast + then have "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ carrier (Q\<^sub>p\<^bsup>n + m + (k - n)\<^esup>)" + using assms A split_cartesian_product_closed[of "partial_pullback n g m S" Q\<^sub>p "n + m" + "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" "k - n" n] + using le_add1 by blast + then show P: "x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)" + by (smt A Nat.add_diff_assoc2 add.commute add_diff_cancel_left' assms(2) le_add1 less_imp_le_nat subsetD) + have "take n x @ drop (n + (k - n)) x \ partial_pullback n g m S" + using 0 A split_cartesian_product_memE[of x "n + m" "k - n" n "partial_pullback n g m S" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p] + le_add1 by blast + have 1: "g (take n x) # drop k x \ S" + using partial_pullback_memE + by (metis (no_types, lifting) \take n x @ drop (n + (k - n)) x \ partial_pullback n g m S\ + \x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)\ add.assoc assms(2) cartesian_power_drop le_add1 + le_add_diff_inverse less_imp_le_nat partial_pullback_memE' take_closed) + have 2: "g (take n x) = (g \ take n) (take k x)" + using assms P comp_apply[of g "take n" "take k x"] + by (metis add.commute append_same_eq append_take_drop_id less_imp_add_positive take_add take_drop) + then show "(g \ take n) (take k x) # drop k x \ S" + using "1" by presburger + qed + qed +qed + +lemma comp_take_is_semialg: + assumes "is_semialg_function n g" + assumes "n < k" + assumes "0 < n" + shows "is_semialg_function k (g \ (take n))" +proof(rule is_semialg_functionI) + show "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" + using assms function_restriction[of g n "carrier Q\<^sub>p" k] dual_order.strict_implies_order + is_semialg_function_closed + by blast + show "\ka S. S \ semialg_sets (1 + ka) \ is_semialgebraic (k + ka) (partial_pullback k (g \ take n) ka S)" + proof- fix l S assume A: "S \ semialg_sets (1 + l)" + have 0: "is_semialgebraic (n + l) (partial_pullback n g l S) " + using assms A is_semialg_functionE is_semialgebraicI + by blast + have "is_semialgebraic (n + l + (k - n)) (split_cartesian_product (n + l) (k - n) n (partial_pullback n g l S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)))" + using A 0 split_cartesian_product_is_semialgebraic[of _ _ + "partial_pullback n g l S" _ "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)"] + add_gr_0 assms(2) assms(3) carrier_is_semialgebraic le_add1 zero_less_diff + by presburger + then show "is_semialgebraic (k + l) (partial_pullback k (g \ take n) l S)" + using partial_pullback_restriction[of g n k l S] + by (metis (no_types, lifting) add.assoc add.commute assms(1) assms(2) is_semialg_function_closed le_add_diff_inverse less_imp_le_nat) + qed +qed + +text\Restriction of a graph to a semialgebraic domain\ + +lemma graph_formula: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "graph n g = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). g (take n as) = as!n}" + using assms graph_memI fun_graph_def[of Q\<^sub>p n g] + by (smt Collect_cong Suc_eq_plus1 graph_memE(1) graph_mem_closed mem_Collect_eq) + +definition restricted_graph where +"restricted_graph n g S = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). take n as \ S \ g (take n as) = as!n }" + +lemma restricted_graph_closed: + "restricted_graph n g S \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + by (metis (no_types, lifting) mem_Collect_eq restricted_graph_def subsetI) + +lemma restricted_graph_memE: + assumes "a \ restricted_graph n g S" + shows "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" "take n a \ S" "g (take n a) = a!n" + using assms + using restricted_graph_closed apply blast + apply (metis (no_types, lifting) assms mem_Collect_eq restricted_graph_def) + using assms unfolding restricted_graph_def + by blast + +lemma restricted_graph_mem_formula: + assumes "a \ restricted_graph n g S" + shows "a = (take n a)@[g (take n a)]" +proof- + have "length a = Suc n" + using assms + by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq restricted_graph_def) + then have "a = (take n a)@[a!n]" + by (metis append_eq_append_conv_if hd_drop_conv_nth lessI take_hd_drop) + then show ?thesis + by (metis assms restricted_graph_memE(3)) +qed + +lemma restricted_graph_memI: + assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + assumes "take n a \ S" + assumes "g (take n a) = a!n" + shows "a \ restricted_graph n g S" + using assms restricted_graph_def + by blast + +lemma restricted_graph_memI': + assumes "a \ S" + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(a@[g a]) \ restricted_graph n g S" +proof- + have "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms(1) assms(3) by blast + then have "g a \ carrier Q\<^sub>p" + using assms by blast + then have 0: "a @ [g a] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + using assms + by (metis (no_types, lifting) add.commute cartesian_power_append plus_1_eq_Suc subsetD) + have 1: "take n (a @ [g a]) \ S" + using assms + by (metis (no_types, lifting) append_eq_conv_conj cartesian_power_car_memE subsetD) + show ?thesis + using assms restricted_graph_memI[of "a@[g a]" n S g] + by (metis "0" \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ append_eq_conv_conj cartesian_power_car_memE nth_append_length) +qed + +lemma restricted_graph_subset: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "restricted_graph n g S \ graph n g" +proof fix x assume A: "x \ restricted_graph n g S" + show "x \ graph n g" + apply(rule graph_memI) + using assms(1) apply blast + using A restricted_graph_memE(3) apply blast + by (metis A add.commute plus_1_eq_Suc restricted_graph_memE(1)) +qed + +lemma restricted_graph_subset': + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "restricted_graph n g S \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" +proof fix a assume A: "a \ restricted_graph n g S" + then have "a = (take n a)@[g (take n a)]" + using restricted_graph_mem_formula by blast + then show "a \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + using cartesian_product_memI' A unfolding restricted_graph_def + by (metis (mono_tags, lifting) assms(2) last_closed' mem_Collect_eq subsetI Qp.to_R1_closed) +qed + +lemma restricted_graph_intersection: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" +proof + show "restricted_graph n g S \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + using assms restricted_graph_subset restricted_graph_subset' + by (meson Int_subset_iff) + show "graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ restricted_graph n g S" + proof fix x assume A: " x \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" + show "x \ restricted_graph n g S" + apply(rule restricted_graph_memI) + using A graph_memE[of g n x] + apply (metis (no_types, lifting) Int_iff add.commute assms(1) graph_mem_closed plus_1_eq_Suc) + using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] + using assms(2) apply blast + using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] + using assms(1) by blast + qed + qed + +lemma restricted_graph_is_semialgebraic: + assumes "is_semialg_function n g" + assumes "is_semialgebraic n S" + shows "is_semialgebraic (n+1) (restricted_graph n g S)" +proof- + have 0: "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" + using assms is_semialg_function_closed is_semialgebraic_closed + restricted_graph_intersection by presburger + have 1: "is_semialgebraic (n + 1) (graph n g)" + using assms semialg_graph + by blast + have 2: "is_semialgebraic (n + 1) (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" + using cartesian_product_is_semialgebraic[of n S 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] assms + carrier_is_semialgebraic less_one + by presburger + then show ?thesis + using 0 1 2 intersection_is_semialg[of "n+1" "graph n g" "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] + by presburger +qed + +lemma take_closed: + assumes "n \ k" + assumes "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + shows "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms take_closed + by blast + +lemma take_compose_closed: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "n < k" + shows "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" +proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + then have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using assms less_imp_le_nat take_closed + by blast + then have "g (take n x) \ carrier Q\<^sub>p" + using assms(1) by blast + then show "(g \ take n) x \ carrier Q\<^sub>p" + using comp_apply[of g "take n" x] + by presburger +qed + +lemma take_graph_formula: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "n < k" + assumes "0 < n" + shows "graph k (g \ (take n)) = {as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>). g (take n as) = as!k}" +proof- + have "\as. as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>) \ (g \ take n) (take k as) = g (take n as) " + using assms comp_apply take_take[of n k] + proof - + fix as :: "((nat \ int) \ (nat \ int)) set list" + show "(g \ take n) (take k as) = g (take n as)" + by (metis (no_types) \n < k\ comp_eq_dest_lhs min.strict_order_iff take_take) + qed + then show ?thesis + using take_compose_closed[of g n k] assms comp_apply[of g "take n"] graph_formula[of "g \ (take n)" k] + by (smt Collect_cong Suc_eq_plus1) +qed + +lemma graph_memI': + assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + assumes "take n a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "g (take n a) = a!n" + shows "a \ graph n g" + using assms fun_graph_def[of Q\<^sub>p n g] + by (smt cartesian_power_car_memE eq_imp_le lessI mem_Collect_eq take_Suc_conv_app_nth take_all) + +lemma graph_memI'': + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "(a@[g a]) \ graph n g " + using assms fun_graph_def + by blast + +lemma graph_as_restricted_graph: + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "graph n f = restricted_graph n f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + apply(rule equalityI) + apply (metis Suc_eq_plus1 assms graph_memE(1) graph_memE(3) graph_mem_closed restricted_graph_memI subsetI) + by (simp add: assms restricted_graph_subset) + +definition double_graph where +"double_graph n f g = {as \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>). f (take n as) = as!n \ g (take n as) = as!(n + 1)}" + +lemma double_graph_rep: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + shows "double_graph n f g = restricted_graph (n + 1) (g \ take n) (graph n f)" +proof + show "double_graph n f g \ restricted_graph (n + 1) (g \ take n) (graph n f)" + proof fix x assume A: "x \ double_graph n f g" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>) \ f (take n x) = x!n \ g (take n x) = x!(n + 1)" + using double_graph_def by blast + have 1: "take (n+1) x \ graph n f" + apply(rule graph_memI) + using assms(2) apply blast + apply (metis "0" append_eq_conv_conj cartesian_power_car_memE le_add1 length_take + less_add_same_cancel1 less_numeral_extra(1) min.absorb2 nth_take take_add) + by (metis (no_types, opaque_lifting) "0" Suc_eq_plus1 Suc_n_not_le_n add_cancel_right_right + dual_order.antisym le_iff_add not_less_eq_eq one_add_one plus_1_eq_Suc take_closed) + show " x \ restricted_graph (n + 1) (g \ take n) (graph n f)" + apply(rule restricted_graph_memI) + apply (metis "0" One_nat_def add_Suc_right numeral_2_eq_2) + using "1" apply blast + using 0 take_take[of n "n + 1" x] comp_apply + by (metis le_add1 min.absorb1) + qed + show "restricted_graph (n + 1) (g \ take n) (graph n f) \ double_graph n f g" + proof fix x + assume A: "x \ restricted_graph (n + 1) (g \ take n) (graph n f)" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc (n + 1)\<^esup>) \ take (n + 1) x \ graph n f \ (g \ take n) (take (n + 1) x) = x ! (n + 1)" + using restricted_graph_memE[of x "n+1" "(g \ take n)" "graph n f" ] + by blast + then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" + using 0 + by (metis Suc_1 add_Suc_right) + have 2: " f (take n x) = x ! n" + using 0 take_take[of n "n + 1" x] graph_memE[of f n "take (n + 1) x"] + by (metis assms(2) le_add1 less_add_same_cancel1 less_numeral_extra(1) min.absorb1 nth_take) + have 3: "g (take n x) = x ! (n + 1)" + using 0 comp_apply take_take[of n "n + 1" x] + by (metis le_add1 min.absorb1) + then show "x \ double_graph n f g" + unfolding double_graph_def using 1 2 3 + by blast + qed +qed + +lemma double_graph_is_semialg: + assumes "n > 0" + assumes "is_semialg_function n f" + assumes "is_semialg_function n g" + shows "is_semialgebraic (n+2) (double_graph n f g)" + using double_graph_rep[of g n f] assms restricted_graph_is_semialgebraic[of n "g \ take n" "graph n f"] + by (metis (no_types, lifting) Suc_eq_plus1 add_Suc_right is_semialg_function_closed + less_add_same_cancel1 less_numeral_extra(1) one_add_one restricted_graph_is_semialgebraic + comp_take_is_semialg semialg_graph) + +definition add_vars :: "nat \ nat \ padic_tuple \ padic_number" where +"add_vars i j as = as!i \\<^bsub>Q\<^sub>p\<^esub> as!j" + +lemma add_vars_rep: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "i < n" + assumes "j < n" + shows "add_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" + unfolding add_vars_def + using assms eval_at_point_add[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] + eval_pvar by (metis pvar_closed) + +lemma add_vars_is_semialg: + assumes "i < n" + assumes "j < n" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "is_semialg_function n (add_vars i j)" +proof- + have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms pvar_closed[of ] + by blast + then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" + using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] + by blast + then show ?thesis + using assms add_vars_rep + semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "add_vars i j" ] + by (metis (no_types, lifting) restrict_ext) +qed + +definition mult_vars :: "nat \ nat \ padic_tuple \ padic_number" where +"mult_vars i j as = as!i \ as!j" + +lemma mult_vars_rep: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "i < n" + assumes "j < n" + shows "mult_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" + unfolding mult_vars_def + using assms eval_at_point_mult[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] + eval_pvar[of i n as] eval_pvar[of j n as ] + by (metis pvar_closed) + +lemma mult_vars_is_semialg: + assumes "i < n" + assumes "j < n" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "is_semialg_function n (mult_vars i j)" +proof- + have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms pvar_closed[of ] + by blast + then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" + using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] + by blast + then show ?thesis + using assms mult_vars_rep + semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "mult_vars i j" ] + by (metis (no_types, lifting) restrict_ext) +qed + +definition minus_vars :: "nat \ padic_tuple \ padic_number" where +"minus_vars i as = \\<^bsub>Q\<^sub>p\<^esub> as!i" + +lemma minus_vars_rep: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "i < n" + shows "minus_vars i as = Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)) as" + unfolding minus_vars_def + using assms eval_pvar[of i n as] eval_at_point_a_inv[of as n "pvar Q\<^sub>p i"] + by (metis pvar_closed) + +lemma minus_vars_is_semialg: + assumes "i < n" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "is_semialg_function n (minus_vars i)" +proof- + have 0: "pvar Q\<^sub>p i \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + using assms pvar_closed[of ] Qp.cring_axioms by presburger + have "is_semialg_function n (Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)))" + apply(rule poly_is_semialg ) + using "0" by blast + then show ?thesis + using assms minus_vars_rep[of a i n] + semialg_function_on_carrier[of n _ "minus_vars i" ] + by (metis (no_types, lifting) minus_vars_rep restrict_ext) +qed + +definition extended_graph where +"extended_graph n f g h = {as \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>). + f (take n as) = as!n \ g (take n as) = as! (n + 1) \ h [(f (take n as)),(g (take n as))] = as! (n + 2) }" + +lemma extended_graph_rep: +"extended_graph n f g h = restricted_graph (n + 2) (h \ (drop n)) (double_graph n f g)" +proof + show "extended_graph n f g h \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" + proof fix x + assume "x \ extended_graph n f g h" + then have A: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \f (take n x) = x!n \ g (take n x) = x! (n + 1) \ + h [(f (take n x)),(g (take n x))] = x! (n + 2)" + unfolding extended_graph_def by blast + then have 0: "take (n + 2) x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" + proof - + have "Suc (Suc n) \ n + numeral (num.One + num.Bit0 num.One)" + by simp + then show ?thesis + by (metis (no_types) \x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \ f (take n x) = x ! n \ g (take n x) = x ! (n + 1) \ h [f (take n x), g (take n x)] = x ! (n + 2)\ add_2_eq_Suc' add_One_commute semiring_norm(5) take_closed) + qed + have 1: "f (take n (take (n + 2) x)) = (take (n + 2) x) ! n" + using A + by (metis Suc_1 add.commute append_same_eq append_take_drop_id + less_add_same_cancel1 nth_take take_add take_drop zero_less_Suc) + have 2: " g (take n (take (n + 2) x)) = (take (n + 2) x) ! (n + 1)" + using A + by (smt add.assoc add.commute append_same_eq append_take_drop_id less_add_same_cancel1 + less_numeral_extra(1) nth_take one_add_one take_add take_drop) + then have 3: "take (n + 2) x \ double_graph n f g" + unfolding double_graph_def + using 0 1 2 + by blast + have 4: "drop n (take (n + 2) x) = [(f (take n x)),(g (take n x))]" + proof- + have 40: "take (n + 2) x ! (n + 1) = x! (n + 1)" + by (metis add.commute add_2_eq_Suc' lessI nth_take plus_1_eq_Suc) + have 41: "take (n + 2) x ! n = x! n" + by (metis Suc_1 less_SucI less_add_same_cancel1 less_numeral_extra(1) nth_take) + have 42: "take (n + 2) x ! (n + 1) = g (take n x)" + using 40 A + by blast + have 43: "take (n + 2) x ! n = f (take n x)" + using 41 A + by blast + show ?thesis using A 42 43 + by (metis "0" add_cancel_right_right cartesian_power_car_memE cartesian_power_drop + le_add_same_cancel1 nth_drop pair_id zero_le_numeral) + qed + then have 5: "(h \ drop n) (take (n + 2) x) = x ! (n + 2)" + using 3 A + by (metis add_2_eq_Suc' comp_eq_dest_lhs) + show "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" + using restricted_graph_def A 3 5 + by (metis (no_types, lifting) One_nat_def Suc_1 + add_Suc_right numeral_3_eq_3 restricted_graph_memI) + qed + show "restricted_graph (n + 2) (h \ drop n) (double_graph n f g) \ extended_graph n f g h" + proof fix x assume A: "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" + then have 0: "take (n+2) x \ double_graph n f g" + using restricted_graph_memE(2) by blast + have 1: "(h \ drop n) (take (n+2) x) = x! (n+2) " + by (meson A restricted_graph_memE(3) padic_fields_axioms) + have 2: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>)" + using A + by (metis (no_types, opaque_lifting) Suc3_eq_add_3 add.commute add_2_eq_Suc' + restricted_graph_closed subsetD) + have 3: "length x = n + 3" + using "2" cartesian_power_car_memE by blast + have 4: "drop n (take (n+2) x) = [x!n, x!(n+1)]" + proof- + have "length (take (n+2) x) = n+2" + by (simp add: "3") + then have 40:"length (drop n (take (n+2) x)) = 2" + by (metis add_2_eq_Suc' add_diff_cancel_left' length_drop) + have 41: "(drop n (take (n+2) x))!0 = x!n" + using 3 + by (metis Nat.add_0_right \length (take (n + 2) x) = n + 2\ add_gr_0 le_add1 less_add_same_cancel1 less_numeral_extra(1) nth_drop nth_take one_add_one) + have 42: "(drop n (take (n+2) x))!1 = x!(n+1)" + using 3 nth_take nth_drop A + by (metis add.commute le_add1 less_add_same_cancel1 less_numeral_extra(1) one_add_one take_drop) + show ?thesis + using 40 41 42 + by (metis pair_id) + qed + have "(take n x) = take n (take (n+2) x)" + using take_take 3 + by (metis le_add1 min.absorb1) + then have 5: "f (take n x) = x ! n" + using 0 double_graph_def[of n f g] 3 + by (smt Suc_1 less_add_same_cancel1 mem_Collect_eq nth_take zero_less_Suc) + have 6: "g (take n x) = x ! (n + 1) " + using 0 double_graph_def[of n f g] 3 take_take[of n "n+2" x] + by (smt Suc_1 \take n x = take n (take (n + 2) x)\ add_Suc_right lessI mem_Collect_eq nth_take) + have 7: " h [f (take n x), g (take n x)] = x ! (n + 2)" + using 4 A comp_apply + by (metis "1" "5" "6") + show " x \ extended_graph n f g h" + unfolding extended_graph_def + using 2 5 6 7 A + by blast + qed +qed + +lemma function_tuple_eval_closed: + assumes "is_function_tuple Q\<^sub>p n fs" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "function_tuple_eval Q\<^sub>p n fs x \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" + using function_tuple_eval_closed[of Q\<^sub>p n fs x] assms by blast + +definition k_graph where +"k_graph n fs = {x \ carrier (Q\<^sub>p\<^bsup>n + length fs\<^esup>). x = (take n x)@ (function_tuple_eval Q\<^sub>p n fs (take n x)) }" + +lemma k_graph_memI: + assumes "is_function_tuple Q\<^sub>p n fs" + assumes "x = as@function_tuple_eval Q\<^sub>p n fs as" + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "x \ k_graph n fs" +proof- + have "take n x = as" + using assms + by (metis append_eq_conv_conj cartesian_power_car_memE) + then show ?thesis unfolding k_graph_def using assms + by (smt append_eq_conv_conj cartesian_power_car_memE cartesian_power_car_memI'' length_append local.function_tuple_eval_closed mem_Collect_eq) +qed + +text\composing a function with a function tuple\ + +lemma Qp_function_tuple_comp_closed: + assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "length fs = n" + assumes "is_function_tuple Q\<^sub>p m fs" + shows "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using assms function_tuple_comp_closed + by blast + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Tuples of Semialgebraic Functions\ + (********************************************************************************************) + (********************************************************************************************) + +text\Predicate for a tuple of semialgebraic functions\ + +definition is_semialg_function_tuple where +"is_semialg_function_tuple n fs = (\ f \ set fs. is_semialg_function n f)" + +lemma is_semialg_function_tupleI: + assumes "\ f. f \ set fs \ is_semialg_function n f" + shows "is_semialg_function_tuple n fs" + using assms is_semialg_function_tuple_def + by blast + +lemma is_semialg_function_tupleE: + assumes "is_semialg_function_tuple n fs" + assumes "i < length fs" + shows "is_semialg_function n (fs ! i)" + by (meson assms(1) assms(2) in_set_conv_nth is_semialg_function_tuple_def padic_fields_axioms) + +lemma is_semialg_function_tupleE': + assumes "is_semialg_function_tuple n fs" + assumes "f \ set fs" + shows "is_semialg_function n f" + using assms(1) assms(2) is_semialg_function_tuple_def + by blast + +lemma semialg_function_tuple_is_function_tuple: + assumes "is_semialg_function_tuple n fs" + shows "is_function_tuple Q\<^sub>p n fs" + apply(rule is_function_tupleI) + using assms is_semialg_function_closed is_semialg_function_tupleE' by blast + +lemma const_poly_function_tuple_comp_is_semialg: + assumes "n > 0" + assumes "is_semialg_function_tuple n fs" + assumes "a \ carrier Q\<^sub>p" + shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (Qp_to_IP a))" + apply(rule semialg_function_on_carrier[of n "Qp_ev (Qp_to_IP a)"]) + using poly_is_semialg[of "(Qp_to_IP a)"] + using assms(1) assms(3) Qp_to_IP_car apply blast + using poly_function_tuple_comp_eq[of n fs "(Qp_to_IP a)"] assms unfolding restrict_def + by (metis (no_types, opaque_lifting) eval_at_point_const poly_function_tuple_comp_constant semialg_function_tuple_is_function_tuple) + +lemma pvar_poly_function_tuple_comp_is_semialg: + assumes "n > 0" + assumes "is_semialg_function_tuple n fs" + assumes "i < length fs" + shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (pvar Q\<^sub>p i))" + apply(rule semialg_function_on_carrier[of n "fs!i"]) + using assms(2) assms(3) is_semialg_function_tupleE apply blast + by (metis assms(2) assms(3) poly_function_tuple_comp_pvar + restrict_ext semialg_function_tuple_is_function_tuple) + +text\Polynomial functions with semialgebraic coefficients\ + +definition point_to_univ_poly :: "nat \ padic_tuple \ padic_univ_poly" where +"point_to_univ_poly n a = ring_cfs_to_univ_poly n a" + +definition tuple_partial_image where +"tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs (take m x))@(drop m x)" + +lemma tuple_partial_image_closed: + assumes "length fs > 0" + assumes "is_function_tuple Q\<^sub>p n fs" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + shows "tuple_partial_image n fs x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" + using assms unfolding tuple_partial_image_def + by (meson cartesian_power_concat(1) cartesian_power_drop + function_tuple_eval_closed le_add1 take_closed) + +lemma tuple_partial_image_indices: + assumes "length fs > 0" + assumes "is_function_tuple Q\<^sub>p n fs" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + assumes "i < length fs" + shows "(tuple_partial_image n fs x) ! i = (fs!i) (take n x)" +proof- + have 0: "(function_tuple_eval Q\<^sub>p n fs (take n x)) ! i = (fs!i) (take n x)" + using assms unfolding function_tuple_eval_def + by (meson nth_map) + have 1: "length (function_tuple_eval Q\<^sub>p n fs (take n x)) > i" + by (metis assms(4) function_tuple_eval_def length_map) + show ?thesis + using 0 1 assms unfolding tuple_partial_image_def + by (metis nth_append) +qed + +lemma tuple_partial_image_indices': + assumes "length fs > 0" + assumes "is_function_tuple Q\<^sub>p n fs" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + assumes "i < l" + shows "(tuple_partial_image n fs x) ! (length fs + i) = x!(n + i)" + using assms unfolding tuple_partial_image_def + by (metis (no_types, lifting) cartesian_power_car_memE function_tuple_eval_closed le_add1 + nth_append_length_plus nth_drop take_closed) + +definition tuple_partial_pullback where +"tuple_partial_pullback n fs l S = ((tuple_partial_image n fs)-`S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + +lemma tuple_partial_pullback_memE: + assumes "as \ tuple_partial_pullback m fs l S" + shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "tuple_partial_image m fs as \ S" + using assms + apply (metis (no_types, opaque_lifting) Int_iff add.commute tuple_partial_pullback_def) + using assms unfolding tuple_partial_pullback_def + by blast + +lemma tuple_partial_pullback_closed: +"tuple_partial_pullback m fs l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + using tuple_partial_pullback_memE by blast + +lemma tuple_partial_pullback_memI: + assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" + assumes "is_function_tuple Q\<^sub>p m fs" + assumes "((function_tuple_eval Q\<^sub>p m fs) (take m as))@(drop m as) \ S" + shows "as \ tuple_partial_pullback m fs k S" + using assms unfolding tuple_partial_pullback_def tuple_partial_image_def + by blast + +lemma tuple_partial_image_eq: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + assumes "x = as @ bs" + shows "tuple_partial_image n fs x = ((function_tuple_eval Q\<^sub>p n fs) 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 assms 0 1 unfolding tuple_partial_image_def + by presburger +qed + +lemma tuple_partial_pullback_memE': + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + assumes "x = as @ bs" + assumes "x \ tuple_partial_pullback n fs k S" + shows "(function_tuple_eval Q\<^sub>p n fs as)@bs \ S" + using tuple_partial_pullback_memE[of x n fs k S] tuple_partial_image_def[of n fs x] + by (metis assms(1) assms(2) assms(3) assms(4) tuple_partial_image_eq) + +text\tuple partial pullbacks have the same algebraic properties as pullbacks\ + +lemma tuple_partial_pullback_intersect: +"tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" + unfolding tuple_partial_pullback_def + by blast + +lemma tuple_partial_pullback_union: +"tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" + unfolding tuple_partial_pullback_def + by blast + +lemma tuple_partial_pullback_complement: + assumes "is_function_tuple Q\<^sub>p m fs" + shows "tuple_partial_pullback m fs l ((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (tuple_partial_pullback m fs l S) " + apply(rule equalityI) + using tuple_partial_pullback_def[of m fs l "((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S)"] + tuple_partial_pullback_def[of m fs l S] + apply blast +proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - tuple_partial_pullback m fs l S" + show " x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>) - S) " + apply(rule tuple_partial_pullback_memI) + using A + apply blast + using assms + apply blast + proof + have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + by (meson A DiffD1 cartesian_power_drop) + have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A + by (meson DiffD1 le_add1 take_closed) + show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x + \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" + using 0 1 assms + using cartesian_power_concat(1) function_tuple_eval_closed by blast + show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ S" + using A unfolding tuple_partial_pullback_def tuple_partial_image_def + by blast + qed +qed + +lemma tuple_partial_pullback_carrier: + assumes "is_function_tuple Q\<^sub>p m fs" + shows "tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + apply(rule equalityI) + using tuple_partial_pullback_memE(1) apply blast +proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + show "x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>))" + apply(rule tuple_partial_pullback_memI) + using A cartesian_power_drop[of x m l] take_closed assms + apply blast + using assms apply blast + proof- + have "function_tuple_eval Q\<^sub>p m fs (take m x) \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" + using A take_closed assms + function_tuple_eval_closed le_add1 + by blast + then show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x + \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" + using cartesian_power_drop[of x m l] A cartesian_power_concat(1) + by blast + qed +qed + +definition is_semialg_map_tuple where +"is_semialg_map_tuple m fs = (is_function_tuple Q\<^sub>p m fs \ + (\l \ 0. \S \ semialg_sets ((length fs) + l). is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)))" + +lemma is_semialg_map_tuple_closed: + assumes "is_semialg_map_tuple m fs" + shows "is_function_tuple Q\<^sub>p m fs" + using is_semialg_map_tuple_def assms by blast + +lemma is_semialg_map_tupleE: + assumes "is_semialg_map_tuple m fs" + assumes "is_semialgebraic ((length fs) + l) S" + shows " is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)" + using is_semialg_map_tuple_def[of m fs] assms is_semialgebraicE[of "((length fs) + l)" S] + by blast + +lemma is_semialg_map_tupleI: + assumes "is_function_tuple Q\<^sub>p m fs" + assumes "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + shows "is_semialg_map_tuple m fs" + using assms unfolding is_semialg_map_tuple_def + by blast + +text\Semialgebraicity for maps can be verified on basic semialgebraic sets\ + +lemma is_semialg_map_tupleI': + assumes "is_function_tuple Q\<^sub>p m fs" + assumes "\k S. S \ basic_semialgs ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + shows "is_semialg_map_tuple m fs" + apply(rule is_semialg_map_tupleI) + using assms(1) apply blast +proof- + show "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + proof- fix k S assume A: "S \ semialg_sets ((length fs) + k)" + show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>length fs + k\<^esup>)" "basic_semialgs ((length fs) + k)"]) + using A unfolding semialg_sets_def + apply blast + using tuple_partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger + using assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg tuple_partial_pullback_carrier tuple_partial_pullback_intersect apply presburger + using tuple_partial_pullback_union union_is_semialgebraic apply presburger + using assms(1) complement_is_semialg tuple_partial_pullback_complement plus_1_eq_Suc by presburger + qed +qed + +text\ + The goal of this section is to show that tuples of semialgebraic functions are semialgebraic maps. +\ + +text\The function $(x_0, x, y) \mapsto (x_0, f(x), y)$\ + +definition twisted_partial_image where +"twisted_partial_image n m f xs = (take n xs)@ partial_image m f (drop n xs)" + +text\The set ${(x_0, x, y) \mid (x_0, f(x), y) \in S}$\ + +text\Convention: a function which produces a subset of (Qp (i + j +k)) will receive the 3 arity +parameters in sequence, at the very beginning of the function\ + +definition twisted_partial_pullback where +"twisted_partial_pullback n m l f S = ((twisted_partial_image n m f)-`S) \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" + +lemma twisted_partial_pullback_memE: + assumes "as \ twisted_partial_pullback n m l f S" + shows "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" "twisted_partial_image n m f as \ S" + using assms + apply (metis (no_types, opaque_lifting) Int_iff add.commute twisted_partial_pullback_def subset_iff) + using assms unfolding twisted_partial_pullback_def + by blast + +lemma twisted_partial_pullback_closed: +"twisted_partial_pullback n m l f S \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" + using twisted_partial_pullback_memE(1) by blast + +lemma twisted_partial_pullback_memI: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" + assumes "(take n as)@((f (take m (drop n as)))#(drop (n + m) as)) \ S" + shows "as \ twisted_partial_pullback n m l f S" + using assms unfolding twisted_partial_pullback_def twisted_partial_image_def + by (metis (no_types, lifting) IntI add.commute drop_drop local.partial_image_def vimageI) + +lemma twisted_partial_image_eq: + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + assumes "x = as @ bs @ cs" + shows "twisted_partial_image n m f x = as@((f bs)#cs)" +proof- + have 0: "(take n x) = as" + by (metis append_eq_conv_conj assms(1) assms(4) + cartesian_power_car_memE) + have 1: "twisted_partial_image n m f x = as@(partial_image m f (bs@cs))" + using 0 assms twisted_partial_image_def + by (metis append_eq_conv_conj cartesian_power_car_memE) + have 2: "(partial_image m f (bs@cs)) = (f bs)#cs" + using partial_image_eq[of bs m cs l "bs@cs" f] assms + by blast + show ?thesis using assms 0 1 2 + by (metis ) +qed + +lemma twisted_partial_pullback_memE': + assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + assumes "x = as @ bs @ cs" + assumes "x \ twisted_partial_pullback n m l f S" + shows "as@((f bs)#cs) \ S" + by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) + twisted_partial_image_eq twisted_partial_pullback_memE(2)) + +text\partial pullbacks have the same algebraic properties as pullbacks\ + +text\permutation which moves the entry at index \i\ to 0\ + +definition twisting_permutation where +"twisting_permutation (i::nat) = (\j. if j < i then j + 1 else + (if j = i then 0 else j))" + +lemma twisting_permutation_permutes: + assumes "i < n" + shows "twisting_permutation i permutes {..x. x > i \ twisting_permutation i x = x" + unfolding twisting_permutation_def + by auto + have 1: "(\x. x \ {.. twisting_permutation i x = x)" + using 0 assms + by auto + have 2: "(\y. \!x. twisting_permutation i x = y)" + proof fix y + show " \!x. twisting_permutation i x = y" + proof(cases "y = 0") + case True + show "\!x. twisting_permutation i x = y" + by (metis Suc_eq_plus1 True add_eq_0_iff_both_eq_0 less_nat_zero_code + nat_neq_iff twisting_permutation_def zero_neq_one) + next + case False + show ?thesis + proof(cases "y \i") + case True + show ?thesis + proof + show "twisting_permutation i (y - 1) = y" + using True + by (metis False add.commute add_diff_inverse_nat diff_less gr_zeroI le_eq_less_or_eq + less_imp_diff_less less_one twisting_permutation_def) + show "\x. twisting_permutation i x = y \ x = y - 1" + using True False twisting_permutation_def by force + qed + next + case F: False + then show ?thesis using False + unfolding twisting_permutation_def + by (metis add_leD1 add_leD2 add_le_same_cancel2 discrete le_numeral_extra(3) + less_imp_not_less ) + qed + qed + qed + show ?thesis + using 1 2 + by (simp add: permutes_def) +qed + +lemma twisting_permutation_action: + assumes "length as = i" + shows "permute_list (twisting_permutation i) (b#(as@bs)) = as@(b#bs)" +proof- + have 0: "length (permute_list (twisting_permutation i) (b#(as@bs))) = length (as@(b#bs))" + by (metis add.assoc length_append length_permute_list list.size(4)) + have "\j. j < length (as@(b#bs)) + \ (permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" + proof- + fix j assume A: "j < length (as@(b#bs))" + show "(permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" + proof(cases "j < i") + case True + then have T0: "twisting_permutation i j = j + 1" + using twisting_permutation_def by auto + then have T1: "(b # as @ bs) ! twisting_permutation i j = as!j" + using assms + by (simp add: assms True nth_append) + have T2: "(permute_list (twisting_permutation i) (b # as @ bs)) ! j = as!j" + proof- + have "twisting_permutation i permutes {..length as = i\) + then have "permute_list (fun_inv TI) (as@(b#bs)) = + permute_list ((fun_inv TI) \ TI) (b#(as@bs))" + using 0 1 + by (metis TI_def fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv + length_permute_list permute_list_compose) + then show ?thesis + by (metis "0" Nil_is_append_conv TI_def fun_inv_permute(3) + length_greater_0_conv list.distinct(1) permute_list_id) +qed + +lemma twisting_semialg: + assumes "is_semialgebraic n S" + assumes "n > i" + shows "is_semialgebraic n ((permute_list ((twisting_permutation i)) ` S))" +proof- + obtain TI where TI_def: "TI = twisting_permutation i" + by blast + have 0: "TI permutes {..<(n::nat)}" + using assms TI_def twisting_permutation_permutes[of i n] + by blast + have "(TI) permutes {.. i" + shows "is_semialgebraic n ((permute_list (fun_inv (twisting_permutation i)) ` S))" +proof- + obtain TI where TI_def: "TI = twisting_permutation i" + by blast + have 0: "TI permutes {..<(n::nat)}" + using assms TI_def twisting_permutation_permutes[of i n] + by blast + have "(fun_inv TI) permutes {..Defining a permutation that does: $(x0, x1, y) \mapsto (x_1, x0, y)$\ + +definition tp_1 where +"tp_1 i j = (\ n. (if n n \ n < i + j then n - i else + n)))" +lemma permutes_I: + assumes "\x. x \ S \ f x = x" + assumes "\y. y \ S \ \!x \ S. f x = y" + assumes "\x. x \ S \ f x \ S" + shows "f permutes S" +proof- + have 0 : "(\x. x \ S \ f x = x) " + using assms(1) by blast + have 1: "(\y. \!x. f x = y)" + proof fix y + show "\!x. f x = y" + apply(cases "y \ S") + apply (metis "0" assms(2)) + proof- + assume "y \ S" + then show "\!x. f x = y" + by (metis assms(1) assms(3)) + qed + qed + show ?thesis + using assms 1 + unfolding permutes_def + by blast +qed + +lemma tp_1_permutes: +"(tp_1 (i::nat) j) permutes {..< i + j}" +proof(rule permutes_I) + show "\x. x \ {.. tp_1 i j x = x" + proof- fix x assume A: "x \ {..y. y \ {.. \!x. x \ {.. tp_1 i j x = y" + proof- fix y assume A: "y \ {..!x. x \ {.. tp_1 i j x = y" + proof(cases "y < j") + case True + then have 0:"tp_1 i j (y + i) = y" + by (simp add: tp_1_def) + have 1: "\x. x \ y + i \ tp_1 i j x \ y" + proof- fix x assume A: " x \ y + i" + show "tp_1 i j x \ y" + apply(cases "x < j") + apply (metis A True add.commute le_add_diff_inverse le_eq_less_or_eq nat_neq_iff not_add_less1 tp_1_def trans_less_add2) + by (metis A True add.commute le_add_diff_inverse less_not_refl tp_1_def trans_less_add1) + qed + show ?thesis using 0 1 + by (metis A \\x. x \ {.. tp_1 i j x = x\) + next + case False + then have "y - j < i" + using A by auto + then have "tp_1 i j (y - j) = y" + using False tp_1_def + by (simp add: tp_1_def) + then show ?thesis + by (smt A False \\x. x \ {.. tp_1 i j x = x\ + add.commute add_diff_inverse_nat add_left_imp_eq + less_diff_conv2 not_less tp_1_def + padic_fields_axioms) + qed + qed + show "\x. x \ {.. tp_1 i j x \ {.. {.. carrier (Q\<^sub>p\<^bsup>i\<^esup>)" + assumes "b \ carrier (Q\<^sub>p\<^bsup>j\<^esup>)" + assumes "c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "permute_list (tp_1 i j) (b@a@c)= a@b@c" +proof- + have 0:"length (permute_list (tp_1 i j) (b@a@c))= length (a@b@c)" + by (metis add.commute append.assoc length_append length_permute_list) + have "\x. x < length (a@b@c) \ (permute_list (tp_1 i j) (b@a@c)) ! x= (a@b@c) ! x" + proof- fix x assume A: "x < length (a@b@c)" + have B: "length (a @ b @ c) = i + j + length c" + using add.assoc assms(1) assms(2) cartesian_power_car_memE length_append + by metis + have C: "tp_1 i j permutes {..length (permute_list (tp_1 i j) (b @ a @ c)) = + length (a @ b @ c)\ length_permute_list nth_append permute_list_nth) + next + case False + show ?thesis + proof(cases "x < i + j") + case True + then have T0: "(tp_1 i j x) = x - i" + by (meson False not_less tp_1_def) + have "x - i < length b" + using E False True by linarith + then have T1: "permute_list (tp_1 i j) (b@ a @ c) ! x = b!(x-i)" + using nth_append + by (metis A C T0 \length (permute_list (tp_1 i j) (b @ a @ c)) = length (a @ b @ c)\ + length_permute_list permute_list_nth) + then show ?thesis + by (metis D False \x - i < length b\ nth_append) + next + case False + then have "(tp_1 i j x) = x" + by (meson tp_1_def trans_less_add1) + then show ?thesis + by (smt A C D E False add.commute add_diff_inverse_nat append.assoc + length_append nth_append_length_plus permute_list_nth) + qed + qed +qed + then show ?thesis + using 0 + by (metis list_eq_iff_nth_eq) +qed + +definition tw where +"tw i j = permute_list (tp_1 j i)" + +lemma tw_is_semialg: + assumes "n > 0" + assumes "is_semialgebraic n S" + assumes "n \ i + j" + shows "is_semialgebraic n ((tw i j)`S)" + unfolding tw_def + using assms tp_1_permutes'[of j i "n - (j + i)"] + permutation_is_semialgebraic[of n S] + by (metis add.commute le_add_diff_inverse) + +lemma twisted_partial_pullback_factored: + assumes "f \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ carrier Q\<^sub>p" + assumes "S \ carrier (Q\<^sub>p\<^bsup>n+1+ l\<^esup>)" + assumes "Y = partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" + shows "twisted_partial_pullback n m l f S = (tw m n) ` Y" +proof + show "twisted_partial_pullback n m l f S \ tw m n ` Y" + proof fix x + assume A: "x \ twisted_partial_pullback n m l f S" + then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" + using twisted_partial_pullback_memE(1) by blast + obtain a where a_def: "a = take n x" + by blast + obtain b where b_def: "b = take m (drop n x)" + by blast + obtain c where c_def: "c = (drop (n + m) x)" + by blast + have x_eq:"x = a@(b@c)" + by (metis a_def append.assoc append_take_drop_id b_def c_def take_add) + have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (metis (no_types, lifting) a_def dual_order.trans le_add1 take_closed x_closed) + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + proof- + have "drop n x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + by (metis (no_types, lifting) add.assoc cartesian_power_drop x_closed) + then show ?thesis + using b_def le_add1 take_closed by blast + qed + have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + using c_def cartesian_power_drop x_closed by blast + have B: "a@((f b)#c) \ S" + using A twisted_partial_pullback_memE' + by (smt a_closed a_def add.commute append_take_drop_id b_closed + b_def c_closed c_def drop_drop) + have "permute_list (fun_inv (twisting_permutation n)) (a@((f b)#c)) = (f b)#(a@c)" + using assms twisting_permutation_action'[of a n "f b" c] + a_closed cartesian_power_car_memE + by blast + then have C: "(f b)#(a@c) \ (permute_list (fun_inv (twisting_permutation n)) ` S)" + by (metis B image_eqI) + have C: "b@(a@c) \ partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" + proof(rule partial_pullback_memI) + show "b @ a @ c \ carrier (Q\<^sub>p\<^bsup>m + (n + l)\<^esup>)" + using a_closed b_closed c_closed cartesian_power_concat(1) + by blast + have 0: "(take m (b @ a @ c)) = b" + by (metis append.right_neutral b_closed cartesian_power_car_memE + diff_is_0_eq diff_self_eq_0 take0 take_all take_append) + have 1: "drop m (b @ a @ c) = a@c" + by (metis "0" append_take_drop_id same_append_eq) + show "f (take m (b @ a @ c)) # drop m (b @ a @ c) \ permute_list (fun_inv (twisting_permutation n)) ` S" + using 0 1 C + by presburger + qed + have D: "tw m n (b@(a@c)) = a@(b@c)" + using assms tw_def a_closed b_closed c_closed + by (metis tp_1_permutation_action x_eq) + then show "x \ tw m n ` Y" + using x_eq C assms + by (metis image_eqI) + qed + show "tw m n ` Y \ twisted_partial_pullback n m l f S" + proof fix x + assume A: "x \ tw m n ` Y" + then obtain y where y_def: "x = tw m n y \ y \ Y" + by blast + obtain as where as_def: "as \ (permute_list (fun_inv (twisting_permutation n)) ` S) \ + as = partial_image m f y" + using partial_pullback_memE + by (metis assms(3) y_def) + obtain s where s_def: "s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as" + using as_def by blast + obtain b where b_def: "b = take m y" + by blast + obtain a where a_def: "a = take n (drop m y)" + by blast + obtain c where c_def: "c = drop (n + m) y" + by blast + have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>m + n + l\<^esup>)" + by (metis add.assoc assms(3) partial_pullback_memE(1) y_def) + then have y_eq: "y = b@a@c" + using a_def b_def c_def + by (metis append_take_drop_id drop_drop) + have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (metis a_def add.commute cartesian_power_drop le_add1 take_closed take_drop y_closed) + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using add_leD2 b_def le_add1 take_closed y_closed + by (meson trans_le_add1) + have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + using c_def cartesian_power_drop y_closed + by (metis add.commute) + have ac_closed: "a@c \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + using a_closed c_closed cartesian_power_concat(1) by blast + then have C: " local.partial_image m f y = f b # a @ c" + using b_closed y_eq partial_image_eq[of b m "a@c" "n + l" y f] + by blast + then have as_eq: "as = (f b)#(a@c)" + using as_def + by force + have B: "(tw m n) y = a@b@c" using y_eq tw_def[of n m] tp_1_permutation_action + by (smt a_closed b_closed c_closed tw_def) + then have "x = a@(b@c)" + by (simp add: y_def) + then have "twisted_partial_image n m f x = a@((f b)# c)" + using a_closed b_closed c_closed twisted_partial_image_eq + by blast + then have D: "permute_list (twisting_permutation n) as = twisted_partial_image n m f x" + using as_eq twisting_permutation_action[of a n "f b" c ] + by (metis a_closed cartesian_power_car_memE) + have "permute_list (twisting_permutation n) as \ S" + proof- + have S: "length s > n" + using s_def assms cartesian_power_car_memE le_add1 le_neq_implies_less + le_trans less_add_same_cancel1 less_one not_add_less1 + by (metis (no_types, lifting) subset_iff) + have "permute_list (twisting_permutation n) as = permute_list (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)" + using fun_inv_def s_def by blast + then have "permute_list (twisting_permutation n) as = + permute_list ((twisting_permutation n) \ (fun_inv (twisting_permutation n))) s" + using fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv + length_permute_list twisting_permutation_permutes[of n "length s"] + permute_list_compose[of "fun_inv (twisting_permutation n)" s "twisting_permutation n"] + by (metis S permute_list_compose) + then have "permute_list (twisting_permutation n) as = + permute_list (id) s" + by (metis S \permute_list (twisting_permutation n) as = permute_list + (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)\ + fun_inv_permute(3) length_greater_0_conv length_permute_list permute_list_compose + twisting_permutation_permutes) + then have "permute_list (twisting_permutation n) as = s" + by simp + then show ?thesis + using s_def + by (simp add: \s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as\) + qed + then show "x \ twisted_partial_pullback n m l f S" + unfolding twisted_partial_pullback_def using D + by (smt \x = a @ b @ c\ a_closed append.assoc append_eq_conv_conj b_closed + c_closed cartesian_power_car_memE cartesian_power_concat(1) length_append + list.inject local.partial_image_def twisted_partial_image_def + twisted_partial_pullback_def twisted_partial_pullback_memI) + qed +qed + +lemma twisted_partial_pullback_is_semialgebraic: + assumes "is_semialg_function m f" + assumes "is_semialgebraic (n + 1 + l) S" + shows "is_semialgebraic (n + m + l)(twisted_partial_pullback n m l f S)" +proof- + have "(fun_inv (twisting_permutation n)) permutes {.. carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + shows "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (smt ab_semigroup_add_class.add_ac(1) add.commute append_take_drop_id + assms augment_def cartesian_power_car_memE cartesian_power_drop length_append) + using assms cartesian_power_car_memE'' unfolding augment_def + by (metis (no_types, opaque_lifting) append_take_drop_id cartesian_power_concat(2) nat_le_iff_add take_closed) + +lemma tuple_partial_image_factor: + assumes "is_function_tuple Q\<^sub>p m fs" + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + assumes "length fs = n" + assumes "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + shows "tuple_partial_image m (fs@[f]) x = twisted_partial_image n m f (tuple_partial_image m fs (augment m x))" +proof- + obtain a where a_def: "a = take m x" + by blast + obtain b where b_def: "b = drop m x" + by blast + have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using a_def assms(4) le_add1 take_closed + by (meson dual_order.trans) + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + using assms(4) b_def cartesian_power_drop + by (metis (no_types, lifting)) + have A: "(augment m x) = a @ (a @ b)" + using a_def augment_def b_def + by blast + have 0: "tuple_partial_image m fs (augment m x) = ((function_tuple_eval Q\<^sub>p m fs) a) @ a @ b" + using A a_closed b_closed tuple_partial_image_eq[of a m "a@b" "m + l" "augment m x" fs] + cartesian_power_concat(1) + by blast + have 1: "tuple_partial_image m (fs@[f]) x = ((function_tuple_eval Q\<^sub>p m fs a) @ [f a])@b" + using 0 tuple_partial_image_eq[of a m b l x "fs@[f]"] unfolding function_tuple_eval_def + by (metis (no_types, lifting) a_closed a_def append_take_drop_id b_closed b_def + list.simps(8) list.simps(9) map_append) + have 2: "tuple_partial_image m (fs@[f]) x = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" + using 1 + by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv2 same_append_eq) + have 3: "tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs a) @ b" + using a_def b_def 2 tuple_partial_image_eq[of a m b l x fs ] assms tuple_partial_image_def + by blast + have 4: "twisted_partial_image n m f (tuple_partial_image m fs (augment m x)) = + (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" + using twisted_partial_image_eq[of _ n _ m _ l] 0 assms(1) assms(3) b_closed + local.a_closed local.function_tuple_eval_closed by blast + show ?thesis using 2 4 + by presburger +qed + +definition diagonalize where +"diagonalize n m S = S \ cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" + +lemma diagaonlize_is_semiaglebraic: + assumes "is_semialgebraic (n + n + m) S" + shows "is_semialgebraic (n + n + m) (diagonalize n m S)" +proof(cases "m = 0") + case True + then have 0: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) = {[]}" + unfolding cartesian_power_def + by simp + have 1: "\ n \ carrier (Q\<^sub>p\<^bsup>n+n\<^esup>)" + using Qp.cring_axioms assms diagonalE(2) + by blast + then have "cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = \ n" + using 0 cartesian_product_empty_right[of "\ n" Q\<^sub>p "n + n" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] + by linarith + then have "diagonalize n m S = S \ (\ n)" + using diagonalize_def + by presburger + then show ?thesis + using intersection_is_semialg True assms diag_is_semialgebraic + by auto +next + case False + have "is_semialgebraic (n + n + m) (cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" + using carrier_is_semialgebraic[of m] + cartesian_product_is_semialgebraic[of "n + n" "\ n" m "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] + diag_is_semialgebraic[of n] False + by blast + then show ?thesis + using intersection_is_semialg assms(1) diagonalize_def + by presburger +qed + +lemma list_segment_take: + assumes "length a \n" + shows "list_segment 0 n a = take n a" +proof- + have 0: "length (list_segment 0 n a) = length (take n a)" + using assms unfolding list_segment_def + by (metis (no_types, lifting) Groups.add_ac(2) add_diff_cancel_left' + append_take_drop_id le_Suc_ex length_append length_drop length_map map_nth) + have "\i. i < n \ list_segment 0 n a !i = take n a ! i" + unfolding list_segment_def using assms + by (metis add.left_neutral diff_zero nth_map_upt nth_take) + then show ?thesis using 0 + by (metis assms diff_zero le0 list_segment_length nth_equalityI) +qed + +lemma augment_inverse_is_semialgebraic: + assumes "is_semialgebraic (n+n+l) S" + shows "is_semialgebraic (n+l) ((augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>))" +proof- + obtain Ps where Ps_def: "Ps = (var_list_segment 0 n)" + by blast + obtain Qs where Qs_def: "Qs = (var_list_segment n (n+l))" + by blast + obtain Fs where Fs_def: "Fs = Ps@Ps@Qs" + by blast + have 0: "is_poly_tuple (n+l) Ps" + by (simp add: Ps_def var_list_segment_is_poly_tuple) + have 1: "is_poly_tuple (n+l) Qs" + by (simp add: Qs_def var_list_segment_is_poly_tuple) + have 2: "is_poly_tuple (n+l) (Ps@Qs)" + using Qp_is_poly_tuple_append[of "n+l" Ps Qs] + by (metis (no_types, opaque_lifting) "0" "1" add.commute) + have "is_poly_tuple (n+l) Fs" + using 0 2 Qp_is_poly_tuple_append[of "n + l" Ps "Ps@Qs"] Fs_def assms + by blast + have 3: "\x. x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ augment n x = poly_map (n + l) Fs x" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + have 30: "poly_map (n+l) Ps x = take n x" + using Ps_def map_by_var_list_segment[of x "n + l" n 0] + list_segment_take[of n x] cartesian_power_car_memE[of x Q\<^sub>p "n+l"] + by (simp add: A) + have 31: "poly_map (n + l) Qs x = drop n x" + using Qs_def map_by_var_list_segment_to_length[of x "n + l" n] A le_add1 + by blast + have 32: "poly_map (n + l) (Ps@Qs) x = take n x @ drop n x" + using poly_map_append[of x "n+l" Ps Qs ] + by (metis "30" "31" A append_take_drop_id) + show "augment n x = poly_map (n + l) Fs x" + using 30 32 poly_map_append + by (metis A Fs_def poly_map_append augment_def) + qed + have 4: "(augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) = poly_tuple_pullback (n + l) S Fs" + proof + show "augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ poly_tuple_pullback (n + l) S Fs" + proof fix x assume A: "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + then have 40: "augment n x \ S" + by blast + have 41: "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" + using 40 assms unfolding augment_def + using is_semialgebraic_closed + by blast + have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + proof- + have "take n x @ x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" + using augment_def A + by (metis "41" append_take_drop_id) + then have 0: "drop n (take n x @ x) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + by (metis (no_types, lifting) add.assoc cartesian_power_drop) + have "drop n (take n x @ x) = x" + proof- + have "length x \ n" + using A + by (metis IntD2 cartesian_power_car_memE le_add1) + then have "length (take n x) = n" + by (metis add_right_cancel append_take_drop_id + le_add_diff_inverse length_append length_drop) + then show ?thesis + by (metis append_eq_conv_conj) + qed + then show ?thesis + using 0 + by presburger + qed + then show "x \ poly_tuple_pullback (n + l) S Fs" + using 41 3 unfolding poly_tuple_pullback_def + by (metis (no_types, opaque_lifting) "40" add.commute cartesian_power_car_memE evimageI evimage_def poly_map_apply) + qed + show "poly_tuple_pullback (n + l) S Fs \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + proof fix x assume A: "x \ poly_tuple_pullback (n + l) S Fs" + have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + using A unfolding poly_tuple_pullback_def by blast + then show "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + using 3 + by (metis (no_types, lifting) A poly_map_apply poly_tuple_pullback_def vimage_inter_cong) + qed + qed + then show ?thesis using assms pullback_is_semialg[of "n + l" Fs] + using poly_tuple_pullback_eq_poly_map_vimage + unfolding restrict_def evimage_def Fs_def + by (smt "4" Ex_list_of_length Fs_def Ps_def Qs_def \is_poly_tuple (n + l) Fs\ add.commute + add_diff_cancel_left' append_assoc diff_zero is_semialgebraic_closed le_add2 length_append + not_add_less1 not_gr_zero padic_fields.is_semialgebraicE padic_fields_axioms var_list_segment_length zero_le) +qed + +lemma tuple_partial_pullback_is_semialg_map_tuple_induct: + assumes "is_semialg_map_tuple m fs" + assumes "is_semialg_function m f" + assumes "length fs = n" + shows "is_semialg_map_tuple m (fs@[f])" +proof(rule is_semialg_map_tupleI) + have 0: "is_function_tuple Q\<^sub>p m fs" + using assms is_semialg_map_tuple_def + by blast + show "is_function_tuple Q\<^sub>p m (fs @ [f])" + proof(rule is_function_tupleI) + have A0: "set (fs @ [f]) = insert f (set fs)" + by simp + have A1: "set fs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using 0 is_function_tuple_def + by blast + then show "set (fs @ [f]) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using assms 0 + by (metis (no_types, lifting) A0 is_semialg_function_closed list.simps(15) set_ConsD subset_code(1)) + qed + show "\k S. S \ semialg_sets (length (fs @ [f]) + k) \ + is_semialgebraic (m + k) (tuple_partial_pullback m (fs @ [f]) k S)" + proof- fix l S + assume A: "S \ semialg_sets (length (fs @ [f]) + l)" + then have B: "S \ semialg_sets (n + l + 1)" + using assms + by (metis (no_types, lifting) add.commute add_Suc_right add_diff_cancel_left' + append_Nil2 diff_Suc_1 length_Suc_conv length_append) + show "is_semialgebraic (m + l) (tuple_partial_pullback m (fs @ [f]) l S)" + proof- + obtain S0 where S0_def: "S0 = tuple_partial_pullback m fs (l+1) S" + by blast + have 0: "is_semialgebraic (m + l + 1) S0" + using B assms is_semialg_map_tupleE[of m fs "l + 1" S] + by (metis S0_def add.assoc is_semialgebraicI) + obtain S1 where S1_def: "S1 = twisted_partial_pullback m m l f S0" + by blast + then have "is_semialgebraic (m + m + l) S1" + using S1_def assms(1) 0 twisted_partial_pullback_is_semialgebraic[of m f m l S0] + by (simp add: assms(2)) + then have L: "is_semialgebraic (m + m + l) (diagonalize m l S1)" + using assms diagaonlize_is_semiaglebraic + by blast + have 1: "(tuple_partial_pullback m (fs @ [f]) l S) + = (augment m -` (diagonalize m l S1)) \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + proof + show "tuple_partial_pullback m (fs @ [f]) l S \ + augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + + proof fix x assume P0: "x \ tuple_partial_pullback m (fs @ [f]) l S " + show "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + proof + show "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + using tuple_partial_pullback_closed P0 + by blast + show "x \ augment m -` diagonalize m l S1" + proof- + obtain a where a_def: "a = take m x" + by blast + then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ le_add1 take_closed + by blast + obtain b where b_def: "b = drop m x" + by blast + then have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ cartesian_power_drop + by blast + have x_eq: "x = a@b" + using a_def b_def + by (metis append_take_drop_id) + have X0: "a @ a @ b = augment m x" + by (metis a_def augment_def b_def) + have "a @ a @ b \ diagonalize m l S1" + proof- + have "length (a@a) = m + m" + using a_closed cartesian_power_car_memE length_append + by blast + then have "take (m + m) (a @ a @ b) = a@a" + by (metis append.assoc append_eq_conv_conj) + then have X00: "take (m + m) (a @ a @ b) \ \ m" + using diagonalI[of "a@a"] a_def a_closed + by (metis append_eq_conv_conj cartesian_power_car_memE) + then have X01: "a @ a @ b \ cartesian_product (\ m) (carrier (Q\<^sub>p\<^bsup>l\<^esup>))" + using a_closed b_closed cartesian_product_memI[of "\ m" Q\<^sub>p "m+m" "carrier (Q\<^sub>p\<^bsup>l\<^esup>)" l "a @ a @ b"] + unfolding diagonal_def + by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed cartesian_power_drop mem_Collect_eq subsetI) + have X02: "twisted_partial_image m m f (a @ a @ b) = a @ ((f a)# b)" + using twisted_partial_image_eq[of a m a m b l _ f] a_closed b_closed + by blast + have "a @ a @ b \ S1" + proof- + have "twisted_partial_image m m f (a @ a @ b) \ S0" + proof- + have X020:"tuple_partial_image m fs (a @ ((f a)# b)) + = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" + using tuple_partial_image_eq[of a m "(f a)# b" "l + 1" _ fs] + by (metis (no_types, lifting) a_closed append_Cons append_eq_conv_conj + cartesian_power_car_memE self_append_conv2 tuple_partial_image_def) + have X021: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b \ S" + proof- + have X0210: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b = + (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" + unfolding function_tuple_eval_def + by (metis (mono_tags, lifting) append.assoc list.simps(8) list.simps(9) map_append) + have X0211: "(function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b = + tuple_partial_image m (fs @ [f]) x" + using x_eq tuple_partial_image_eq[of a m b l x "fs@[f]"] + by (simp add: a_closed b_closed) + have "tuple_partial_image m (fs @ [f]) x \ S" + using P0 tuple_partial_pullback_memE(2) + by blast + then show ?thesis using X0211 X0210 by presburger + qed + have X022: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) + = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" + proof- + have X0220: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = + tuple_partial_image m fs (a @ ((f a)# b))" + using X02 by presburger + have X0221: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = + (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)# b)" + using tuple_partial_image_eq + by (metis X02 X020 append_Cons self_append_conv2) + then show ?thesis + unfolding function_tuple_eval_def + by (metis X02 X020 X0221 append_same_eq) + qed + have X023: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) \ S" + using X02 X020 X021 by presburger + have "twisted_partial_image m m f (a @ a @ b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" + proof- + have "a @ ((f a)# b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (metis a_closed add.commute b_closed cartesian_power_car_memE + length_Cons length_append plus_1_eq_Suc) + proof- + have "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using assms(2) is_semialg_function_closed by blast + then have "f a \ carrier Q\<^sub>p" + using a_closed assms + by blast + then show "set (a @ f a # b) \ carrier Q\<^sub>p" + using assms a_closed b_closed + by (meson cartesian_power_car_memE'' cartesian_power_concat(1) cartesian_power_cons) + qed + then show ?thesis + using X02 + by presburger + qed + then show ?thesis + using S0_def X023 tuple_partial_pullback_def[of m fs "l+1" S ] + by blast + qed + then show ?thesis using X02 S1_def twisted_partial_pullback_def + by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed + drop_drop local.partial_image_def twisted_partial_image_def + twisted_partial_pullback_memI) + qed + then show ?thesis using X01 diagonalize_def[of m l S1] + by blast + qed + then show ?thesis + by (metis X0 vimageI) + qed + qed + qed + show "augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) \ tuple_partial_pullback m (fs @ [f]) l S" + proof + fix x + assume A: "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + then have X0: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" + by blast + obtain a where a_def: "a = take m x" + by blast + then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using X0 le_add1 take_closed by blast + obtain b where b_def: "b = drop m x" + by blast + then have a_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + using X0 cartesian_power_drop + by blast + have X1: "augment m x = a@a@b" + using a_def augment_def b_def + by blast + have X2: "a@a@b \ diagonalize m l S1" + using A X1 + by (metis Int_iff vimage_eq) + have X3: "a@a@b \ S1" + using X2 diagonalize_def + by blast + have X4: "twisted_partial_image m m f (a@a@b) \ S0" + using X3 S1_def twisted_partial_pullback_memE(2) + by blast + have X5: "a@((f a)#b) \ S0" + using X4 twisted_partial_image_eq[of a m a m b l _ f] + by (metis X0 a_closed a_def le_add1 take_closed) + have X6: "tuple_partial_image m fs (a@((f a)#b)) \ S" + using S0_def X5 tuple_partial_pullback_memE(2) + by blast + have X7: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) \ S" + using X6 using tuple_partial_image_eq + by (metis X0 a_def append_eq_conv_conj cartesian_power_car_memE + le_add1 take_closed tuple_partial_image_def) + have X8: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) = + tuple_partial_image m (fs @ [f]) x" + proof- + have X80: "tuple_partial_image m (fs @ [f]) x = (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" + using tuple_partial_image_def a_def b_def + by blast + then show ?thesis unfolding function_tuple_eval_def + by (metis (no_types, lifting) append_Cons append_eq_append_conv2 list.simps(9) map_append self_append_conv2) + qed + show "x \ tuple_partial_pullback m (fs @ [f]) l S" + using X8 X7 tuple_partial_pullback_def + by (metis X0 \is_function_tuple Q\<^sub>p m (fs @ [f])\ + tuple_partial_image_def tuple_partial_pullback_memI) + qed + qed + then show ?thesis + using augment_inverse_is_semialgebraic + by (simp add: L) + qed + qed +qed + +lemma singleton_tuple_partial_pullback_is_semialg_map_tuple: + assumes "is_semialg_function_tuple m fs" + assumes "length fs = 1" + shows "is_semialg_map_tuple m fs" +proof(rule is_semialg_map_tupleI) + show "is_function_tuple Q\<^sub>p m fs" + by (simp add: assms(1) semialg_function_tuple_is_function_tuple) + show "\k S. S \ semialg_sets (length fs + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + proof- fix k S assume A: "S \ semialg_sets (length fs + k)" + show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + proof- + obtain f where f_def: "fs = [f]" + using assms + by (metis One_nat_def length_0_conv length_Suc_conv) + have 0: "is_semialg_function m f" + using f_def assms is_semialg_function_tupleE'[of m fs f] + by simp + have 1: "\x. tuple_partial_image m fs x = partial_image m f x" + unfolding function_tuple_eval_def tuple_partial_image_def partial_image_def + by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv_if + f_def list.simps(8) list.simps(9)) + have 2: "tuple_partial_pullback m fs k S = partial_pullback m f k S" + proof + show "tuple_partial_pullback m fs k S \ partial_pullback m f k S" + using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def + by (metis (no_types, lifting) set_eq_subset vimage_inter_cong) + show "partial_pullback m f k S \ tuple_partial_pullback m fs k S" + using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def + by blast + qed + then show ?thesis + by (metis "0" A assms(2) is_semialg_functionE is_semialgebraicI) + qed + qed +qed + +lemma empty_tuple_partial_pullback_is_semialg_map_tuple: + assumes "is_semialg_function_tuple m fs" + assumes "length fs = 0" + shows "is_semialg_map_tuple m fs" + apply(rule is_semialg_map_tupleI) + using assms(1) semialg_function_tuple_is_function_tuple apply blast +proof- + fix k S assume A: "S \ semialg_sets (length fs + k)" + then have 0: "is_semialgebraic k S" + by (metis add.left_neutral assms(2) is_semialgebraicI) + have 1: "tuple_partial_pullback m fs k S = cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" + proof + have 1: "\x. function_tuple_eval Q\<^sub>p m fs (take m x) = []" + using assms unfolding function_tuple_eval_def + by blast + show "tuple_partial_pullback m fs k S \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" + apply(rule subsetI) apply(rule cartesian_product_memI[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k]) + apply blast using 0 is_semialgebraic_closed apply blast + using 0 assms unfolding 1 tuple_partial_pullback_def tuple_partial_image_def + apply (meson IntD2 le_add1 take_closed) + by (metis append_Nil evimageD evimage_def) + have 2: "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" + using is_semialgebraic_closed[of k S] 0 assms cartesian_product_closed[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k] by blast + show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ tuple_partial_pullback m fs k S" + apply(rule subsetI) apply(rule tuple_partial_pullback_memI) + using 2 apply blast + using assms semialg_function_tuple_is_function_tuple apply blast + unfolding 1 + by (metis carrier_is_semialgebraic cartesian_product_memE(2) is_semialgebraic_closed self_append_conv2) + qed + show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" + unfolding 1 + using "0" car_times_semialg_is_semialg by blast +qed + +lemma tuple_partial_pullback_is_semialg_map_tuple: + assumes "is_semialg_function_tuple m fs" + shows "is_semialg_map_tuple m fs" +proof- + have "\n fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" + proof- fix n + show " \ fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" + apply(induction n) + using singleton_tuple_partial_pullback_is_semialg_map_tuple empty_tuple_partial_pullback_is_semialg_map_tuple apply blast + proof- + fix n fs + assume IH: "(\fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs)" + assume A: "is_semialg_function_tuple m fs \ length fs = Suc n" + then obtain gs f where gs_f_def: "fs = gs@[f]" + by (metis length_Suc_conv list.discI rev_exhaust) + have gs_length: "length gs = n" + using gs_f_def + by (metis A length_append_singleton nat.inject) + have 0: "set gs \ set fs" + by (simp add: gs_f_def subsetI) + have 1: "is_semialg_function_tuple m gs" + apply(rule is_semialg_function_tupleI) + using 0 A gs_f_def is_semialg_function_tupleE'[of m fs] + by blast + then have 2: "is_semialg_map_tuple m gs" + using IH gs_length + by blast + have 3: "is_semialg_function m f" + using gs_f_def A + by (metis gs_length is_semialg_function_tupleE lessI nth_append_length) + then show "is_semialg_map_tuple m fs" + using assms 2 gs_f_def tuple_partial_pullback_is_semialg_map_tuple_induct + by blast + qed + qed + then show ?thesis + using assms by blast +qed + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Semialgebraic Functions are Closed under Composition with Semialgebraic Tuples\ + (********************************************************************************************) + (********************************************************************************************) + +lemma function_tuple_comp_partial_pullback: + assumes "is_semialg_function_tuple m fs" + assumes "length fs = n" + assumes "is_semialg_function n f" + assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + shows "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = + tuple_partial_pullback m fs k (partial_pullback n f k S)" +proof- + have 0: "\x. partial_image m (function_tuple_comp Q\<^sub>p fs f) x = + partial_image n f (tuple_partial_image m fs x)" + unfolding partial_image_def function_tuple_comp_def tuple_partial_image_def + using comp_apply[of f "function_tuple_eval Q\<^sub>p 0 fs"] + unfolding function_tuple_eval_def + proof - + fix x :: "((nat \ int) \ (nat \ int)) set list" + assume a1: "\x. (f \ (\x. map (\f. f x) fs)) x = f (map (\f. f x) fs)" + have f2: "\f rs. drop n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = rs" + by (simp add: assms(2)) + have "\f rs. take n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = map f fs" + by (simp add: assms(2)) + then show "(f \ (\rs. map (\f. f rs) fs)) (take m x) # drop m x = + f (take n (map (\f. f (take m x)) fs @ drop m x)) # drop n (map (\f. f (take m x)) fs @ drop m x)" + using f2 a1 by presburger + qed + show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = + tuple_partial_pullback m fs k (partial_pullback n f k S)" + proof + show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S \ tuple_partial_pullback m fs k (partial_pullback n f k S)" + proof fix x assume A: "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" + then have 1: "partial_image m (function_tuple_comp Q\<^sub>p fs f) x \ S" + using partial_pullback_memE(2) by blast + have 2: " partial_image n f (tuple_partial_image m fs x) \ S" + using 0 1 + by presburger + have 3: "x \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" + using A assms + by (metis partial_pullback_memE(1)) + have 4: "tuple_partial_image m fs x \ partial_pullback n f k S" + apply(rule partial_pullback_memI) + apply (metis "0" "3" add_cancel_left_left assms(1) assms(2) cartesian_power_drop drop0 + list.inject local.partial_image_def not_gr_zero semialg_function_tuple_is_function_tuple + tuple_partial_image_closed) + by (metis "2" local.partial_image_def) + show " x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" + apply(rule tuple_partial_pullback_memI) + apply (simp add: "3") + using assms(1) semialg_function_tuple_is_function_tuple apply blast + by (metis "4" tuple_partial_image_def) + qed + show " tuple_partial_pullback m fs k (partial_pullback n f k S) \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" + proof fix x assume A: "x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" + show "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S " + proof- + have "partial_image n f (tuple_partial_image m fs x) \ S" + using A partial_pullback_memE(2) tuple_partial_pullback_memE(2) + by blast + show ?thesis + apply(rule partial_pullback_memI) + apply (meson A subset_eq tuple_partial_pullback_closed) + by (metis "0" \local.partial_image n f (tuple_partial_image m fs x) \ S\ + local.partial_image_def) + qed + qed + qed +qed + +lemma semialg_function_tuple_comp: + assumes "is_semialg_function_tuple m fs" + assumes "length fs = n" + assumes "is_semialg_function n f" + shows "is_semialg_function m (function_tuple_comp Q\<^sub>p fs f)" +proof(rule is_semialg_functionI) + show "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + using function_tuple_comp_closed[of f Q\<^sub>p n fs] assms(1) assms(2) + assms(3) is_semialg_function_closed semialg_function_tuple_is_function_tuple + by blast + show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" + proof- fix k S + assume A0: "S \ semialg_sets (1 + k)" + show "is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" + proof- + have 0: "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = + tuple_partial_pullback m fs k (partial_pullback n f k S)" + using function_tuple_comp_partial_pullback[of m fs n f S k] assms + \S \ semialg_sets (1 + k)\ is_semialgebraicI is_semialgebraic_closed + by blast + have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" + using assms A0 is_semialg_functionE is_semialgebraicI + by blast + have 2: "is_semialgebraic (m + k) (tuple_partial_pullback m fs k (partial_pullback n f k S))" + using 1 0 assms tuple_partial_pullback_is_semialg_map_tuple[of m fs] + is_semialg_map_tupleE[of m fs k "partial_pullback n f k S"] + by blast + then show ?thesis + using 0 + by simp + qed + qed +qed + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Algebraic Operations on Semialgebraic Functions\ + (********************************************************************************************) + (********************************************************************************************) + + +text\Defining the set of extensional semialgebraic functions\ + +definition Qp_add_fun where +"Qp_add_fun xs = xs!0 \\<^bsub>Q\<^sub>p\<^esub> xs!1" + +definition Qp_mult_fun where +"Qp_mult_fun xs = xs!0 \ xs!1" + +text\Inversion function on first coordinates of Qp tuples. Arbitrarily redefined at 0 to map to 0\ + +definition Qp_invert where +"Qp_invert xs = (if ((xs!0) = \) then \ else (inv (xs!0)))" + +text\Addition is semialgebraic\ + +lemma addition_is_semialg: +"is_semialg_function 2 Qp_add_fun" +proof- + have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = (Qp_ev (pvar Q\<^sub>p 0) x) \\<^bsub>Q\<^sub>p\<^esub> (Qp_ev (pvar Q\<^sub>p 1) x)" + by (metis A One_nat_def eval_at_point_add pvar_closed less_Suc_eq numeral_2_eq_2) + then show " Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" + by (metis A Qp_add_fun_def add_vars_def add_vars_rep one_less_numeral_iff + pos2 semiring_norm(76)) + qed + then have 1: "restrict Qp_add_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = + restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" + by (meson restrict_ext) + have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" + using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] + by (meson MP.add.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) + then show ?thesis + using 1 semialg_function_on_carrier[of 2 "Qp_add_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] + semialg_function_on_carrier + by presburger +qed + +text\Multiplication is semialgebraic:\ + +lemma multiplication_is_semialg: +"is_semialg_function 2 Qp_mult_fun" +proof- + have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" + have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = + (Qp_ev (pvar Q\<^sub>p 0) x) \ (Qp_ev (pvar Q\<^sub>p 1) x)" + by (metis A One_nat_def eval_at_point_mult pvar_closed less_Suc_eq numeral_2_eq_2) + then show " Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" + by (metis A Qp_mult_fun_def mult_vars_def mult_vars_rep + one_less_numeral_iff pos2 semiring_norm(76)) + qed + then have 1: "restrict Qp_mult_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = + restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" + by (meson restrict_ext) + have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" + using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] + by (meson MP.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) + thus ?thesis + using 1 semialg_function_on_carrier[of 2 "Qp_mult_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] + semialg_function_on_carrier + by presburger +qed + +text\Inversion is semialgebraic:\ + +lemma(in field) field_nat_pow_inv: + assumes "a \ carrier R" + assumes "a \ \" + shows "inv (a [^] (n::nat)) = (inv a) [^] (n :: nat)" + apply(induction n) + using inv_one local.nat_pow_0 apply presburger + using assms nat_pow_of_inv + by (metis Units_one_closed field_inv(2) field_inv(3) unit_factor) + +lemma Qp_invert_basic_semialg: + assumes "is_basic_semialg (1 + k) S" + shows "is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" +proof- + obtain P n where P_n_def: "(n::nat) \ 0 \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>]) \ S = basic_semialg_set (1+k) n P \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>])" + using assms is_basic_semialg_def + by meson + obtain d::nat where d_def: "d = deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P)" + by auto + obtain l where l_def: "l = ((- d) mod n)" + by blast + have 1: "(l + d) mod n = 0" + by (metis add_eq_0_iff equation_minus_iff l_def mod_0 mod_add_cong mod_minus_eq zmod_int) + then obtain m::int where m_def: " (l + d) = m*n " + using d_def l_def + by (metis mult_of_nat_commute zmod_eq_0D) + have 2: "m \0" + proof- + have 10: "n > 0" + using P_n_def + by blast + have 11: "l \ 0" + using l_def 10 Euclidean_Division.pos_mod_sign of_nat_0_less_iff + by blast + then show ?thesis + using m_def + by (metis "10" le_add_same_cancel1 minus_add_cancel negative_zle + neq0_conv of_nat_le_0_iff zero_le_imp_eq_int zero_le_mult_iff) + qed + obtain N where N_def: "N = m*n" + by blast + have 3: "N \ d" + proof- + have "l \ 0" + using l_def d_def m_def Euclidean_Division.pos_mod_sign[of n "-d"] P_n_def + by linarith + then show ?thesis + using d_def N_def m_def + by linarith + qed + have 4: "deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P) \ nat N" + using d_def N_def 3 + by linarith + have 5: " P \ carrier (coord_ring Q\<^sub>p (Suc k))" + by (metis P_n_def plus_1_eq_Suc) + have 6: " \q\carrier (coord_ring Q\<^sub>p (Suc k)). + \x\carrier Q\<^sub>p - {\}. \a\carrier (Q\<^sub>p\<^bsup>k\<^esup>). Qp_ev q (insert_at_index a x 0) = (x[^]nat N) \ Qp_ev P (insert_at_index a (inv x) 0)" + using 3 4 d_def to_univ_poly_one_over_poly''[of 0 k P "nat N"] "5" Qp.field_axioms + by blast + obtain q where q_def: "q \ carrier (coord_ring Q\<^sub>p (Suc k)) \ ( \ x \ carrier Q\<^sub>p - {\}. ( \ a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). + eval_at_point Q\<^sub>p (insert_at_index a x 0) q = (x[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index a (inv x) 0) P)))" + using 6 + by blast + obtain T where T_def: "T = basic_semialg_set (1+k) n q" + by auto + have "is_basic_semialg (1 + k) T" + proof- + have "q \ carrier ( Q\<^sub>p[\\<^bsub>Suc k\<^esub>])" + using q_def + by presburger + then show ?thesis + using T_def is_basic_semialg_def + by (metis P_n_def plus_1_eq_Suc) + qed + then have T_semialg: "is_semialgebraic (1+k) T" + using T_def basic_semialg_is_semialg[of "1+k" T] is_semialgebraicI + by blast + obtain Nz where Nz_def: "Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 \ \}" + by blast + have Nz_semialg: "is_semialgebraic (1+k) Nz" + proof- + obtain Nzc where Nzc_def: "Nzc = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" + by blast + have 0: "Nzc = zero_set Q\<^sub>p (Suc k) (pvar Q\<^sub>p 0)" + unfolding zero_set_def + using Nzc_def + by (metis (no_types, lifting) Collect_cong eval_pvar zero_less_Suc) + have 1: "is_algebraic Q\<^sub>p (1+k) Nzc" + using 0 pvar_closed[of ] + by (metis is_algebraicI' plus_1_eq_Suc zero_less_Suc) + then have 2: "is_semialgebraic (1+k) Nzc" + using is_algebraic_imp_is_semialg by blast + have 3: "Nz = carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>) - Nzc" + using Nz_def Nzc_def + by blast + then show ?thesis + using 2 + by (simp add: complement_is_semialg) + qed + have 7: "(partial_pullback 1 Qp_invert k S) \ Nz = T \ Nz" + proof + show "partial_pullback 1 Qp_invert k S \ Nz \ T \ Nz" + proof fix c assume A: "c \ partial_pullback 1 Qp_invert k S \ Nz" + show "c \ T \ Nz" + proof- + have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + using A partial_pullback_closed[of 1 Qp_invert k S] + by blast + obtain x a where xa_def: "c = (x#a)" + using c_closed + by (metis Suc_eq_plus1 add.commute cartesian_power_car_memE length_Suc_conv) + have x_closed: "x \ carrier Q\<^sub>p" + using xa_def c_closed + by (metis (no_types, lifting) append_Cons cartesian_power_decomp + list.inject Qp.to_R1_to_R Qp.to_R_pow_closed) + have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + using xa_def c_closed + by (metis One_nat_def cartesian_power_drop drop0 drop_Suc_Cons) + have 0: "c \ Nz" + using A by blast + then have "x \ \" + using Nz_def xa_def + by (metis (mono_tags, lifting) mem_Collect_eq nth_Cons_0) + have 1: "Qp_invert [x] = inv x" + unfolding Qp_invert_def + by (metis \x \ \\ nth_Cons_0) + have 2: "partial_image 1 Qp_invert c \ S" + using A partial_pullback_memE[of c 1 "Qp_invert" k S] + by blast + have 3: "inv x # a \ S" + proof- + have 30: "[x] = take 1 c" + by (simp add: xa_def) + have 31: "a = drop 1 c" + by (simp add: xa_def) + show ?thesis + using 1 30 31 partial_image_def[of 1 "Qp_invert" c] xa_def "2" + by metis + qed + obtain y where y_def: "y \ carrier Q\<^sub>p \ eval_at_point Q\<^sub>p (inv x # a) P = y [^] n" + using 3 P_n_def basic_semialg_set_memE(2) + by blast + then have 4: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P + = x [^] (nat N) \ y [^] n" + by presburger + have 5: "x [^] (nat N) \ y [^] n = ((x[^]m)\y)[^]n" + proof- + have 50: "x [^] (N) \ y [^] n = x [^] (m*n) \ y [^] n" + using N_def by blast + have 51: "x [^] (m*n) = (x[^]m)[^]n" + using Qp_int_nat_pow_pow \x \ \\ not_nonzero_Qp x_closed + by metis + have 52: "x [^] (m*n)\ y [^] n = ((x[^]m) \ y) [^] n" + proof- + have 0: "x [^] (m*n)\ y [^] n= (x[^]m)[^]n \ (y[^] n)" + using "51" by presburger + have 1: "(x[^]m)[^]n \ (y[^] n) = ((x[^]m) \ y) [^] n" + apply(induction n) + using Qp.nat_pow_0 Qp.one_closed Qp.r_one apply presburger + using x_closed y_def + by (metis Qp.nat_pow_distrib Qp.nonzero_closed Qp_int_pow_nonzero \x \ \\ not_nonzero_Qp) + then show ?thesis + using "0" by blast + qed + have 53: "x [^] N = x [^] (nat N)" + proof- + have "N \ 0" + by (metis (full_types) Euclidean_Division.pos_mod_sign N_def P_n_def + add_increasing2 int.lless_eq l_def m_def of_nat_0_le_iff of_nat_le_0_iff ) + then show ?thesis + by (metis pow_nat) + qed + then show ?thesis + using 50 52 + by presburger + qed + have 6: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P = ((x[^]m)\y)[^]n" + using "4" "5" + by blast + have 7: "eval_at_point Q\<^sub>p c q = ((x[^]m)\y)[^]n" + proof- + have 70: "(insert_at_index a (inv x) 0) = inv x # a" + using insert_at_index.simps + by (metis (no_types, lifting) append_eq_append_conv2 append_same_eq append_take_drop_id drop0 same_append_eq) + have 71: "(insert_at_index a x) 0 = x # a" + by simp + then show ?thesis using 6 q_def + by (metis "70" DiffI \x \ \\ a_closed empty_iff insert_iff x_closed xa_def) + qed + have 8: "(x[^]m)\y \ carrier Q\<^sub>p" + proof- + have 80: "x[^]m \ carrier Q\<^sub>p" + using \x \ \\ x_closed Qp_int_pow_nonzero[of x m] unfolding nonzero_def + by blast + then show ?thesis + using y_def by blast + qed + then have "c \ T" + using T_def basic_semialg_set_def "7" c_closed by blast + then show ?thesis + by (simp add: \c \ T\ "0") + qed + qed + show "T \ Nz \ partial_pullback 1 Qp_invert k S \ Nz" + proof fix x assume A: "x \ T \ Nz" + show " x \ partial_pullback 1 Qp_invert k S \ Nz " + proof- + have " x \ partial_pullback 1 Qp_invert k S" + proof(rule partial_pullback_memI) + show x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + using T_def A + by (meson IntD1 basic_semialg_set_memE(1)) + show "Qp_invert (take 1 x) # drop 1 x \ S" + proof- + have 00: "x!0 \ \" + using A Nz_def + by blast + then have 0: "Qp_invert (take 1 x) # drop 1 x = inv (x!0) # drop 1 x" + unfolding Qp_invert_def + by (smt One_nat_def lessI nth_take) + have "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + using \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ cartesian_power_drop by blast + obtain a where a_def: "a = (x!0)" + by blast + have a_closed: "a \ carrier Q\<^sub>p" + using 00 a_def A Nz_def cartesian_power_car_memE'[of x Q\<^sub>p "Suc k" 0] inv_in_frac(1) + by blast + have a_nz: "a \ \" + using a_def Nz_def A + by blast + obtain b where b_def: "b = drop 1 x" + by blast + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + using b_def A Nz_def \drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)\ + by blast + have abx: "x = a#b" + using a_def b_def x_closed + by (metis (no_types, lifting) One_nat_def append_Cons append_Nil + append_eq_conv_conj cartesian_power_car_memE cartesian_power_decomp + lessI nth_take Qp.to_R1_to_R) + have 1: "Qp_invert (take 1 x) # drop 1 x = (inv a)#b" + using "0" a_def b_def + by blast + have 22: "eval_at_point Q\<^sub>p (insert_at_index b a 0) q = + (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" + using q_def a_closed a_nz b_closed + by blast + obtain c where c_def: "c \ carrier Q\<^sub>p \ Qp_ev q x = (c[^]n)" + using A T_def unfolding basic_semialg_set_def + by blast + obtain c' where c'_def: "c' = (inv a)[^]m \ c" + by blast + have c'_closed: "c' \ carrier Q\<^sub>p" + using c_def a_def a_closed a_nz Qp_int_pow_nonzero nonzero_def + c'_def inv_in_frac(3) Qp.m_closed Qp.nonzero_closed by presburger + have 3: "(eval_at_point Q\<^sub>p ((inv a) # b) P) = (c'[^]n)" + proof- + have 30: "x = insert_at_index b a 0" + using abx + by simp + have 31: "(c[^]n) = + (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" + using 22 30 c_def + by blast + have 32: "insert_at_index b (inv a) 0 = (inv a) # b" + using insert_at_index.simps + by (metis drop0 self_append_conv2 take0) + have 33: "(c[^]n) = + (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" + using "31" "32" by presburger + have 34: "(inv a) # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + apply(rule cartesian_power_car_memI'') + apply (metis b_closed cartesian_power_car_memE length_Suc_conv plus_1_eq_Suc) + using a_closed a_nz b_closed + apply (metis One_nat_def inv_in_frac(1) take0 take_Suc_Cons Qp.to_R1_closed) + by (metis abx b_closed b_def drop_Cons' not_Cons_self2) + have 35: "(eval_at_point Q\<^sub>p ((inv a) # b) P) \ carrier Q\<^sub>p" + using 34 P_n_def eval_at_point_closed + by blast + have "inv(a[^] (nat N)) \ (c[^]n) = + inv(a[^] (nat N)) \ ((a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P))" + using 31 "33" by presburger + then have 6: "inv(a[^] (nat N)) \ (c[^]n) = + inv(a[^] (nat N)) \ (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" + using 35 monoid.m_assoc[of Q\<^sub>p] Qp.monoid_axioms Qp.nat_pow_closed + Qp.nonzero_pow_nonzero a_nz inv_in_frac(1) local.a_closed by presburger + have 37:"inv(a[^] (nat N)) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" + proof- + have "inv(a[^] (nat N)) \ (a[^] (nat N)) = \ " + using a_closed a_nz Qp.nat_pow_closed Qp.nonzero_pow_nonzero field_inv(1) + by blast + then have "inv(a[^] (nat N)) \ (c[^]n) = + \ \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" + using 6 by presburger + then show ?thesis using 35 Qp.l_one by blast + qed + have 38:"(inv a)[^] (nat N) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" + using 37 group.nat_pow_inv[of Q\<^sub>p a "nat N"] a_closed Qp.field_axioms field.field_nat_pow_inv[of Q\<^sub>p] + by (metis a_nz) + have 39:"((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" + using 2 38 monoid.nat_pow_pow[of Q\<^sub>p "inv a" ] N_def + by (smt "3" Qp_int_nat_pow_pow a_closed a_nz inv_in_frac(3) of_nat_0_le_iff pow_nat) + have 310:"((((inv a)[^]m) \ c)[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" + proof- + have AA: "(inv a)[^]m \ carrier Q\<^sub>p" + using Qp_int_pow_nonzero nonzero_def a_closed a_nz inv_in_frac(3) Qp.nonzero_closed + by presburger + have "((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = ((((inv a)[^]m) \ c)[^]n)" + using Qp.nat_pow_distrib[of "(inv a)[^]m" c n] a_closed a_def c_def AA by blast + then show ?thesis + using "39" by blast + qed + then show ?thesis using c'_def + by blast + qed + have 4: "inv a # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + by (metis a_closed a_nz add.commute b_closed cartesian_power_cons inv_in_frac(1)) + then have 5: "((inv a) # b) \ S" + using 3 P_n_def c'_closed basic_semialg_set_memI[of "(inv a) # b" "1 + k" c' P n] + by blast + have 6: "Qp_invert (take 1 x) # drop 1 x = inv a # b" + using a_def b_def unfolding Qp_invert_def using "1" Qp_invert_def + by blast + show ?thesis using 5 6 + by presburger + qed + qed + then show ?thesis + using A by blast + qed + qed + qed + have 8: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) \ Nz)" + using "7" Nz_semialg T_semialg intersection_is_semialg + by auto + have 9: "(partial_pullback 1 Qp_invert k S) - Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S" + proof + show "partial_pullback 1 Qp_invert k S - Nz \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" + proof fix x assume A: " x \ partial_pullback 1 Qp_invert k S - Nz" + have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" + using A + by (metis DiffD1 partial_pullback_memE(1) plus_1_eq_Suc) + have 1: "take 1 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + by (metis "0" le_add1 plus_1_eq_Suc take_closed) + have 2: "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + using "0" cartesian_power_drop plus_1_eq_Suc + by presburger + have 3: " x = take 1 x @ drop 1 x " + using 0 + by (metis append_take_drop_id) + have 4: "Qp_invert (take 1 x) # drop 1 x \ S" + using A partial_pullback_memE'[of "take 1 x" 1 "drop 1 x" k x Qp_invert S] 1 2 3 + by blast + have 5: "x!0 = \" + using A 0 Nz_def by blast + have 6: "Qp_invert (take 1 x) # drop 1 x = x" + proof- + have "(take 1 x) =[x!0]" + using 0 + by (metis "1" "3" append_Cons nth_Cons_0 Qp.to_R1_to_R) + then have "Qp_invert (take 1 x) = \" + unfolding Qp_invert_def using 5 + by (metis less_one nth_take) + then show ?thesis using 0 5 + by (metis "3" Cons_eq_append_conv \take 1 x = [x ! 0]\ self_append_conv2) + qed + have "x \ S" + using 6 4 + by presburger + then show "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" + using Nz_def A 0 + by blast + qed + show "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S \ partial_pullback 1 Qp_invert k S - Nz" + proof fix x assume A: "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" + have A0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" + using A by blast + have A1: "x!0 = \" + using A by blast + have A2: "x \ S" + using A by blast + show " x \ partial_pullback 1 Qp_invert k S - Nz" + proof + show "x \ Nz" + using Nz_def A1 by blast + show " x \ partial_pullback 1 Qp_invert k S" + proof(rule partial_pullback_memI) + show "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" + using A0 + by (simp add: A0) + show "Qp_invert (take 1 x) # drop 1 x \ S" + proof- + have "Qp_invert (take 1 x) = \" + unfolding Qp_invert_def using A0 A1 + by (metis less_numeral_extra(1) nth_take) + then have "Qp_invert (take 1 x) # drop 1 x = x" + using A0 A1 A2 + by (metis (no_types, lifting) Cons_eq_append_conv Qp_invert_def \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ + append_take_drop_id inv_in_frac(2) le_add_same_cancel1 self_append_conv2 + take_closed Qp.to_R1_to_R Qp.to_R_pow_closed zero_le) + then show ?thesis + using A2 by presburger + qed + qed + qed + qed + qed + have 10: "is_semialgebraic (1+k) {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" + proof- + have "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} = V\<^bsub>Q\<^sub>p\<^esub> (Suc k) (pvar Q\<^sub>p 0)" + unfolding zero_set_def using eval_pvar[of 0 "Suc k"] Qp.cring_axioms + by blast + then show ?thesis using + is_zero_set_imp_basic_semialg pvar_closed[of 0 "Suc k"] Qp.cring_axioms + is_zero_set_imp_semialg plus_1_eq_Suc zero_less_Suc + by presburger + qed + have 11: "is_semialgebraic (1+k) ({xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S)" + using 10 assms basic_semialg_is_semialgebraic intersection_is_semialg + by blast + have 12: "(partial_pullback 1 Qp_invert k S) = ((partial_pullback 1 Qp_invert k S) \ Nz) \ + ((partial_pullback 1 Qp_invert k S) - Nz)" + by blast + have 13: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) - Nz)" + using 11 9 by metis + show ?thesis + using 8 12 13 + by (metis "7" Int_Diff_Un Int_commute plus_1_eq_Suc union_is_semialgebraic) +qed + +lemma Qp_invert_is_semialg: +"is_semialg_function 1 Qp_invert" +proof(rule is_semialg_functionI') + show 0: "Qp_invert \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ carrier Q\<^sub>p" + proof fix x + assume A: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + then obtain a where a_def: "x = [a]" + by (metis Qp.to_R1_to_R) + have a_closed: "a \ carrier Q\<^sub>p" + using a_def A cartesian_power_concat(1) last_closed' + by blast + show " Qp_invert x \ carrier Q\<^sub>p" + apply(cases "a = \") + unfolding Qp_invert_def using a_def a_closed + apply (meson Qp.to_R_to_R1) + by (metis a_closed a_def inv_in_frac(1) Qp.to_R_to_R1) + qed + show "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" + using Qp_invert_basic_semialg + by blast +qed + +lemma Taylor_deg_1_expansion'': + assumes "f \ carrier Q\<^sub>p_x" + assumes "\n. f n \ \\<^sub>p" + assumes "a \ \\<^sub>p " + assumes "b \ \\<^sub>p" + shows "\c c' c''. c = to_fun f a \ c' = deriv f a \ c \ \\<^sub>p \ c' \ \\<^sub>p \c'' \ \\<^sub>p \ + to_fun f (b) = c \ c' \ (b \ a) \ (c'' \ (b \ a)[^](2::nat))" +proof- + obtain S where S_def: "S = (Q\<^sub>p \ carrier := \\<^sub>p \)" + by blast + have 1: "f \ carrier (UP S)" + unfolding S_def using val_ring_subring UPQ.poly_cfs_subring[of \\<^sub>p f] assms + by blast + have 2: " f \ carrier (UP (Q\<^sub>p\carrier := \\<^sub>p\))" + using val_ring_subring 1 assms poly_cfs_subring[of \\<^sub>p] + by blast + have 3: "\c\\\<^sub>p. f \ b = f \ a \ UPQ.deriv f a \ (b \ a) \ c \ (b \ a) [^] (2::nat)" + using UP_subring_taylor_appr'[of \\<^sub>p f b a] UP_subring_taylor_appr[of \\<^sub>p f b a] val_ring_subring 1 2 assms + by blast + then show ?thesis + using UP_subring_taylor_appr[of \\<^sub>p f b a] assms UP_subring_deriv_closed[of \\<^sub>p f a] + UP_subring_eval_closed[of \\<^sub>p f a] 2 val_ring_subring by blast +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Sets Defined by Residues of Valuation Ring Elements\ +(**************************************************************************************************) +(**************************************************************************************************) + +sublocale padic_fields < Res: cring "Zp_res_ring (Suc n)" + using p_residues residues.cring + by blast + +context padic_fields +begin + +definition Qp_res where +"Qp_res x n = to_Zp x n " + +lemma Qp_res_closed: + assumes "x \ \\<^sub>p" + shows "Qp_res x n \ carrier (Zp_res_ring n)" + unfolding Qp_res_def using assms val_ring_memE residue_closed to_Zp_closed by blast + +lemma Qp_res_add: + assumes "x \ \\<^sub>p" + assumes "y \ \\<^sub>p" + shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" + unfolding Qp_res_def + using assms residue_of_sum to_Zp_add by presburger + +lemma Qp_res_mult: + assumes "x \ \\<^sub>p" + assumes "y \ \\<^sub>p" + shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" + unfolding Qp_res_def + using assms residue_of_prod to_Zp_mult by presburger + +lemma Qp_res_diff: + assumes "x \ \\<^sub>p" + assumes "y \ \\<^sub>p" + shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" + unfolding Qp_res_def + using assms residue_of_diff to_Zp_minus + by (meson val_ring_res) + +lemma Qp_res_zero: + shows "Qp_res \ n = 0" + unfolding Qp_res_def to_Zp_zero + using residue_of_zero(2) by blast + +lemma Qp_res_one: + assumes "n > 0" + shows "Qp_res \ n = (1::int)" + using assms + unfolding Qp_res_def to_Zp_one + using residue_of_one(2) by blast + +lemma Qp_res_nat_inc: + shows "Qp_res ([(n::nat)]\\) n = n mod p^n" + unfolding Qp_res_def unfolding to_Zp_nat_inc + using Zp_nat_inc_res by blast + +lemma Qp_res_int_inc: + shows "Qp_res ([(k::int)]\\) n = k mod p^n" + unfolding Qp_res_def unfolding to_Zp_int_inc + using Zp_int_inc_res by blast + +lemma Qp_poly_res_monom: +assumes "a \ \\<^sub>p" +assumes "x \ \\<^sub>p" +assumes "Qp_res a n = 0" +assumes "k > 0" +shows "Qp_res (up_ring.monom (UP Q\<^sub>p) a k \ x) n = 0" +proof- + have 0: "up_ring.monom (UP Q\<^sub>p) a k \ x = a \ x [^] k" + apply(rule UPQ.to_fun_monom[of a x k]) + using assms val_ring_memE apply blast + using assms val_ring_memE by blast + have 1: "x[^]k \ \\<^sub>p" + using assms val_ring_nat_pow_closed by blast + show ?thesis unfolding 0 + using Qp_res_mult[of a "x[^]k" n] assms + using "1" residue_times_zero_r by presburger +qed + +lemma Qp_poly_res_zero: + assumes "q \ carrier (UP Q\<^sub>p)" + assumes "\i. q i \ \\<^sub>p" + assumes "\i. Qp_res (q i) n = 0" + assumes "x \ \\<^sub>p" + shows "Qp_res (q \ x) n = 0" +proof- + have "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0" + proof(rule UPQ.poly_induct[of q], rule assms, rule ) + fix p assume A: "p \ carrier (UP Q\<^sub>p)" " deg Q\<^sub>p p = 0" " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" + have 0: "p \ x = p 0" + using assms + by (metis A(1) A(2) val_ring_memE UPQ.ltrm_deg_0 UPQ.to_fun_ctrm) + show "Qp_res (p \ x) n = 0" + unfolding 0 using A by blast + next + fix p + assume A0: "(\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0)" + "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p" + show "(\i. p i \ \\<^sub>p \ Qp_res (p i) n = 0) \ Qp_res (p \ x) n = 0" + proof assume A1: " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" + obtain k where k_def: "k = deg Q\<^sub>p p" + by blast + obtain q where q_def: "q = UPQ.trunc p" + by blast + have q_closed: "q \ carrier (UP Q\<^sub>p)" + unfolding q_def + using A0(2) UPQ.trunc_closed by blast + have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p p" + unfolding q_def + using A0(2) A0(3) UPQ.trunc_degree by blast + have 9: "\i. i < deg Q\<^sub>p p \ q i = p i" + unfolding q_def + using A0(2) UPQ.trunc_cfs by blast + have 90: "\i. \ i < deg Q\<^sub>p p \ q i = \" + unfolding q_def + proof - + fix i :: nat + assume "\ i < deg Q\<^sub>p p" + then have "deg Q\<^sub>p q < i" + using q_deg by linarith + then show "Cring_Poly.truncate Q\<^sub>p p i = \" + using UPQ.deg_gtE q_closed q_def by blast + qed + have 10: "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0)" + proof fix i + show "q i \ \\<^sub>p \ Qp_res (q i) n = 0" + apply(cases "i < deg Q\<^sub>p p") + using A1 9[of i] apply presburger + unfolding q_def using Qp_res_zero 90 + by (metis q_def zero_in_val_ring) + qed + have 11: "Qp_res (q \ x) n = 0" + using 10 A1 A0 q_closed q_deg by blast + have 12: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k" + unfolding k_def q_def + using A0(2) UPQ.trunc_simps(1) by blast + have 13: "p \ x = q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x" + proof- + have 0: " (q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k) \ x = q \ x \ up_ring.monom (UP Q\<^sub>p) (p k) k \ x" + apply(rule UPQ.to_fun_plus) + using A0(2) UPQ.ltrm_closed k_def apply blast + unfolding q_def apply(rule UPQ.trunc_closed, rule A0) + using assms val_ring_memE by blast + show ?thesis + using 0 12 by metis + qed + have 14: "(up_ring.monom (UP Q\<^sub>p) (p k) k) \ x \ \\<^sub>p" + apply(rule val_ring_poly_eval) + using A0(2) UPQ.ltrm_closed k_def apply blast + using UPQ.cfs_monom[of "p k" k ] A1 zero_in_val_ring + using A0(2) UPQ.ltrm_cfs k_def apply presburger + using assms(4) by blast + have 15: "Qp_res ((up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" + apply(rule Qp_poly_res_monom) + using A1 apply blast using assms apply blast + using A1 apply blast unfolding k_def using A0 by blast + have 16: "Qp_res (q \ x) n = 0" + using A0 10 11 by blast + have 17: "q \ x \ \\<^sub>p" + apply(rule val_ring_poly_eval, rule q_closed) + using 10 apply blast by(rule assms) + have 18: "Qp_res (q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" + using Qp_res_add[of "q \ x" "up_ring.monom (UP Q\<^sub>p) (p k) k \ x" n] 14 17 + unfolding 15 16 + by (metis "10" Qp_res_add UPQ.cfs_add UPQ.coeff_of_sum_diff_degree0 q_closed q_deg) + show "Qp_res (p \ x) n = 0" + using 13 18 by metis + qed + qed + thus ?thesis using assms by blast +qed + +lemma Qp_poly_res_eval_0: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "x \ \\<^sub>p" + assumes "\i. f i \ \\<^sub>p" + assumes "\i. g i \ \\<^sub>p" + assumes "\i. Qp_res (f i) n = Qp_res (g i) n" + shows "Qp_res (f \ x) n = Qp_res (g \ x) n" +proof- + obtain F where F_def: "F = f \\<^bsub>UP Q\<^sub>p\<^esub>g" + by blast + have F_closed: "F \ carrier (UP Q\<^sub>p)" + unfolding F_def + using assms by blast + have F_cfs: "\i. F i = (f i) \ (g i)" + unfolding F_def + using assms UPQ.cfs_minus by blast + have F_cfs_res: "\i. Qp_res (F i) n = Qp_res (f i) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g i) n" + unfolding F_cfs apply(rule Qp_res_diff) + using assms apply blast using assms by blast + have 0: "\i. Qp_res (f i) n = Qp_res (g i) n" + using assms by blast + have F_cfs_res': "\i. Qp_res (F i) n = 0" + unfolding F_cfs_res 0 + by (metis diff_self mod_0 residue_minus) + have 1: "\i. F i \ \\<^sub>p" + unfolding F_cfs using assms + using val_ring_minus_closed by blast + have 2: "Qp_res (F \ x) n = 0" + by(rule Qp_poly_res_zero, rule F_closed, rule 1, rule F_cfs_res', rule assms) + have 3: "F \ x = f \ x \ g \ x" + unfolding F_def using assms + by (meson assms UPQ.to_fun_diff val_ring_memE) + have 4: "Qp_res (F \ x) n = Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n" + unfolding 3 apply(rule Qp_res_diff, rule val_ring_poly_eval, rule assms) + using assms apply blast using assms apply blast + apply(rule val_ring_poly_eval, rule assms) + using assms apply blast by(rule assms) + have 5: "f \ x \ \\<^sub>p" + apply(rule val_ring_poly_eval, rule assms) + using assms apply blast using assms by blast + have 6: "g \ x \ \\<^sub>p" + apply(rule val_ring_poly_eval, rule assms) + using assms apply blast by(rule assms) + show "Qp_res (f \ x) n = Qp_res (g \ x) n" + using 5 6 2 Qp_res_closed[of "f \ x" n] Qp_res_closed[of "g \ x" n] + unfolding 4 + proof - + assume "Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n = 0" + then show ?thesis + by (metis (no_types) Qp_res_def 5 6 res_diff_zero_fact(1) residue_of_diff to_Zp_closed val_ring_memE) + qed +qed + +lemma Qp_poly_res_eval_1: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "x \ \\<^sub>p" + assumes "y \ \\<^sub>p" + assumes "\i. f i \ \\<^sub>p" + assumes "Qp_res x n = Qp_res y n" + shows "Qp_res (f \ x) n = Qp_res (f \ y) n" +proof- + have "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" + apply(rule UPQ.poly_induct[of f], rule assms) + proof fix f assume A: "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f = 0" "\i. f i \ \\<^sub>p" + show "Qp_res (f \ x) n = Qp_res (f \ y) n" + proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ f = to_polynomial Q\<^sub>p a" + using assms + by (metis A(1) A(2) UPQ.lcf_closed UPQ.to_poly_inverse) + have a_eq: "f = to_polynomial Q\<^sub>p a" + using a_def by blast + have 0: "f \ x = a" + using a_def assms unfolding a_eq + by (meson UPQ.to_fun_to_poly val_ring_memE) + have 1: "f \ y = a" + using a_def assms unfolding a_eq + by (meson UPQ.to_fun_to_poly val_ring_memE) + show " Qp_res (f \ x) n = Qp_res (f \ y) n" + unfolding 0 1 by blast + qed + next + fix f + assume A: " (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p f \ (\i. q i \ \\<^sub>p) \ Qp_res (q \ x) n = Qp_res (q \ y) n)" + "f \ carrier (UP Q\<^sub>p)" " 0 < deg Q\<^sub>p f" + show "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" + proof assume A1: "\i. f i \ \\<^sub>p" + obtain q where q_def: "q = UPQ.trunc f" + by blast + have q_closed: "q \ carrier (UP Q\<^sub>p)" + using q_def A UPQ.trunc_closed by presburger + have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p f" + using q_def A UPQ.trunc_degree by blast + have q_cfs: "\i. q i \ \\<^sub>p" + proof fix i show "q i \ \\<^sub>p" + apply(cases "i < deg Q\<^sub>p f") + unfolding q_def using A A1 UPQ.trunc_cfs + apply presburger + using q_deg q_closed + proof - + assume "\ i < deg Q\<^sub>p f" + then have "deg Q\<^sub>p f \ i" + by (meson diff_is_0_eq neq0_conv zero_less_diff) + then show "Cring_Poly.truncate Q\<^sub>p f i \ \\<^sub>p" + by (metis (no_types) UPQ.deg_eqI diff_is_0_eq' le_trans nat_le_linear neq0_conv q_closed q_def q_deg zero_in_val_ring zero_less_diff) + qed + qed + hence 0: "Qp_res (q \ x) n = Qp_res (q \ y) n" + using A q_closed q_deg by blast + have 1: "Qp_res (UPQ.ltrm f \ x) n = Qp_res (UPQ.ltrm f \ y) n" + proof- + have 10: "UPQ.ltrm f \ x = (f (deg Q\<^sub>p f)) \ x[^](deg Q\<^sub>p f)" + using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger + have 11: "UPQ.ltrm f \ y = (f (deg Q\<^sub>p f)) \ y[^](deg Q\<^sub>p f)" + using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger + obtain d where d_def: "d = deg Q\<^sub>p f" + by blast + have 12: "Qp_res (x[^]d) n = Qp_res (y[^]d) n" + apply(induction d) + using Qp.nat_pow_0 apply presburger + using Qp_res_mult assms Qp.nat_pow_Suc val_ring_nat_pow_closed by presburger + hence 13: "Qp_res (x [^] deg Q\<^sub>p f) n = Qp_res (y [^] deg Q\<^sub>p f) n" + unfolding d_def by blast + have 14: "x [^] deg Q\<^sub>p f \ \\<^sub>p" + using assms val_ring_nat_pow_closed by blast + have 15: "y [^] deg Q\<^sub>p f \ \\<^sub>p" + using assms val_ring_nat_pow_closed by blast + have 16: "Qp_res (f (deg Q\<^sub>p f) \ x [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (x [^] deg Q\<^sub>p f) n" + apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " x[^](deg Q\<^sub>p f)" n]) + using A1 apply blast by(rule 14) + have 17: "Qp_res (f (deg Q\<^sub>p f) \ y [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (y [^] deg Q\<^sub>p f) n" + apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " y[^](deg Q\<^sub>p f)" n]) + using A1 apply blast by(rule 15) + show ?thesis + unfolding 10 11 16 17 13 by blast + qed + have f_decomp: "f = q \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.ltrm f" + using A unfolding q_def + using UPQ.trunc_simps(1) by blast + have 2: "f \ x = q \ x \ (UPQ.ltrm f \ x)" + using A f_decomp q_closed q_cfs + by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(2)) + have 3: "f \ y = q \ y \ (UPQ.ltrm f \ y)" + using A f_decomp q_closed q_cfs + by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(3)) + show 4: " Qp_res (f \ x) n = Qp_res (f \ y) n " + unfolding 2 3 using assms q_cfs Qp_res_add 0 1 + by (metis (no_types, opaque_lifting) "2" "3" A(2) A1 Qp_res_def poly_eval_cong) + qed + qed + thus ?thesis using assms by blast +qed + +lemma Qp_poly_res_eval_2: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "x \ \\<^sub>p" + assumes "y \ \\<^sub>p" + assumes "\i. f i \ \\<^sub>p" + assumes "\i. g i \ \\<^sub>p" + assumes "\i. Qp_res (f i) n = Qp_res (g i) n" + assumes "Qp_res x n = Qp_res y n" + shows "Qp_res (f \ x) n = Qp_res (g \ y) n" +proof- + have 0: "Qp_res (f \ x) n = Qp_res (g \ x) n" + using Qp_poly_res_eval_0 assms by blast + have 1: "Qp_res (g \ x) n = Qp_res (g \ y) n" + using Qp_poly_res_eval_1 assms by blast + show ?thesis unfolding 0 1 by blast +qed + +definition poly_res_class where +"poly_res_class n d f = {q \ carrier (UP Q\<^sub>p). deg Q\<^sub>p q \ d \ (\i. q i \ \\<^sub>p \ Qp_res (f i) n = Qp_res (q i) n) }" + +lemma poly_res_class_closed: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "g \ carrier (UP Q\<^sub>p)" + assumes "deg Q\<^sub>p f \ d" + assumes "deg Q\<^sub>p g \ d" + assumes "g \ poly_res_class n d f" + shows "poly_res_class n d f = poly_res_class n d g" + unfolding poly_res_class_def + apply(rule equalityI) +apply(rule subsetI) + unfolding mem_Collect_eq apply(rule conjI, blast, rule conjI, blast) + using assms unfolding poly_res_class_def mem_Collect_eq + apply presburger +apply(rule subsetI) unfolding mem_Collect_eq + apply(rule conjI, blast, rule conjI, blast) + using assms unfolding poly_res_class_def mem_Collect_eq + by presburger + +lemma poly_res_class_memE: + assumes "f \ poly_res_class n d g" + shows "f \ carrier (UP Q\<^sub>p)" + "deg Q\<^sub>p f \ d" + "f i \ \\<^sub>p" + "Qp_res (g i) n = Qp_res (f i) n" + using assms unfolding poly_res_class_def mem_Collect_eq apply blast + using assms unfolding poly_res_class_def mem_Collect_eq apply blast + using assms unfolding poly_res_class_def mem_Collect_eq apply blast + using assms unfolding poly_res_class_def mem_Collect_eq by blast + +definition val_ring_polys where +"val_ring_polys = {f \ carrier (UP Q\<^sub>p). (\i. f i \ \\<^sub>p)} " + +lemma val_ring_polys_closed: +"val_ring_polys \ carrier (UP Q\<^sub>p)" + unfolding val_ring_polys_def by blast + +lemma val_ring_polys_memI: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "\i. f i \ \\<^sub>p" + shows "f \ val_ring_polys" + using assms unfolding val_ring_polys_def by blast + +lemma val_ring_polys_memE: + assumes "f \ val_ring_polys" + shows "f \ carrier (UP Q\<^sub>p)" + "f i \ \\<^sub>p" + using assms unfolding val_ring_polys_def apply blast + using assms unfolding val_ring_polys_def by blast + +definition val_ring_polys_grad where +"val_ring_polys_grad d = {f \ val_ring_polys. deg Q\<^sub>p f \ d}" + +lemma val_ring_polys_grad_closed: +"val_ring_polys_grad d \ val_ring_polys" + unfolding val_ring_polys_grad_def by blast + +lemma val_ring_polys_grad_closed': +"val_ring_polys_grad d \ carrier (UP Q\<^sub>p)" + unfolding val_ring_polys_grad_def val_ring_polys_def by blast + +lemma val_ring_polys_grad_memI: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "\i. f i \ \\<^sub>p" + assumes "deg Q\<^sub>p f \ d" + shows "f \ val_ring_polys_grad d" + using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast + +lemma val_ring_polys_grad_memE: + assumes "f \ val_ring_polys_grad d" + shows "f \ carrier (UP Q\<^sub>p)" + "deg Q\<^sub>p f \ d" + "f i \ \\<^sub>p" + using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast + using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast + using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast + +lemma poly_res_classes_in_val_ring_polys_grad: + assumes "f \ val_ring_polys_grad d" + shows "poly_res_class n d f \ val_ring_polys_grad d" + apply(rule subsetI, rule val_ring_polys_grad_memI) + apply(rule poly_res_class_memE[of _ n d f], blast) + apply(rule poly_res_class_memE[of _ n d f], blast) + by(rule poly_res_class_memE[of _ n d f], blast) + +lemma poly_res_class_disjoint: + assumes "f \ val_ring_polys_grad d" + assumes "f \ poly_res_class n d g" + shows "poly_res_class n d f \ poly_res_class n d g = {}" + apply(rule equalityI) + apply(rule subsetI) + using assms + unfolding poly_res_class_def mem_Collect_eq Int_iff + apply (metis val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3)) + by blast + +lemma poly_res_class_refl: + assumes "f \ val_ring_polys_grad d" + shows "f \ poly_res_class n d f" + unfolding poly_res_class_def mem_Collect_eq + using assms val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3) by blast + +lemma poly_res_class_memI: + assumes "f \ carrier (UP Q\<^sub>p)" + assumes "deg Q\<^sub>p f \ d" + assumes "\i. f i \ \\<^sub>p" + assumes "\i. Qp_res (f i) n = Qp_res (g i) n" + shows "f \ poly_res_class n d g" + unfolding poly_res_class_def mem_Collect_eq using assms + by metis + +definition poly_res_classes where +"poly_res_classes n d = poly_res_class n d ` val_ring_polys_grad d" + +lemma poly_res_classes_disjoint: + assumes "A \ poly_res_classes n d" + assumes "B \ poly_res_classes n d" + assumes "g \ A - B" + shows "A \ B = {}" +proof- + obtain a where a_def: "a \ val_ring_polys_grad d \ A = poly_res_class n d a" + using assms unfolding poly_res_classes_def by blast + obtain b where b_def: "b \ val_ring_polys_grad d \ B = poly_res_class n d b" + using assms unfolding poly_res_classes_def by blast + have 0: "\f. f \ A \ B \ False" + proof- + fix f assume A: "f \ A \ B" + have 1: "\i. Qp_res (g i) n \ Qp_res (f i) n" + proof(rule ccontr) + assume B: "\i. Qp_res (g i) n \ Qp_res (f i) n" + then have 2: "\i. Qp_res (g i) n = Qp_res (f i) n" + by blast + have 3: "g \ poly_res_class n d a" + using a_def assms by blast + have 4: "\i. Qp_res (b i) n = Qp_res (f i) n" + apply(rule poly_res_class_memE[of _ n d]) + using assms A b_def by blast + have 5: "\i. Qp_res (a i) n = Qp_res (g i) n" + apply(rule poly_res_class_memE[of _ n d]) + using assms A a_def by blast + have 6: "g \ poly_res_class n d b" + apply(rule poly_res_class_memI, rule poly_res_class_memE[of _ n d a], rule 3, + rule poly_res_class_memE[of _ n d a], rule 3, rule poly_res_class_memE[of _ n d a], rule 3) + unfolding 2 4 by blast + show False using 6 b_def assms by blast + qed + then obtain i where i_def: "Qp_res (g i) n \ Qp_res (f i) n" + by blast + have 2: "\i. Qp_res (a i) n = Qp_res (f i) n" + apply(rule poly_res_class_memE[of _ n d]) + using A a_def by blast + have 3: "\i. Qp_res (b i) n = Qp_res (f i) n" + apply(rule poly_res_class_memE[of _ n d]) + using A b_def by blast + have 4: "\i. Qp_res (a i) n = Qp_res (g i) n" + apply(rule poly_res_class_memE[of _ n d]) + using assms a_def by blast + show False using i_def 2 unfolding 4 2 by blast + qed + show ?thesis using 0 by blast +qed + +definition int_fun_to_poly where +"int_fun_to_poly (f::nat \ int) i = [(f i)]\\" + +lemma int_fun_to_poly_closed: + assumes "\i. i > d \ f i = 0" + shows "int_fun_to_poly f \ carrier (UP Q\<^sub>p)" + apply(rule UPQ.UP_car_memI[of d]) + using assms unfolding int_fun_to_poly_def + using Qp.int_inc_zero apply presburger + by(rule Qp.int_inc_closed) + +lemma int_fun_to_poly_deg: + assumes "\i. i > d \ f i = 0" + shows "deg Q\<^sub>p (int_fun_to_poly f) \ d" + apply(rule UPQ.deg_leqI, rule int_fun_to_poly_closed, rule assms, blast) + unfolding int_fun_to_poly_def using assms + using Qp.int_inc_zero by presburger + +lemma Qp_res_mod_triv: + assumes "a \ \\<^sub>p" + shows "Qp_res a n mod p ^ n = Qp_res a n" + using assms Qp_res_closed[of a n] + by (meson mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) + +lemma int_fun_to_poly_is_class_wit: + assumes "f \ poly_res_class n d g" + shows "(int_fun_to_poly (\i::nat. Qp_res (f i) n)) \ poly_res_class n d g" +proof(rule poly_res_class_memI[of ], rule int_fun_to_poly_closed[of d]) + show 0: "\i. d < i \ Qp_res (f i) n = 0" + proof- fix i assume A: "d < i" + hence 0: "deg Q\<^sub>p f < i" + using A assms poly_res_class_memE(2)[of f n d g] + by linarith + have 1: "f i = \" + using 0 assms poly_res_class_memE[of f n d g] + using UPQ.UP_car_memE(2) by blast + show "Qp_res (f i) n = 0" + unfolding 1 Qp_res_zero by blast + qed + show "deg Q\<^sub>p (int_fun_to_poly (\i. Qp_res (f i) n)) \ d" + by(rule int_fun_to_poly_deg, rule 0, blast) + show "\i. int_fun_to_poly (\i. Qp_res (f i) n) i \ \\<^sub>p" + unfolding int_fun_to_poly_def + using Qp.int_mult_closed Qp_val_ringI val_of_int_inc by blast + show "\i. Qp_res (int_fun_to_poly (\i. Qp_res (f i) n) i) n = Qp_res (g i) n" + unfolding int_fun_to_poly_def Qp_res_int_inc + using Qp_res_mod_triv assms poly_res_class_memE(4) Qp_res_closed UPQ.cfs_closed + by (metis poly_res_class_memE(3)) +qed + +lemma finite_support_funs_finite: +"finite (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" +proof- + have 0: "finite (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n))" + apply(rule finite_PiE, blast) + using residue_ring_card[of n] by blast + obtain g where g_def: "g = (\f. (\i::nat. if i \ {..d} then f i else (0::int)))" + by blast + have 1: "g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) = (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" + proof(rule equalityI, rule subsetI) + fix x assume A: "x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" + obtain f where f_def: "f \ (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) \ x = g f" + using A by blast + have x_eq: "x = g f" + using f_def by blast + show "x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" + proof(rule, rule) + fix i assume A: "i \ {..d}" + show "x i \ carrier (Zp_res_ring n)" + proof(cases "i \ {..d}") + case True + then have T0: "f i \ carrier (Zp_res_ring n)" + using f_def by blast + have "x i = f i" + unfolding x_eq g_def + using True by metis + thus ?thesis using T0 by metis + next + case False + then have F0: "x i = 0" + unfolding x_eq g_def by metis + show ?thesis + unfolding F0 + by (metis residue_mult_closed residue_times_zero_r) + qed + next + show "x \ {f. \i>d. f i = 0}" + proof(rule, rule, rule) + fix i assume A: "d < i" + then have 0: "i \ {..d}" + by simp + thus "x i = 0" + unfolding x_eq g_def + by metis + qed + qed + next + show "({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} + \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" + proof(rule subsetI) + fix x + assume A: " x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" + show " x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" + proof- + obtain h where h_def: "h = restrict x {..d}" + by blast + have 0: "\i. i \ {..d} \ h i = x i" + unfolding h_def restrict_apply by metis + have 1: "\i. i \ {..d} \ h i = undefined" + unfolding h_def restrict_apply by metis + have 2: "\i. i \ {..d} \ h i \ carrier (Zp_res_ring n)" + using A 0 unfolding 0 by blast + have 3: "h \ {..d} \\<^sub>E carrier (residue_ring (p ^ n))" + by(rule, rule 2, blast, rule 1, blast) + have 4: "\i. i \ {..d} \ x i = 0" + using A unfolding Int_iff mem_Collect_eq + by (metis atMost_iff eq_imp_le le_simps(1) linorder_neqE_nat) + have 5: "x = g h" + proof fix i + show "x i = g h i" + unfolding g_def + apply(cases "i \ {..d}") + using 0 apply metis unfolding 4 + by metis + qed + show ?thesis unfolding 5 using 3 by blast + qed + qed + qed + have 2: "finite (g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)))" + using 0 by blast + show ?thesis using 2 unfolding 1 by blast +qed + +lemma poly_res_classes_finite: +"finite (poly_res_classes n d)" +proof- + have 0: "poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}) = poly_res_classes n d" + proof(rule equalityI, rule subsetI) + fix x assume A: " x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" + then obtain f where f_def: "f \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} \ + x = poly_res_class n d (int_fun_to_poly f)" + by blast + have x_eq: "x = poly_res_class n d (int_fun_to_poly f)" + using f_def by blast + show "x \ poly_res_classes n d" + proof- + have 0: "int_fun_to_poly f \ val_ring_polys_grad d" + apply(rule val_ring_polys_grad_memI, rule int_fun_to_poly_closed[of d]) + using f_def apply blast + using int_fun_to_poly_def + apply (metis Qp.int_inc_closed padic_fields.int_fun_to_poly_def padic_fields.val_of_int_inc padic_fields_axioms val_ring_memI) + apply(rule int_fun_to_poly_deg) + using f_def by blast + show ?thesis unfolding poly_res_classes_def x_eq + using 0 by blast + qed + next + show "poly_res_classes n d + \ poly_res_class n d ` + int_fun_to_poly ` + (({..d} \ carrier (residue_ring (p ^ n))) \ + {f. \i>d. f i = 0})" + proof(rule subsetI) + fix x assume A: " x \ poly_res_classes n d" + show "x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" + proof- + obtain f where f_def: "f \ val_ring_polys_grad d \ x = poly_res_class n d f" + using A unfolding poly_res_classes_def by blast + have x_eq: "x = poly_res_class n d f" + using f_def by blast + obtain h where h_def: "h = (\i::nat. Qp_res (f i) n)" + by blast + have 0: "\i. i > d \ f i = \" + proof- fix i assume A: "i > d" + have "i > deg Q\<^sub>p f" + apply(rule le_less_trans[of _ d]) + using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq + apply blast + by(rule A) + then show "f i = \" + using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq + using UPQ.deg_leE by blast + qed + have 1: "\i. i > d \ h i = 0" + unfolding h_def 0 Qp_res_zero by blast + have 2: "x = poly_res_class n d (int_fun_to_poly h)" + unfolding x_eq + apply(rule poly_res_class_closed) + using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast + apply(rule int_fun_to_poly_closed[of d], rule 1, blast) + using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast + apply(rule int_fun_to_poly_deg, rule 1, blast) + unfolding h_def + apply(rule int_fun_to_poly_is_class_wit, rule poly_res_class_refl) + using f_def by blast + have 3: "h \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" + apply(rule , rule ) + unfolding h_def apply(rule Qp_res_closed, rule val_ring_polys_grad_memE[of _ d]) + using f_def apply blast + unfolding mem_Collect_eq apply(rule, rule) + unfolding 0 Qp_res_zero by blast + show ?thesis + unfolding 2 using 3 by blast + qed + qed + qed + have 1: "finite (poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}))" + using finite_support_funs_finite by blast + show ?thesis using 1 unfolding 0 by blast +qed + +lemma Qp_res_eq_zeroI: + assumes "a \ \\<^sub>p" + assumes "val a \ n" + shows "Qp_res a n = 0" +proof- + have 0: "val_Zp (to_Zp a) \ n" + using assms to_Zp_val by presburger + have 1: "to_Zp a n = 0" + apply(rule zero_below_val_Zp, rule to_Zp_closed) + using val_ring_closed assms apply blast + by(rule 0) + thus ?thesis unfolding Qp_res_def by blast +qed + +lemma Qp_res_eqI: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "Qp_res (a \ b) n = 0" + shows "Qp_res a n = Qp_res b n" + using assms by (metis Qp_res_def val_ring_memE res_diff_zero_fact(1) to_Zp_closed to_Zp_minus) + +lemma Qp_res_eqI': + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "val (a \ b) \ n" + shows "Qp_res a n = Qp_res b n" + apply(rule Qp_res_eqI, rule assms, rule assms, rule Qp_res_eq_zeroI) + using assms Q\<^sub>p_def Zp_def \_def padic_fields.val_ring_minus_closed padic_fields_axioms apply blast + by(rule assms) + +lemma Qp_res_eqE: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "Qp_res a n = Qp_res b n" + shows "val (a \ b) \ n" +proof- + have 0: "val (a \ b) = val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b)" + using assms + by (metis to_Zp_minus to_Zp_val val_ring_minus_closed) + have 1: "(to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) n = 0" + using assms unfolding Qp_res_def + by (meson val_ring_memE res_diff_zero_fact'' to_Zp_closed) + have 2: "val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) \ n" + apply(cases "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>") + proof - + assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>" + have "\n. eint (int n) \ val_Zp \\<^bsub>Z\<^sub>p\<^esub>" + by (metis (no_types) Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) + then show ?thesis + using a1 by presburger + next + assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ \\<^bsub>Z\<^sub>p\<^esub>" + have 00: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ carrier Z\<^sub>p" + using assms + by (meson val_ring_memE Zp.cring_simprules(4) to_Zp_closed) + show ?thesis + using 1 a1 ord_Zp_geq[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b" n] 00 + val_ord_Zp[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b"] eint_ord_code by metis + qed + thus ?thesis unfolding 0 by blast +qed + +lemma notin_closed: +"(\ ((c::eint) \ x \ x \ d)) = (x < c \ d < x)" + by auto + +lemma Qp_res_neqI: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "val (a \ b) < n" + shows "Qp_res a n \ Qp_res b n" + apply(rule ccontr) + using Qp_res_eqE[of a b n] assms + using notin_closed by blast + +lemma Qp_res_equal: + assumes "a \ \\<^sub>p" + assumes "l = Qp_res a n" + shows "Qp_res a n = Qp_res ([l]\\) n " + unfolding Qp_res_int_inc assms using assms Qp_res_mod_triv by presburger + +definition Qp_res_class where +"Qp_res_class n b = {a \ \\<^sub>p. Qp_res a n = Qp_res b n}" + +definition Qp_res_classes where +"Qp_res_classes n = Qp_res_class n ` \\<^sub>p" + +lemma val_ring_int_inc_closed: +"[(k::int)]\\ \ \\<^sub>p" +proof- + have 0: "[(k::int)]\\ = \ ([(k::int)]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using inc_of_int by blast + thus ?thesis + by blast +qed + +lemma val_ring_nat_inc_closed: +"[(k::nat)]\\ \ \\<^sub>p" +proof- + have 0: "[k]\\ = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using inc_of_nat by blast + thus ?thesis + by blast +qed + +lemma Qp_res_classes_wits: +"Qp_res_classes n = (\l::int. Qp_res_class n ([l]\\)) ` (carrier (Zp_res_ring n))" +proof- + obtain F where F_def: "F = (\l::int. Qp_res_class n ([l]\\))" + by blast + have 0: "Qp_res_classes n = F ` (carrier (Zp_res_ring n))" + proof(rule equalityI, rule subsetI) + fix x assume A: "x \ Qp_res_classes n" + then obtain a where a_def: "a \ \\<^sub>p \ x = Qp_res_class n a" + unfolding Qp_res_classes_def by blast + have 1: "Qp_res a n = Qp_res ([(Qp_res a n)]\\) n " + apply(rule Qp_res_equal) + using a_def apply blast by blast + have 2: "Qp_res_class n a = Qp_res_class n ([(Qp_res a n)]\\)" + unfolding Qp_res_class_def using 1 by metis + have 3: "x = Qp_res_class n ([(Qp_res a n)]\\)" + using a_def unfolding 2 by blast + have 4: "a \ \\<^sub>p" + using a_def by blast + show " x \ F ` carrier (Zp_res_ring n)" + unfolding F_def 3 + using Qp_res_closed[of a n] 4 by blast + next + show "F ` carrier (residue_ring (p ^ n)) \ Qp_res_classes n" + proof(rule subsetI) + fix x assume A: "x \ F ` (carrier (Zp_res_ring n))" + then obtain l where l_def: "l \ carrier (Zp_res_ring n) \ x = F l" + using A by blast + have 0: "x = F l" + using l_def by blast + show "x \ Qp_res_classes n" + unfolding 0 F_def Qp_res_classes_def using val_ring_int_inc_closed by blast + qed + qed + then show ?thesis unfolding F_def by blast +qed + +lemma Qp_res_classes_finite: +"finite (Qp_res_classes n)" +by (metis Qp_res_classes_wits finite_atLeastLessThan_int finite_imageI p_res_ring_car) + +definition Qp_cong_set where +"Qp_cong_set \ a = {x \ \\<^sub>p. to_Zp x \ = a \}" + +lemma Qp_cong_set_as_ball: + assumes "a \ carrier Z\<^sub>p" + assumes "a \ = 0" + shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[\]" +proof- + have 0: "\ a \ carrier Q\<^sub>p" + using assms inc_closed[of a] by blast + show ?thesis + proof + show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\]" + proof fix x assume A: "x \ Qp_cong_set \ a" + show "x \ B\<^bsub>\ \<^esub>[\]" + proof(rule c_ballI) + show t0: "x \ carrier Q\<^sub>p" + using A unfolding Qp_cong_set_def + using val_ring_memE by blast + show "eint (int \) \ val (x \ \)" + proof- + have t1: "to_Zp x \ = 0" + using A unfolding Qp_cong_set_def + by (metis (mono_tags, lifting) assms(2) mem_Collect_eq) + have t2: "val_Zp (to_Zp x) \ \" + apply(cases "to_Zp x = \\<^bsub>Z\<^sub>p\<^esub>") + apply (metis Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) + using ord_Zp_geq[of "to_Zp x" \] A unfolding Qp_cong_set_def + by (metis (no_types, lifting) val_ring_memE eint_ord_simps(1) t1 to_Zp_closed to_Zp_def val_ord_Zp) + then show ?thesis using A unfolding Qp_cong_set_def mem_Collect_eq + using val_ring_memE + by (metis Qp_res_eqE Qp_res_eq_zeroI Qp_res_zero to_Zp_val zero_in_val_ring) + qed + qed + qed + show "B\<^bsub>int \\<^esub>[\] \ Qp_cong_set \ a" + proof fix x assume A: "x \ B\<^bsub>int \\<^esub>[\]" + then have 0: "val x \ \" + using assms c_ballE[of x \ \] + by (smt Qp.minus_closed Qp.r_right_minus_eq Qp_diff_diff) + have 1: "to_Zp x \ carrier Z\<^sub>p" + using A 0 assms c_ballE(1) to_Zp_closed by blast + have 2: "x \ \\<^sub>p" + using 0 A val_ringI c_ballE + by (smt Q\<^sub>p_def Zp_def \_def eint_ord_simps(1) of_nat_0 of_nat_le_0_iff val_ring_ord_criterion padic_fields_axioms val_ord' zero_in_val_ring) + then have "val_Zp (to_Zp x) \ \" + using 0 1 A assms c_ballE[of x \ \] to_Zp_val by presburger + then have "to_Zp x \ = 0" + using 1 zero_below_val_Zp by blast + then show " x \ Qp_cong_set \ a" + unfolding Qp_cong_set_def using assms(2) 2 + by (metis (mono_tags, lifting) mem_Collect_eq) + qed + qed +qed + +lemma Qp_cong_set_as_ball': + assumes "a \ carrier Z\<^sub>p" + assumes "val_Zp a < eint (int \)" + shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[(\ a)]" +proof + show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\ a]" + proof fix x + assume A: "x \ Qp_cong_set \ a" + then have 0: "to_Zp x \ = a \" + unfolding Qp_cong_set_def by blast + have 1: "x \ \\<^sub>p" + using A unfolding Qp_cong_set_def by blast + have 2: "to_Zp x \ carrier Z\<^sub>p" + using 1 val_ring_memE to_Zp_closed by blast + have 3: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" + using 0 assms 2 val_Zp_dist_def val_Zp_dist_res_eq2 by presburger + have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) > val_Zp a" + using 3 assms(2) less_le_trans[of "val_Zp a" "eint (int \)" "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a)" ] + by blast + then have 5: "val_Zp (to_Zp x) = val_Zp a" + using assms 2 equal_val_Zp by blast + have 7: "val (x \ (\ a)) \ \" + using 3 5 1 by (metis "2" Zp.minus_closed assms(1) inc_of_diff to_Zp_inc val_of_inc) + then show "x \ B\<^bsub>int \\<^esub>[\ a]" + using c_ballI[of x \ "\ a"] 1 assms val_ring_memE by blast + qed + show "B\<^bsub>int \\<^esub>[\ a] \ Qp_cong_set \ a" + proof fix x + assume A: "x \ B\<^bsub>int \\<^esub>[\ a]" + then have 0: "val (x \ \ a) \ \" + using c_ballE by blast + have 1: "val (\ a) = val_Zp a" + using assms Zp_def \_def padic_fields.val_of_inc padic_fields_axioms + by metis + then have 2: "val (x \ \ a) > val (\ a)" + using 0 assms less_le_trans[of "val (\ a)" "eint (int \)" "val (x \ \ a)"] + by metis + have "\ a \ carrier Q\<^sub>p" + using assms(1) inc_closed by blast + then have B: "val x = val (\ a)" + using 2 A assms c_ballE(1)[of x \ "\ a"] + by (metis ultrametric_equal_eq) + have 3: "val_Zp (to_Zp x) = val_Zp a" + by (metis "1" A \val x = val (\ a)\ assms(1) c_ballE(1) to_Zp_val val_pos val_ringI) + have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" + using 0 A 3 + by (metis B Zp.minus_closed assms(1) c_ballE(1) inc_of_diff to_Zp_closed to_Zp_inc val_of_inc val_pos val_ring_val_criterion) + then have 5: "to_Zp x \ = a \" + by (meson A Zp.minus_closed assms(1) c_ballE(1) res_diff_zero_fact(1) to_Zp_closed zero_below_val_Zp) + have 6: "x \ \\<^sub>p" + proof- + have "val x \ 0" + using B assms 1 val_pos by presburger + then show ?thesis + using A c_ballE(1) val_ringI by blast + qed + then show "x \ Qp_cong_set \ a" unfolding Qp_cong_set_def + using "5" by blast + qed +qed + +lemma Qp_cong_set_is_univ_semialgebraic: + assumes "a \ carrier Z\<^sub>p" + shows "is_univ_semialgebraic (Qp_cong_set \ a)" +proof(cases "a \ = 0") + case True + then show ?thesis + using ball_is_univ_semialgebraic[of \ \] Qp.zero_closed Qp_cong_set_as_ball assms + by metis +next + case False + then have "\ \ 0" + using assms residues_closed[of a 0] + by (meson p_res_ring_0') + then obtain n where n_def: "Suc n = \" + by (metis lessI less_Suc_eq_0_disj) + then have "val_Zp a < eint (int \)" + using below_val_Zp_zero[of a n] + by (smt False assms eint_ile eint_ord_simps(1) eint_ord_simps(2) zero_below_val_Zp) + then show ?thesis + using ball_is_univ_semialgebraic[of "\ a" \] Qp.zero_closed Qp_cong_set_as_ball'[of a \] assms + inc_closed by presburger +qed + +lemma constant_res_set_semialg: + assumes "l \ carrier (Zp_res_ring n)" + shows "is_univ_semialgebraic {x \ \\<^sub>p. Qp_res x n = l}" +proof- + have 0: "{x \ \\<^sub>p. Qp_res x n = l} = Qp_cong_set n ([l]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + apply(rule equalityI') + unfolding mem_Collect_eq Qp_cong_set_def Qp_res_def + apply (metis val_ring_memE Zp_int_inc_rep nat_le_linear p_residue_padic_int to_Zp_closed) + using assms + by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) + show ?thesis unfolding 0 + apply(rule Qp_cong_set_is_univ_semialgebraic) + by(rule Zp.int_inc_closed) +qed + +end + +end diff --git a/thys/Padic_Field/Padic_Field_Topology.thy b/thys/Padic_Field/Padic_Field_Topology.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Padic_Field_Topology.thy @@ -0,0 +1,1326 @@ +theory Padic_Field_Topology + imports Padic_Fields +begin + +(**************************************************************************************************) +(**************************************************************************************************) +section\Topology of $p$-adic Fields\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + In this section we develop some basic properties of the topology on the $p$-adics. Open and + closed sets are defined, convex subsets of the value group are characterized. +\ +type_synonym padic_univ_poly = "nat \ padic_number" + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\$p$-adic Balls\ +(**************************************************************************************************) +(**************************************************************************************************) + +context padic_fields +begin + +definition c_ball :: "int \ padic_number \ padic_number set" ("B\<^bsub>_\<^esub>[_]") where +"c_ball n c = {x \ carrier Q\<^sub>p. val (x \ c) \ n}" + +lemma c_ballI: + assumes "x \ carrier Q\<^sub>p" + assumes " val (x \ c) \ n" + shows "x \ c_ball n c" + using assms c_ball_def + by blast + +lemma c_ballE: + assumes "x \ c_ball n c" + shows "x \ carrier Q\<^sub>p" + " val (x \ c) \ n" + using assms c_ball_def apply blast + using assms c_ball_def by blast + +lemma c_ball_in_Qp: + "B\<^bsub>n\<^esub>[c] \ carrier Q\<^sub>p" + unfolding c_ball_def + by blast + +definition +q_ball :: "nat \ int \ int \ padic_number \ padic_number set" where +"q_ball n k m c = {x \ carrier Q\<^sub>p. (ac n (x \ c) = k \ (ord (x \ c)) = m) }" + +lemma q_ballI: + assumes "x \ carrier Q\<^sub>p" + assumes "ac n (x \ c) = k" + assumes "(ord (x \ c)) = m" + shows "x \ q_ball n k m c" + using assms q_ball_def + by blast + +lemma q_ballE: + assumes "x \ q_ball n k m c " + shows "x \ carrier Q\<^sub>p" + + using assms q_ball_def by blast + +lemma q_ballE': + assumes "x \ q_ball n k m c " + shows "ac n (x \ c) = k" + "(ord (x \ c)) = m" + using assms q_ball_def apply blast + using assms q_ball_def by blast + +lemma q_ball_in_Qp: + "q_ball n k m c \ carrier Q\<^sub>p" + unfolding q_ball_def by blast + +lemma ac_ord_prop: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "ord a = ord b" + assumes "ord a = n" + assumes "ac m a = ac m b" + assumes "m > 0" + shows "val (a \ b) \ m + n " +proof- + have 0: "a = (\[^]n) \ \ (angular_component a)" + using angular_component_factors_x assms(1) assms(4) by blast + have 1: "b = (\[^]n) \ \ (angular_component b)" + using angular_component_factors_x assms(4) assms(2) assms(3) + by presburger + have 2: "a \b = (\[^]n) \ \ (angular_component a) \ + (\[^]n) \ \ (angular_component b) " + using 0 1 by auto + have 3: "a \b = (\[^]n) \( \ (angular_component a) \ \ (angular_component b))" + proof- + have 30: "(\[^]n) \ carrier Q\<^sub>p" + by (simp add: p_intpow_closed(1)) + have 31: " \ (angular_component a) \ carrier Q\<^sub>p" + using Zp.nonzero_one_closed angular_component_closed assms(1) frac_closed local.inc_def + by presburger + have 32: " \ (angular_component b) \ carrier Q\<^sub>p" + using Zp.nonzero_one_closed angular_component_closed assms(2) frac_closed local.inc_def + by presburger + show ?thesis + using 2 30 31 32 ring.ring_simprules(23)[of Q\<^sub>p "(\[^]n)"] + unfolding a_minus_def + by (metis Qp.domain_axioms cring.cring_simprules(25) cring.cring_simprules(29) + cring.cring_simprules(3) domain.axioms(1)) + qed + have 4: "a \b = (\[^]n) \( \ ((angular_component a) \\<^bsub>Z\<^sub>p\<^esub> (angular_component b)))" + using 3 + by (simp add: angular_component_closed assms(1) assms(2) inc_of_diff) + have 5: "val_Zp ((angular_component a) \\<^bsub>Z\<^sub>p\<^esub> (angular_component b)) \ m " + proof- + have "((angular_component a) \\<^bsub>Z\<^sub>p\<^esub> (angular_component b)) m = 0" + using assms(5) + unfolding ac_def + using Q\<^sub>p_def Qp.nonzero_memE(2) angular_component_closed assms(1) assms(2) residue_of_diff' + by auto + then show ?thesis + using Zp.minus_closed angular_component_closed assms(1) assms(2) ord_Zp_geq val_Zp_def val_ord_Zp + by auto + qed + have 6: "val (a \ b) \ n + val ( \ (angular_component a) \ \ (angular_component b))" + using 3 Qp.minus_closed angular_component_closed assms(1) assms(2) inc_closed + ord_p_pow_int p_intpow_closed(1) p_intpow_closed(2) val_mult val_ord + by simp + have 7: "n + val ( \ (angular_component a) \ \ (angular_component b)) + = n + val_Zp ((angular_component a) \\<^bsub>Z\<^sub>p\<^esub> (angular_component b))" + using Zp.minus_closed angular_component_closed assms(1) assms(2) inc_of_diff val_of_inc + by simp + have 8: "n + val_Zp ( (angular_component a) \\<^bsub>Z\<^sub>p\<^esub>(angular_component b)) + \ n + m" + using 5 + by (metis add_mono_thms_linordered_semiring(2) plus_eint_simps(1)) + then have 9: "n + val ( \ (angular_component a) \ \ (angular_component b)) + \ n + m" + using "7" by presburger + then show ?thesis + by (metis (no_types, opaque_lifting) "6" add.commute order_trans plus_eint_simps(1)) +qed + +lemma c_ball_q_ball: + assumes "b \ nonzero Q\<^sub>p" + assumes "n > 0" + assumes "k = ac n b" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ q_ball n k m c" + shows "q_ball n k m c = c_ball (m + n) d" +proof + show "q_ball n k m c \ B\<^bsub>m + int n\<^esub>[d]" + proof + fix x + assume A0: "x \ q_ball n k m c" + show "x \ B\<^bsub>m + int n\<^esub>[d]" + proof- + have A1: "(ac n (x \ c) = k \ (ord (x \ c)) = m)" + using A0 q_ball_def + by blast + have "val (x \ d) \ m + n" + proof- + have A2: "(ac n (d \ c) = k \ (ord (d \ c)) = m)" + using assms(5) q_ball_def + by blast + have A3: "(x \ c) \ nonzero Q\<^sub>p" + proof- + have "k \0" + using A2 assms(1) assms(3) assms(5) ac_units[of b n] + by (metis One_nat_def Suc_le_eq assms(2) zero_not_in_residue_units) + then show ?thesis + by (smt A0 Qp.domain_axioms ac_def assms(4) cring.cring_simprules(4) domain.axioms(1) + mem_Collect_eq not_nonzero_Qp q_ball_def) + qed + have A4: "(d \ c) \ nonzero Q\<^sub>p" + proof- + have "k \0" + using A2 assms(1) assms(3) assms(5) ac_units[of b n] + by (metis One_nat_def Suc_le_eq assms(2) zero_not_in_residue_units) + then show ?thesis + by (metis (no_types, lifting) A2 Qp.domain_axioms ac_def assms(4) assms(5) + cring.cring_simprules(4) domain.axioms(1) mem_Collect_eq not_nonzero_Qp q_ball_def) + qed + then have " val ((x \ c) \(d \ c)) \ n + m" + using ac_ord_prop[of "(x \ c)" "(d \ c)" m n ] A1 A2 assms A3 + by simp + then show ?thesis + by (smt A0 Qp_diff_diff assms(4) assms(5) q_ballE) + qed + then show ?thesis + by (metis (no_types, lifting) A0 c_ball_def mem_Collect_eq q_ball_def) + qed + qed + show "B\<^bsub>m + int n\<^esub>[d] \ q_ball n k m c" + proof + fix x + assume A: "x \ B\<^bsub>m + int n\<^esub>[d]" + show "x \ q_ball n k m c" + proof- + have A0: "val (x \ d) \ m + n" + using A c_ball_def + by blast + have A1: "ord (d \ c) = m" + using assms(5) q_ball_def + by blast + have A2: "ac n (d \ c) = k" + using assms(5) q_ball_def + by blast + have A3: "(d \ c) \ \" + using A2 assms + by (metis ac_def ac_units le_eq_less_or_eq le_neq_implies_less less_one nat_le_linear + padic_integers.zero_not_in_residue_units padic_integers_def prime zero_less_iff_neq_zero) + have A4: "val (d \ c) =m" + by (simp add: A1 A3 val_def) + have A5: "val (x \ d) > val (d \ c)" + by (smt A0 A4 assms(2) eint_ord_code(4) eint_ord_simps(1) eint_ord_simps(2) of_nat_0_less_iff val_def) + have A6: "val ((x \ d) \ (d \ c)) = m" + using A4 A0 A5 + by (metis (mono_tags, opaque_lifting) A Qp.minus_closed assms(4) assms(5) + c_ballE(1) q_ballE val_ultrametric_noteq) + have A7: "val (x \ c) = m" + proof- + have "(x \ d) \ (d \ c) = ((x \ d) \ d) \ c" + by (metis (no_types, lifting) A Qp.domain_axioms a_minus_def assms(4) assms(5) + c_ball_def cring.cring_simprules(3) cring.cring_simprules(4) + cring.cring_simprules(7) domain.axioms(1) mem_Collect_eq q_ball_def) + have "(x \ d) \ (d \ c) = (x \ (\ d \ d)) \ c" + by (metis (no_types, opaque_lifting) A Qp.add.l_cancel_one Qp.add.m_comm Qp.l_neg + Qp.minus_closed Qp.plus_diff_simp Qp.zero_closed assms(4) assms(5) + c_ballE(1) q_ballE) + then show ?thesis + by (metis (no_types, lifting) A A6 Qp.domain_axioms assms(5) c_ball_def + cring.cring_simprules(16) cring.cring_simprules(9) domain.axioms(1) + mem_Collect_eq q_ball_def) + qed + have A8: "ac n (x \ c) = ac n (d \ c)" + proof- + have A80: "(x \ c) \ nonzero Q\<^sub>p" + by (metis (no_types, lifting) A A4 A5 A7 Qp.domain_axioms + assms(4) cring.cring_simprules(4) domain.axioms(1) + mem_Collect_eq c_ball_def val_nonzero) + have A81: "(d \ c) \ nonzero Q\<^sub>p" + by (metis (no_types, lifting) A3 Qp.domain_axioms assms(4) assms(5) + cring.cring_simprules(4) domain.axioms(1) mem_Collect_eq not_nonzero_Qp q_ball_def) + have A82: "n + m= val (x \ c) + n" + by (simp add: A7) + show ?thesis + using A0 A4 A7 ac_val[of "(x \ c)" "(d \ c)" n] A A80 A81 Qp_diff_diff assms(4) assms(5) c_ballE(1) q_ballE + by auto + qed + show ?thesis using A8 A3 A7 A2 q_ball_def[of n k m c] q_ballI[of x n c k m] + by (metis (no_types, lifting) A A4 A5 Qp.minus_closed assms(4) c_ballE(1) eint.inject val_nonzero val_ord) + qed + qed +qed + +definition is_ball :: "padic_number set \ bool" where +"is_ball B = (\(m::int). \ c \ carrier Q\<^sub>p. (B = B\<^bsub>m\<^esub>[c]))" + +lemma is_ball_imp_in_Qp: + assumes "is_ball B" + shows "B \ carrier Q\<^sub>p" + unfolding is_ball_def + using assms c_ball_in_Qp is_ball_def + by auto + +lemma c_ball_centers: + assumes "is_ball B" + assumes "B = B\<^bsub>n\<^esub>[c]" + assumes "d \ B" + assumes "c \ carrier Q\<^sub>p" + shows "B = B\<^bsub>n\<^esub>[d]" +proof + show "B \ B\<^bsub>n\<^esub>[d]" + proof + fix x + assume A0: "x \ B" + have "val (x \ d) \ n" + proof- + have A00: "val (x \ c) \ n" + using A0 assms(2) c_ballE(2) by blast + have A01: "val (d \ c) \ n" + using assms(2) assms(3) c_ballE(2) by blast + then show ?thesis + using Qp_isosceles[of x c d "n"] assms A0 A00 c_ballE(1) + by blast + qed + then show "x \ B\<^bsub>n\<^esub>[d]" + using A0 assms(1) c_ballI is_ball_imp_in_Qp + by blast + qed + show "B\<^bsub>n\<^esub>[d] \ B" + proof + fix x + assume "x \ B\<^bsub>n\<^esub>[d]" + show "x \ B" + using Qp_isosceles[of x d c "n"] + assms + unfolding c_ball_def + by (metis (no_types, lifting) Qp.domain_axioms Qp_isosceles \x \ B\<^bsub>n\<^esub>[d]\ + a_minus_def assms(2) c_ballE(2) c_ballI cring.cring_simprules(17) domain.axioms(1) + c_ballE(1)) + qed +qed + +lemma c_ball_center_in: + assumes "is_ball B" + assumes "B = B\<^bsub>n\<^esub>[c]" + assumes "c \ carrier Q\<^sub>p" + shows "c \ B" + using assms unfolding c_ball_def + by (metis (no_types, lifting) Qp.r_right_minus_eq assms(2) c_ballI eint_ord_code(3) local.val_zero) + +text \Every point a has a point b of distance exactly n away from it.\ +lemma dist_nonempty: + assumes "a \ carrier Q\<^sub>p" + shows "\b \ carrier Q\<^sub>p. val (b \ a) = eint n" +proof- + obtain b where b_def: "b = (\ [^] n) \ a" + by simp + have "val (b \a) = n" + using b_def assms + by (metis (no_types, lifting) Qp.domain_axioms a_minus_def cring.cring_simprules(16) + cring.cring_simprules(17) cring.cring_simprules(3) cring.cring_simprules(7) + domain.axioms(1) ord_p_pow_int p_intpow_closed(1) p_intpow_closed(2) val_ord) + then show ?thesis + by (metis Qp.domain_axioms assms b_def cring.cring_simprules(1) domain.axioms(1) p_intpow_closed(1)) +qed + +lemma dist_nonempty': + assumes "a \ carrier Q\<^sub>p" + shows "\b \ carrier Q\<^sub>p. val (b \ a) = \" +proof(cases "\ = \") + case True + then have "val (a \ a) = \" + using assms + by (metis Qp.r_right_minus_eq local.val_zero) + then show ?thesis + using assms + by blast +next + case False + then obtain n where n_def: "eint n = \" + by blast + then show ?thesis + using assms dist_nonempty[of a n] + by blast +qed + +lemma ball_rad_0: + assumes "is_ball B" + assumes "B\<^bsub>m\<^esub>[c] \ B\<^bsub>n\<^esub>[c]" + assumes "c \ carrier Q\<^sub>p" + shows "n \ m" +proof- + obtain b where b_def: "b \ carrier Q\<^sub>p \ val (b \c) = m" + by (meson assms(3) dist_nonempty) + then have "b \ B\<^bsub>n\<^esub>[c]" + using assms c_ballI + by auto + + then have "m \ n" + using Q\<^sub>p_def Zp_def b_def c_ballE(2) padic_integers_axioms + by force + then show ?thesis + by (simp ) +qed + +lemma ball_rad: + assumes "is_ball B" + assumes "B = B\<^bsub>n\<^esub>[c]" + assumes "B = B\<^bsub>m\<^esub>[c]" + assumes "c \ carrier Q\<^sub>p" + shows "n = m" +proof- + have 0: "n \m" + using assms ball_rad_0 + by (metis order_refl) + have 1: "m \n" + using assms ball_rad_0 + by (metis order_refl) + show ?thesis + using 0 1 + by auto +qed + +definition radius :: "padic_number set \ int" ("rad") where +"radius B = (SOME n. (\c \ carrier Q\<^sub>p . B = B\<^bsub>n\<^esub>[c]))" + +lemma radius_of_ball: + assumes "is_ball B" + assumes "c \ B" + shows "B = B\<^bsub>rad B\<^esub>[c]" +proof- + obtain d m where d_m_def: "d \ carrier Q\<^sub>p \ B = B\<^bsub>m\<^esub>[d]" + using assms(1) is_ball_def + by blast + then have "B = B\<^bsub>m\<^esub>[c]" + using assms(1) assms(2) c_ball_centers by blast + then have "rad B = m" + proof- + have "\n. (\c \ carrier Q\<^sub>p . B = B\<^bsub>n\<^esub>[c])" + using d_m_def by blast + then have "(\c \ carrier Q\<^sub>p . B = B\<^bsub>rad B\<^esub>[c])" + using radius_def[of B] + by (smt someI_ex) + then show ?thesis + using radius_def ball_rad[of B m ] + by (metis (mono_tags, lifting) \B = B\<^bsub>m\<^esub>[c]\ assms(1) assms(2) c_ballE(1) c_ball_centers) + qed + then show ?thesis + using \B = B\<^bsub>m\<^esub>[c]\ by blast +qed + +lemma ball_rad': + assumes "is_ball B" + assumes "B = B\<^bsub>n\<^esub>[c]" + assumes "B = B\<^bsub>m\<^esub>[d]" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + shows "n = m" + by (metis assms(1) assms(2) assms(3) assms(4) assms(5) ball_rad c_ball_center_in c_ball_centers) + +lemma nested_balls: + assumes "is_ball B" + assumes "B = B\<^bsub>n\<^esub>[c]" + assumes "B' = B\<^bsub>m\<^esub>[c]" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + shows "n \m \ B \ B'" +proof + show "m \ n \ B \ B'" + proof + assume A0: "m \n" + then have A0': "m \ n" + by (simp add: ) + fix x + assume A1: "x \ B" + show "x \ B'" + using assms c_ballI[of x m c] A0' A1 c_ballE(2)[of x n c] c_ball_in_Qp + by (meson c_ballE(1) dual_order.trans eint_ord_simps(1)) + qed + show "B \ B' \ m \ n" + using assms(1) assms(2) assms(3) assms(4) ball_rad_0 + by blast +qed + +lemma nested_balls': + assumes "is_ball B" + assumes "is_ball B'" + assumes "B \ B' \ {}" + shows "B \ B' \ B' \ B" +proof- + obtain b where b_def: "b \ B \ B'" + using assms(3) by blast + show "B \ B' \ B' \ B" + proof- + have "\ B \ B' \ B' \ B" + proof- + assume A: "\ B \ B' " + have 0: "B = B\<^bsub>rad B\<^esub>[b]" + using assms(1) b_def radius_of_ball by auto + have 1: "B' = B\<^bsub>rad B'\<^esub>[b]" + using assms(2) b_def radius_of_ball by auto + show "B' \ B" using 0 1 A nested_balls + by (smt IntD2 Q\<^sub>p_def Zp_def assms(1) assms(2) b_def + c_ballE(1) padic_integers_axioms) + qed + then show ?thesis by blast + qed +qed + +definition is_bounded:: "padic_number set \ bool" where +"is_bounded S = (\n. \c \ carrier Q\<^sub>p. S \ B\<^bsub>n\<^esub>[c] )" + +lemma empty_is_bounded: +"is_bounded {}" + unfolding is_bounded_def + by blast + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\$p$-adic Open Sets\ +(**************************************************************************************************) +(**************************************************************************************************) + + +definition is_open:: "padic_number set \ bool" where +"is_open U \ (U \ carrier Q\<^sub>p) \ (\c \ U. \n. B\<^bsub>n\<^esub>[c]\ U )" + +lemma is_openI: + assumes "U \carrier Q\<^sub>p" + assumes "\c. c \ U \ \n. B\<^bsub>n\<^esub>[c]\ U" + shows "is_open U" + by (simp add: assms(1) assms(2) is_open_def) + +lemma ball_is_open: + assumes "is_ball B" + shows "is_open B" + by (metis (mono_tags, lifting) assms is_ball_imp_in_Qp is_open_def radius_of_ball subset_iff) + +lemma is_open_imp_in_Qp: + assumes "is_open U" + shows "U \ carrier Q\<^sub>p" + using assms unfolding is_open_def + by linarith + +lemma is_open_imp_in_Qp': + assumes "is_open U" + assumes " x \ U" + shows "x \ carrier Q\<^sub>p" + using assms(1) assms(2) is_open_imp_in_Qp + by blast + +text\ + Owing to the total disconnectedness of the $p$-adic field, every open set can be decomposed + into a disjoint union of balls which are maximal with respect to containment in that set. + This unique decomposition is occasionally useful. +\ + +definition is_max_ball_of ::"padic_number set \ padic_number set \ bool" where +"is_max_ball_of U B \ (is_ball B) \ (B \ U) \ (\B'. ((is_ball B') \ (B' \ U) \ B \ B') \ B' \ B)" + +lemma is_max_ball_ofI: + assumes "U \ carrier Q\<^sub>p" + assumes "(B\<^bsub>m\<^esub>[c]) \ U" + assumes "c \ carrier Q\<^sub>p" + assumes "\m'. m' < m \ \ B\<^bsub>m'\<^esub>[c] \ U" + shows "is_max_ball_of U (B\<^bsub>m\<^esub>[c])" +proof(rule ccontr) + assume " \ is_max_ball_of U B\<^bsub>m\<^esub>[c]" + then have "\ (\B'. is_ball B' \ B' \ U \ B\<^bsub>m\<^esub>[c] \ B'\ B' \ B\<^bsub>m\<^esub>[c])" + using assms is_max_ball_of_def[of U "B\<^bsub>m\<^esub>[c]" ] + unfolding is_ball_def + by blast + then obtain B' where B'_def: "is_ball B' \ B' \ U \ B\<^bsub>m\<^esub>[c] \ B' \ \ B' \ B\<^bsub>m\<^esub>[c]" + by auto + obtain n where n_def: "B' = B\<^bsub>n\<^esub>[c]" + by (meson B'_def assms(3) c_ball_center_in is_ball_def radius_of_ball subset_iff) + then show False + using assms + by (smt B'_def Q\<^sub>p_def Zp_def ball_rad_0 padic_integers_axioms) +qed + +lemma int_prop: + fixes P:: "int \ bool" + assumes "P n" + assumes "\m. m \N \ \ P m" + shows "\n. P n \ (\n'. P n'\ n' \n)" +proof- + have "n > N" + by (meson assms(1) assms(2) not_less) + obtain k::nat where k_def: "k = nat (n - N)" + by blast + obtain l::nat where l_def: "l = (LEAST M. P (N + int M))" + by simp + have 0: " P (N + int l)" + by (metis (full_types) LeastI \N < n\ assms(1) l_def zless_iff_Suc_zadd) + have 1: "l > 0" + using "0" assms(2) of_nat_0_less_iff by fastforce + have 2: "\M. M < l \ \ P (N + M)" + by (metis Least_le l_def less_le_trans nat_less_le) + obtain n where n_def: "n = (N + int l)" + by simp + have "P n \ (\n'. P n'\ n' \ n)" + proof- + have "P n" + by (simp add: "0" n_def) + have "(\n'. P n'\ n' \ n)" + proof + fix n' + show "P n' \ n \ n'" + proof + assume "P n'" + show "n \n'" + proof(rule ccontr) + assume " \ n \ n'" + then have C: "n' < n" + by auto + show "False" + proof(cases "n' \ N") + case True + obtain M where M_def: "nat (n' - N) = M" + by simp + then have M0: "n' = N + int M " + using True by linarith + have M1: "M < l" + using n_def C M0 + by linarith + then show ?thesis using 2 M_def M0 M1 + using \P n'\ by auto + next + case False + then show ?thesis using assms + using \P n'\ by auto + qed + qed + qed + qed + then show ?thesis + using \P n\ by blast + qed + then show ?thesis by blast +qed + +lemma open_max_ball: + assumes "is_open U" + assumes "U \ carrier Q\<^sub>p" + assumes "c \ U" + shows "\B. is_max_ball_of U B \ c \ B" +proof- + obtain B where B_def: "is_ball B \ B \ U \ c \ B" + by (meson assms(1) assms(3) c_ball_center_in is_ball_def is_open_imp_in_Qp' is_open_def padic_integers_axioms) + show P: "\B. is_max_ball_of U B \ c \ B" + proof(rule ccontr) + assume C: "\B. is_max_ball_of U B \ c \ B" + show False + proof- + have C': "\B. c \ B \ \ is_max_ball_of U B " + using C + by auto + have C'': "\N. \m m\<^esub>[c] \ U " + proof + fix N + show "\mm\<^esub>[c] \ U" + proof(rule ccontr) + assume A: "\ (\mm\<^esub>[c] \ U)" + obtain P where P_def: "P = (\ n. \mm\<^esub>[c] \ U)" + by simp + have A0: "\n. P n" + by (metis B_def P_def gt_ex radius_of_ball) + have A1: "\m. m \N \ \ P m" + using A P_def by auto + have A2: "\n. P n \ (\n'. P n'\ n' \n)" + using A0 A1 int_prop + by auto + obtain n where n_def: " P n \ (\n'. P n'\ n' \n)" + using A2 by blast + have " B\<^bsub>n\<^esub>[c] \ U" + by (smt B_def P_def c_ball_def is_ball_def mem_Collect_eq n_def nested_balls order_trans) + obtain m where m_def: "m < n \B\<^bsub>m\<^esub>[c] \ U" + using P_def n_def by blast + have "m = n-1" + proof- + have "P (m +1)" + using P_def m_def + by auto + then have "m + 1 \ n" + using n_def by blast + then show ?thesis using m_def by auto + qed + have "\m'. m' < m \ \ B\<^bsub>m'\<^esub>[c] \ U" + proof + fix m' + show " m' < m \ \ B\<^bsub>m'\<^esub>[c] \ U" + proof + assume "m' < m" + show "\ B\<^bsub>m'\<^esub>[c] \ U" + proof + assume "B\<^bsub>m'\<^esub>[c] \ U" + then have "P (m' + 1)" + using P_def by auto + have "m'+ 1 < n" + using \m = n - 1\ \m' < m\ by linarith + then show False + using n_def \P (m' + 1)\ by auto + qed + qed + qed + then have "is_max_ball_of U B\<^bsub>m\<^esub>[c]" + using is_max_ball_ofI assms(1) assms(3) is_open_imp_in_Qp is_open_imp_in_Qp' m_def + by presburger + then show False + using C assms(1) assms(3) c_ball_center_in is_open_imp_in_Qp' + is_max_ball_of_def padic_integers_axioms + by blast + qed + qed + have "U = carrier Q\<^sub>p" + proof + show "carrier Q\<^sub>p \ U" + proof + fix x + assume "x \ carrier Q\<^sub>p" + show "x \ U" + proof(cases "x = c") + case True + then show ?thesis using assms by auto + next + case False + obtain m where m_def: "eint m = val(x \ c)" + using False + by (metis (no_types, opaque_lifting) Qp_diff_diff Qp.domain_axioms \x \ carrier Q\<^sub>p\ a_minus_def + assms(1) assms(3) cring.cring_simprules(16) cring.cring_simprules(17) + cring.cring_simprules(4) domain.axioms(1) is_open_imp_in_Qp' val_def val_minus) + obtain m' where m'_def: "m' < m \ B\<^bsub>m'\<^esub>[c] \ U " + using C'' + by blast + have "val (x \ c) \ m'" + by (metis eint_ord_simps(1) less_imp_le m'_def m_def) + then have "x \ B\<^bsub>m'\<^esub>[c]" + using \x \ carrier Q\<^sub>p\ c_ballI by blast + then show "x \ U" + using m'_def by blast + qed + qed + show "U \ carrier Q\<^sub>p " + using assms + by (simp add: is_open_imp_in_Qp) + qed + then show False using assms by auto + qed + qed +qed + +definition interior where + "interior U = {a. \B. is_open B \ B \ U \ a \ B}" + +lemma interior_subset: + assumes "U \ carrier Q\<^sub>p" + shows "interior U \ U" +proof + fix x + assume "x \ interior U" + show "x \ U" + proof- + obtain B where B_def: "is_open B \ B \ U \ x \ B" + using \x \ interior U\ interior_def + by auto + then show "x \ U" + by blast + qed +qed + +lemma interior_open: + assumes "U \ carrier Q\<^sub>p" + shows "is_open (interior U)" +proof(rule is_openI) + show "interior U \ carrier Q\<^sub>p" + using assms interior_subset by blast + show "\c. c \ interior U \ \n. B\<^bsub>n\<^esub>[c] \ interior U" + proof- + fix c + assume "c \ interior U" + show "\n. B\<^bsub>n\<^esub>[c] \ interior U" + proof- + obtain B where B_def: "is_open B \ B \ U \ c \ B" + using \c \ interior U\ interior_def padic_integers_axioms + by auto + then show ?thesis + proof - + obtain ii :: "((nat \ int) \ (nat \ int)) set \ ((nat \ int) \ (nat \ int)) set set \ int" + where + "B\<^bsub>ii c B\<^esub>[c] \ B" + by (meson B_def is_open_def) + then show ?thesis + using B_def interior_def padic_integers_axioms by auto +qed + qed + qed +qed + +lemma interiorI: + assumes "W \ U" + assumes "is_open W" + shows "W \ interior U" + using assms(1) assms(2) interior_def by blast + +lemma max_ball_interior: + assumes "U \ carrier Q\<^sub>p" + assumes "is_max_ball_of (interior U) B" + shows "is_max_ball_of U B" +proof(rule ccontr) + assume C: " \ is_max_ball_of U B" + then obtain B' where B'_def: "is_ball B' \ B' \ U \ B \ B' \ B \ B'" + by (metis (no_types, lifting) assms(1) assms(2) dual_order.trans + interior_subset is_max_ball_of_def ) + then have "B' \ interior U" + using interior_def padic_integers_axioms ball_is_open + by auto + then show False using assms B'_def is_max_ball_of_def[of "interior U" "B"] + by blast +qed + +lemma ball_in_max_ball: + assumes "U \ carrier Q\<^sub>p" + assumes "U \ carrier Q\<^sub>p" + assumes "c \ U" + assumes "\B. B \ U \ is_ball B \ c \ B" + shows "\B'. is_max_ball_of U B' \ c \ B'" +proof- + obtain B where " B \ U \ is_ball B \ c \ B" + using assms(4) + by blast + then have 0: "B \ interior U" + using ball_is_open interior_def by blast + have 1: "c \ interior U" + using "0" \B \ U \ is_ball B \ c \ B\ by blast + then have "\B'. is_max_ball_of (interior U) B' \ c \ B'" + using open_max_ball[of "interior U" c] assms(1) assms(2) interior_open interior_subset + by blast + then show ?thesis + using assms(1) max_ball_interior + by blast +qed + +lemma ball_in_max_ball': + assumes "U \ carrier Q\<^sub>p" + assumes "U \ carrier Q\<^sub>p" + assumes "B \ U \ is_ball B" + shows "\B'. is_max_ball_of U B' \ B \ B'" +proof- + obtain c where c_def: "c \ B" + by (metis assms(3) c_ball_center_in is_ball_def) + obtain B' where B'_def: " is_max_ball_of U B' \ c \ B'" + using assms ball_in_max_ball[of U c] c_def + by blast + then show ?thesis + by (meson assms(3) c_def disjoint_iff_not_equal nested_balls' + is_max_ball_of_def padic_integers_axioms) +qed + +lemma max_balls_disjoint: + assumes "U \ carrier Q\<^sub>p" + assumes "is_max_ball_of U B" + assumes "is_max_ball_of U B'" + assumes "B \B'" + shows "B \ B' = {}" + by (meson assms(2) assms(3) assms(4) nested_balls' is_max_ball_of_def + padic_integers_axioms subset_antisym) + +definition max_balls :: "padic_number set \ padic_number set set" where +"max_balls U = {B. is_max_ball_of U B }" + +lemma max_balls_interior: + assumes "U \ carrier Q\<^sub>p" + assumes "U \ carrier Q\<^sub>p" + shows "interior U = {x \ carrier Q\<^sub>p. (\B \ (max_balls U). x \ B)}" +proof + show "interior U \ {x \ carrier Q\<^sub>p. \B\max_balls U. x \ B}" + proof + fix x + assume A: " x \ interior U" + show "x \ {x \ carrier Q\<^sub>p. \B\max_balls U. x \ B}" + by (metis (mono_tags, lifting) A assms(1) assms(2) interior_open + interior_subset is_open_imp_in_Qp' max_ball_interior max_balls_def + mem_Collect_eq open_max_ball subset_antisym) + qed + show "{x \ carrier Q\<^sub>p. \B\max_balls U. x \ B} \ interior U" + proof + fix x + assume A: " x \ {x \ carrier Q\<^sub>p. \B\max_balls U. x \ B} " + show "x \ interior U" + proof- + obtain B where B_def: "B\max_balls U \ x \ B" + using A by blast + then have "B \ interior U" + by (metis (no_types, lifting) interior_def is_max_ball_of_def mem_Collect_eq + ball_is_open max_balls_def subsetI) + then show ?thesis + using B_def by blast + qed + qed +qed + +lemma max_balls_interior': + assumes "U \ carrier Q\<^sub>p" + assumes "U \ carrier Q\<^sub>p" + assumes "B \ max_balls U" + shows "B \ interior U" + using assms(1) assms(2) assms(3) is_max_ball_of_def max_balls_interior + max_balls_def padic_integers_axioms + by auto + +lemma max_balls_interior'': + assumes "U \ carrier Q\<^sub>p" + assumes "U \ carrier Q\<^sub>p" + assumes "a \ interior U" + shows "\B \ max_balls U. a \ B" + using assms(1) assms(2) assms(3) max_balls_interior + by blast + +lemma open_interior: + assumes "is_open U" + shows "interior U = U" + unfolding interior_def using assms + by blast + +lemma interior_idempotent: + assumes "U \ carrier Q\<^sub>p" + shows "interior (interior U) = interior U" + using assms interior_open[of U] open_interior[of "interior U"] + by auto + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Convex Subsets of the Value Group\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + The content of this section will be useful for defining and reasoning about $p$-adic cells in the + proof of Macintyre's theorem. It is proved that every convex set in the extended integers is + either an open ray, a closed ray, a closed interval, or a left-closed interval.\ + +definition is_convex :: "eint set \ bool" where +"is_convex A = (\ x \ A. \y \ A. \c. x \ c \ c \ y \ c \ A)" + +lemma is_convexI: + assumes "\x y c. x \ A \ y \ A \ x \ c \ c \y \ c \ A" + shows "is_convex A" + unfolding is_convex_def + using assms + by blast + +lemma is_convexE: + assumes "is_convex A" + assumes "x \ A" + assumes "y \ A" + assumes "x \ a" + assumes "a \ y" + shows "a \ A" + using assms is_convex_def + by blast + +lemma empty_convex: +"is_convex {}" + apply(rule is_convexI) + by blast + +lemma UNIV_convex: +"is_convex UNIV" + apply(rule is_convexI) + by blast + +definition closed_interval ("I[_ _]") where + "closed_interval \ \ = {a . \ \ a \ a \ \}" + +lemma closed_interval_is_convex: + assumes "A = closed_interval \ \" + shows "is_convex A" + apply(rule is_convexI) + using assms unfolding closed_interval_def + by auto + +lemma empty_closed_interval: +"{} = closed_interval \ (eint 1)" + unfolding closed_interval_def + by auto + +definition left_closed_interval where +"left_closed_interval \ \ = {a . \ \ a \ a < \}" + +lemma left_closed_interval_is_convex: + assumes "A = left_closed_interval \ \" + shows "is_convex A" + apply(rule is_convexI) + using assms unfolding left_closed_interval_def + using leD order.trans by auto + +definition closed_ray where +"closed_ray \ \ = {a . a \ \ }" + +lemma closed_ray_is_convex: + assumes "A = closed_ray \ \" + shows "is_convex A" + apply(rule is_convexI) + using assms unfolding closed_ray_def + by auto + +lemma UNIV_closed_ray: +"(UNIV::eint set)= closed_ray \ \" + unfolding closed_ray_def + by simp + +definition open_ray :: "eint \ eint \ eint set" where +"open_ray \ \ = {a . a < \ }" + +lemma open_ray_is_convex: + assumes "A = open_ray \ \" + shows "is_convex A" + apply(rule is_convexI) + using assms unfolding open_ray_def + using leD by auto + +lemma open_rayE: + assumes "a < \" + shows "a \ open_ray \ \" + unfolding open_ray_def using assms + by blast + +lemma value_group_is_open_ray: +"UNIV - {\} = open_ray \ \" +proof + show "UNIV - {\} \ open_ray \ \" + using open_rayE[of _ \ "\"] + by (simp add: open_rayE subset_eq) + show "open_ray \ \ \ UNIV - {\}" + unfolding open_ray_def + by blast +qed + +text\ + This is a predicate which identifies a certain kind of set-valued function on the extended + integers. Convex conditions will be important in the definition of $p$-adic cells later, and + it will be proved that every convex set is induced by a convex condition.\ +definition is_convex_condition :: "(eint \ eint \ eint set) \ bool" + where "is_convex_condition I \ + I = closed_interval \ I = left_closed_interval \ I = closed_ray \ I = open_ray" + +lemma convex_condition_imp_convex: + assumes "is_convex_condition I" + shows "is_convex (I \ \)" + using assms + unfolding is_convex_condition_def + by (meson closed_interval_is_convex closed_ray_is_convex left_closed_interval_is_convex open_ray_is_convex) + +lemma bounded_order: + assumes "(a::eint) < \" + assumes "b \ a" + obtains k::nat where "a = b + k" +proof- + obtain m::int where m_def: "a = m" + using assms(1) less_infinityE by blast + obtain n::int where n_def: "b = n" + using assms(2) eint_ile m_def by blast + have 0: "a - b = eint (m - n)" + by (simp add: m_def n_def) + then have "a = b + nat (m - n)" + using assms m_def n_def + by simp + thus ?thesis + using that by blast +qed + +text\Every convex set is given by a convex condition\ +lemma convex_imp_convex_condition: + assumes "is_convex A" + shows "\ I \ \. is_convex_condition I \ A = (I \ \)" +proof(cases "\ a \ A. \ b \ A. a \ b") + case True + then obtain \ where alpha_def: "\ \ A \ (\ b \ A. \ \ b)" + by blast + then show ?thesis + proof(cases "\ a \ A. \ b \ A. b \ a") + case True + then obtain \ where beta_def: "\ \ A \ (\ b \ A. b \ \)" + by blast + have "A = closed_interval \ \" + unfolding closed_interval_def + using alpha_def beta_def assms is_convexE[of A \ \] + by blast + then show ?thesis + using is_convex_condition_def + by blast + next + case False + have F0: "\a. \ \ a \ a < \ \ a \ A" + proof(rule ccontr) + assume A: "\ (\a. a \ \ \ a < \ \ a \ A)" + then obtain a where a_def: " \ \ a \ a < \ \ a \ A" + by blast + obtain n where n_def: "\ = eint n" + using False alpha_def by force + obtain m where m_def: "a = eint m" + using a_def less_infinityE by blast + have "\k::nat. \c. (\ + eint (int k)) < c \ c \ A" + proof fix k + show "\c>\ + eint (int k). c \ A" + apply(induction k) + using False alpha_def le_less_linear zero_eint_def apply fastforce + proof- fix k + assume IH: "\c>\ + eint (int k). c \ A" + then obtain c where c_def: "c>\ + eint (int k) \ c \ A" + by blast + then obtain c' where c'_def: "c' \ A \ c < c'" + using False + by (meson le_less_linear) + then have "(\ + eint (int (Suc k))) \ c'" + proof- + have 0: "eint (int (Suc k)) = eint (int k) + eint 1" + by simp + have "\ + eint (int k) \c'" + by (meson c'_def c_def le_less le_less_trans) + then have 1: "(\ + eint (int k)) < c'" + using c'_def c_def le_less_trans + by auto + have 2: "\ + eint (int k) + eint 1 = \ + eint (int (Suc k))" + using 0 + by (simp add: n_def) + then show "(\ + eint (int (Suc k))) \ c'" + by (metis "1" ileI1 one_eint_def) + qed + then show "\c>\ + eint (int (Suc k)). c \ A" + using False c'_def not_less by fastforce + qed + qed + obtain k where k_def: "k = a - \" + using a_def n_def + by blast + hence "k \ 0" + using a_def n_def + by (metis alpha_def eint_minus_le le_less) + hence 0: "\n::nat. k = n" + using a_def n_def k_def + by (metis eint.distinct(2) eint_0_iff(2) eint_add_cancel_fact eint_ord_simps(1) fromeint.cases m_def nonneg_int_cases plus_eint_simps(3)) + have 1: "a = \ + k" + using k_def a_def n_def + by simp + then obtain c where "c \ A \ a < c" + by (metis "0" \\k. \c>\ + eint (int k). c \ A\ the_eint.simps) + then have "a \ A" + using is_convexE[of A \ a c] a_def alpha_def assms is_convexE + by (meson linear not_less) + then show False + using a_def by blast + qed + have "A = closed_interval \ \ \ A = left_closed_interval \ \" + apply(cases "\ \ A") + using False eint_ord_code(3) apply blast + proof- + assume A0: "\ \ A " + have "A = left_closed_interval \ \" + proof + show "A \ left_closed_interval \ \" + unfolding left_closed_interval_def + using alpha_def A0 + by (metis (mono_tags, lifting) False eint_ord_code(3) le_less_linear less_le_trans mem_Collect_eq subsetI) + show "left_closed_interval \ \ \ A" + unfolding left_closed_interval_def + using alpha_def A0 F0 + by blast + qed + then show ?thesis + by blast + qed + then show ?thesis + unfolding is_convex_condition_def + by blast + qed +next + case False + show ?thesis apply(cases "A = {}") + using empty_closed_interval is_convex_condition_def apply blast + proof- + assume A0: "A \ {}" + have "A \ {\}" + using False + by blast + then obtain \ where alpha_def: "\ \ A \ \ \\" + using A0 + by fastforce + have A1: "\k::nat. \ b \ A. b + eint (int k) \ \" + proof- fix k + show " \ b \ A. b + eint (int k) \ \" + proof(induction k) + case 0 + then have "\ + eint (int 0) = \" + by (simp add: zero_eint_def) + then show ?case + using alpha_def by auto + next + case (Suc k) fix k + assume IH: "\b\A. \ \ b + eint (int k)" + show "\b\A. \ \ b + eint (int (Suc k))" + proof- + obtain b where b_def: "b \ A \ \ \ b + eint (int k)" + using IH by blast + then obtain c where c_def: "c \ A \ c < b" + using False le_less_linear by blast + have 0: "(c + eint (int (Suc k))) < (b + eint (int (Suc k)))" + using c_def + by simp + have 1: "b + eint (int (Suc k)) = (b + eint (int k)) + eint 1" + by simp + then show ?thesis + by (metis "0" b_def c_def eSuc_ile_mono ileI1 le_less one_eint_def) + qed + qed + qed + show ?thesis + proof(cases "\ a \ A. \ b \ A. b \ a") + case True + then obtain \ where beta_def: "\ \ A \ (\ b \ A. b \ \)" + by blast + have "A = closed_ray \ \" + unfolding closed_ray_def + proof + show "A \ {a. \ \ a}" + using assms beta_def + by blast + show "{a. \ \ a} \ A" + proof fix x assume "x \ {a. \ \ a}" + then have 0: "\ \ x" by blast + show "x \ A" + proof(cases "x \ \") + case True + obtain n where n_def: "\= eint n" + using alpha_def + by blast + obtain m where m_def: "x = eint m" + using True eint_ile n_def by blast + have 1: "m \ n" + using True m_def n_def + by simp + have 2: "eint (int (nat (n - m))) = eint (n - m)" + by (simp add: "1") + then obtain b where b_def: "b \ A \ b + eint (n - m) \ \" + using A1[of "nat (n - m)"] + by (metis) + then have "b + eint (n - m) \ x + eint (n - m)" + using b_def + by (simp add: m_def n_def) + then have "b \ x" + by auto + then show ?thesis + using "0" assms b_def beta_def is_convex_def + by blast + next + case False + then show ?thesis + using "0" alpha_def assms beta_def is_convexE + by (meson linear) + qed + qed + qed + then show ?thesis + using is_convex_condition_def + by blast + next + case f: False + have F0: "\a. \ \ a \ a < \ \ a \ A" + proof(rule ccontr) + assume A: "\ (\a. a \ \ \ a < \ \ a \ A)" + then obtain a where a_def: " \ \ a \ a < \ \ a \ A" + by blast + obtain n where n_def: "\ = eint n" + using alpha_def by blast + obtain m where m_def: "a = eint m" + using a_def less_infinityE by blast + have 0: "\k::nat. \c. (\ + eint (int k)) < c \ c \ A" + proof fix k + show "\c>\ + eint (int k). c \ A" + apply(induction k) + using alpha_def f le_less_linear apply fastforce + proof- fix k + assume IH: "\c>\ + eint (int k). c \ A" + then obtain c where c_def: "c>\ + eint (int k) \ c \ A" + by blast + then obtain c' where c'_def: "c' \ A \ c < c'" + using False f le_less_linear by blast + then have "(\ + eint (int (Suc k))) \ c'" + proof- + have 0: "eint (int (Suc k)) = eint (int k) + eint 1" + by simp + have "\ + eint (int k) \c'" + using c_def c'_def dual_order.strict_trans le_less by blast + then have 1: "(\ + eint (int k)) < c'" + using c'_def c_def le_less_trans by auto + have 2: "\ + eint (int k) + eint 1 = \ + eint (int (Suc k))" + using 0 by (simp add: n_def) + then show "(\ + eint (int (Suc k))) \ c'" + by (metis "1" ileI1 one_eint_def) + qed + then show "\c>\ + eint (int (Suc k)). c \ A" + using False c'_def + by (smt c_def eSuc_eint iadd_Suc_right ileI1 le_less of_nat_Suc) + qed + qed + obtain k::nat where "a = \ + eint (int k)" + using bounded_order a_def + by blast + then obtain c where "c \ A \ a A" + using is_convexE[of A \ a c] a_def alpha_def assms is_convexE + by (meson linear not_less) + then show False + using a_def by blast + qed + have "A = UNIV - {\}" + proof + show "A \ UNIV - {\}" + using f by auto + show "UNIV - {\} \ A" + proof fix x ::eint + assume A: "x \ UNIV - {\}" + show "x \ A" + proof(cases "x \ \") + case True + obtain k::nat where k_def: "x + k = \" + by (metis True alpha_def bounded_order eint_ord_simps(4)) + obtain c where c_def: "c \ A \ c + k = \" + by (metis A1 True add.commute alpha_def assms eint_add_left_cancel_le is_convexE k_def not_eint_eq) + have "x = c" + using k_def c_def + by auto + thus ?thesis + by (simp add: c_def) + next + case False + thus ?thesis + using A F0 by auto + qed + qed + qed + then show ?thesis + by (meson is_convex_condition_def value_group_is_open_ray) + qed + qed +qed + +lemma ex_val_less: + shows "\ (\::eint). \ < \" + apply(induction \) + using eint_ord_simps(2) lt_ex apply blast + using eint_ord_simps(4) by blast + +lemma ex_dist_less: + assumes "c \ carrier Q\<^sub>p" + shows "\ a \ carrier Q\<^sub>p. val (a \ c) < \" + using ex_val_less[of \] assms + by (metis dist_nonempty' ex_val_less) +end +end \ No newline at end of file diff --git a/thys/Padic_Field/Padic_Fields.thy b/thys/Padic_Field/Padic_Fields.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Padic_Fields.thy @@ -0,0 +1,3063 @@ +theory Padic_Fields + imports Fraction_Field Padic_Ints.Hensels_Lemma + +begin + + +(**************************************************************************************************) +(**************************************************************************************************) +section\Constructing the $p$-adic Valued Field\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + As a field, we can define the field $\mathbb{Q}_p$ immediately as the fraction field of + $\mathbb{Z}_p$. The valuation can then be extended from to $\mathbb{Z}_p$ to $\mathbb{Q}_p$ by + defining $\text{val}(a/b) = \text{val}\ a - \text{val}\ b$ where $a, b \in \mathbb{Z}_p$. \ + +(**********************************************************************************************) +(**************************************************************************************************) +subsection\A Locale for $p$-adic Fields\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + This section builds a locale for reasoning about general $p$-adic fields for a fixed $p$. + The locale fixes constants for the ring of $p$-adic integers ($\mathbb{Z}_p$) and the inclusion + map $\iota: \mathbb{Z}_p \to \mathbb{Q}_p$. \ +type_synonym padic_number = "((nat \ int) \ (nat \ int)) set" +locale padic_fields= +fixes Q\<^sub>p:: "_ ring" (structure) +fixes Z\<^sub>p:: "_ ring" (structure) +fixes p +fixes \ +defines "Z\<^sub>p \ padic_int p" +defines "Q\<^sub>p \ Frac Z\<^sub>p" +defines "\ \ domain_frac.inc Z\<^sub>p" +assumes prime: "prime p" + +sublocale padic_fields < Zp?: domain_frac Z\<^sub>p + by (simp add: Z\<^sub>p_def domain_frac.intro padic_int_is_domain prime) + +sublocale padic_fields < Qp?: ring Q\<^sub>p + unfolding Q\<^sub>p_def + by (simp add: Fraction_Field.domain_frac_def domain_axioms domain_frac.fraction_field_is_field field.is_ring) + +sublocale padic_fields < Qp?: cring Q\<^sub>p + unfolding Q\<^sub>p_def + by (simp add: Fraction_Field.domain_frac_def domain.axioms(1) domain_axioms domain_frac.fraction_field_is_domain) + +sublocale padic_fields < Qp?: field Q\<^sub>p + unfolding Q\<^sub>p_def + by (simp add: Fraction_Field.domain_frac_def domain_axioms domain_frac.fraction_field_is_field) + +sublocale padic_fields < Qp?: domain Q\<^sub>p + by (simp add: Qp.domain_axioms) + +sublocale padic_fields < padic_integers Z\<^sub>p +apply (simp add: padic_integers_def prime) +using Z\<^sub>p_def by auto + +sublocale padic_fields < UPQ?: UP_cring Q\<^sub>p "UP Q\<^sub>p" + using Qp.is_cring UP_cring_def apply blast + by auto + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\The Valuation Ring in $\mathbb{Q}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + The valuation ring $\mathcal{O}_p$ is the subring of elements in $\mathbb{Q}_p$ with positive + valuation. It is an isomorphic copy of $\mathbb{Z}_p$.\ +context padic_fields +begin + +abbreviation \\<^sub>p where +"\\<^sub>p \ \ ` carrier Z\<^sub>p" + +lemma inc_closed: + assumes "a \ carrier Z\<^sub>p" + shows "\ a \ carrier Q\<^sub>p" + using Q\<^sub>p_def \_def assms Zp.inc_is_hom ring_hom_closed + by fastforce + +lemma inc_is_hom: +"\ \ ring_hom Z\<^sub>p Q\<^sub>p" + unfolding Q\<^sub>p_def \_def +by(rule Zp.inc_is_hom) + +text\An alternate formula of the map $\iota$\ + +lemma inc_def: + assumes "a \ carrier Z\<^sub>p" + shows "\ a = frac a \\<^bsub>Z\<^sub>p\<^esub>" + using assms inc_equation[of a] \_def by auto + +lemma inc_of_nonzero: + assumes "a \ nonzero Z\<^sub>p" + shows "\ a \ nonzero Q\<^sub>p" + using \_def assms Q\<^sub>p_def Qp.nonzero_memI + Zp.nonzero_closed Zp.nonzero_one_closed frac_closed local.inc_def nonzero_fraction + by (metis Zp.nonzero_memE(2) inc_closed inc_inj1) + +lemma inc_of_one: +"\ \\<^bsub>Z\<^sub>p\<^esub> = \" + by (simp add: inc_is_hom ring_hom_one) + +lemma inc_of_zero: +"\ \\<^bsub>Z\<^sub>p\<^esub> = \" + apply(rule ring_hom_zero[of \], rule inc_is_hom) + apply (simp add: Zp.ring_axioms) + by (simp add: Qp.ring_axioms) + +lemma inc_of_sum: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ carrier Z\<^sub>p" + shows "\ (a \\<^bsub>Z\<^sub>p\<^esub> b) = (\ a) \ (\ b)" +by(rule ring_hom_add[of _ Z\<^sub>p], rule inc_is_hom, rule assms, rule assms) + +lemma inc_of_prod: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ carrier Z\<^sub>p" + shows "\ (a \\<^bsub>Z\<^sub>p\<^esub> b) = (\ a) \ (\ b)" + by (simp add: assms(1) assms(2) inc_is_hom ring_hom_mult) + +lemma inc_pow: + assumes "a \ nonzero Z\<^sub>p" + shows "\ (a[^]\<^bsub>Z\<^sub>p\<^esub>(n::nat)) = (\ a)[^] n" + apply(induction n) + apply (simp add: inc_of_one) + by (simp add: assms inc_of_prod Zp.nonzero_closed) + +lemma inc_of_diff: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ carrier Z\<^sub>p" + shows "\ (a \\<^bsub>Z\<^sub>p\<^esub> b) = (\ a) \ (\ b)" + using assms unfolding a_minus_def + using inc_is_hom Qp.ring_axioms Q\<^sub>p_def Zp.ring_hom_a_inv inc_of_sum by fastforce + +lemma Units_nonzero_Qp: +assumes "u \ Units Q\<^sub>p" +shows "u \ nonzero Q\<^sub>p" + by (simp add: Qp.Units_nonzero assms) + +lemma Units_eq_nonzero: + "Units Q\<^sub>p = nonzero Q\<^sub>p" + using frac_nonzero_Units unfolding Q\<^sub>p_def Z\<^sub>p_def +by blast + +lemma Units_inverse_Qp: + assumes "u \ Units Q\<^sub>p" + shows "inv\<^bsub>Q\<^sub>p\<^esub> u \ Units Q\<^sub>p" + using Q\<^sub>p_def Units_eq_nonzero assms frac_nonzero_inv_Unit by auto + +lemma nonzero_inverse_Qp: + assumes "u \ nonzero Q\<^sub>p" + shows "inv\<^bsub>Q\<^sub>p\<^esub> u \ nonzero Q\<^sub>p" + using Units_eq_nonzero Units_inverse_Qp assms by auto + +lemma frac_add: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "c \ carrier Z\<^sub>p" + assumes "d \ nonzero Z\<^sub>p" + shows "(frac a b) \ (frac c d) = (frac ((a \\<^bsub>Z\<^sub>p\<^esub> d) \\<^bsub>Z\<^sub>p\<^esub> (b \\<^bsub>Z\<^sub>p\<^esub> c)) (b \\<^bsub>Z\<^sub>p\<^esub> d))" + by (simp add: Q\<^sub>p_def assms(1) assms(2) assms(3) assms(4) local.frac_add) + +lemma frac_add_common_denom: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ carrier Z\<^sub>p" + assumes "c \ nonzero Z\<^sub>p" + shows "(frac a c) \ (frac b c) = frac (a \\<^bsub>Z\<^sub>p\<^esub> b) c" + by (simp add: Q\<^sub>p_def assms(1) assms(2) assms(3) frac_add_common_denom) + +lemma frac_mult: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "c \ carrier Z\<^sub>p" + assumes "d \ nonzero Z\<^sub>p" + shows "(frac a b) \ (frac c d) = (frac (a \\<^bsub>Z\<^sub>p\<^esub> c) (b \\<^bsub>Z\<^sub>p\<^esub> d))" + by (simp add: Q\<^sub>p_def assms(1) assms(2) assms(3) assms(4) frac_mult) + +lemma frac_one: + assumes "a \ nonzero Z\<^sub>p" + shows "frac a a = \" + by (simp add: Q\<^sub>p_def assms frac_one) + +lemma frac_closed: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "frac a b \ carrier Q\<^sub>p" + by (simp add: Q\<^sub>p_def assms(1) assms(2) frac_closed) + +lemma inv_in_frac: + assumes "a \ carrier Q\<^sub>p" + assumes "a \\" + shows "inv\<^bsub>Q\<^sub>p\<^esub> a \ carrier Q\<^sub>p" + "inv\<^bsub>Q\<^sub>p\<^esub> a \\" + "inv\<^bsub>Q\<^sub>p\<^esub> a \ nonzero Q\<^sub>p" + apply (simp add: assms(1) assms(2) field_inv(3)) + using assms(1) assms(2) field_inv(1) apply fastforce + using Qp.not_nonzero_memE assms(1) assms(2) nonzero_inverse_Qp by blast + +lemma nonzero_numer_imp_nonzero_fraction: + assumes "a \ nonzero Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "frac a b \ \" + by (simp add: Q\<^sub>p_def assms(1) assms(2) nonzero_fraction) + +lemma nonzero_fraction_imp_numer_not_zero: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "frac a b \ \" + shows "a \ \\<^bsub>Z\<^sub>p\<^esub>" + using assms fraction_zero Q\<^sub>p_def by blast + +lemma nonzero_fraction_imp_nonzero_numer: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "frac a b \ \" + shows "a \ nonzero Z\<^sub>p" + using assms(1) assms(2) assms(3) nonzero_fraction_imp_numer_not_zero not_nonzero_Zp by blast + +lemma(in padic_fields) frac_inv_id: + assumes "a \ nonzero Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "c \ nonzero Z\<^sub>p" + assumes "d \ nonzero Z\<^sub>p" + assumes "frac a b = frac c d" + shows "frac b a = frac d c" + using frac_inv assms + by metis + +lemma(in padic_fields) frac_uminus: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "\ (frac a b) = frac (\\<^bsub>Z\<^sub>p\<^esub> a) b" + by (simp add: Q\<^sub>p_def assms(1) assms(2) frac_uminus) + +lemma(in padic_fields) i_mult: + assumes "a \ carrier Z\<^sub>p" + assumes "c \ carrier Z\<^sub>p" + assumes "d \ nonzero Z\<^sub>p" + shows "(\ a) \ (frac c d) = frac (a \\<^bsub>Z\<^sub>p\<^esub> c) d" + by (simp add: Q\<^sub>p_def \_def assms(1) assms(2) assms(3) i_mult) + +lemma numer_denom_facts: + assumes "a \ carrier Q\<^sub>p" + shows "(numer a) \ carrier Z\<^sub>p" + "(denom a) \ nonzero Z\<^sub>p" + "a \ \ \ numer a \ \\<^bsub>Z\<^sub>p\<^esub>" + "a \ (\ (denom a)) = \ (numer a)" + "a = frac (numer a) (denom a)" + unfolding Q\<^sub>p_def + using Q\<^sub>p_def assms numer_denom_facts(2) apply auto[1] + using Q\<^sub>p_def assms numer_denom_facts(3) apply blast + using Q\<^sub>p_def assms numer_denom_facts(4) apply blast + using Q\<^sub>p_def \_def assms numer_denom_facts(5) apply auto[1] + using Q\<^sub>p_def assms numer_denom_facts(1) by auto + +lemma get_common_denominator: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + obtains a b c where + "a \ carrier Z\<^sub>p" + "b \ carrier Z\<^sub>p" + "c \ nonzero Z\<^sub>p" + "x = frac a c" + "y = frac b c" + using Q\<^sub>p_def assms(1) assms(2) common_denominator[of x y] + by blast + +abbreviation fract :: "_ \ _ \ _" (infixl "\
" 50) where +"(fract a b) \ (a \ (inv\<^bsub>Q\<^sub>p\<^esub> b))" + +text\fract generalizes frac\ + +lemma fract_frac: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "(frac a b) = (\ a \
\ b)" +proof- + have B: "b \ carrier Z\<^sub>p" + using Zp.nonzero_closed assms(2) by auto + have P0:"(inv\<^bsub>Q\<^sub>p\<^esub> (\ b)) = frac \\<^bsub>Z\<^sub>p\<^esub> b" + by (simp add: B Q\<^sub>p_def Zp.nonzero_one_closed assms(2) frac_inv local.inc_def) + have P1: "(frac a b) = (\ a) \ (frac \\<^bsub>Z\<^sub>p\<^esub> b)" + by (simp add: assms(1) assms(2) i_mult) + show ?thesis + by (simp add: P0 P1) +qed + +lemma frac_eq: + assumes "a \ nonzero Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "frac a b = \" + shows "a = b" +proof- + have "frac a b = frac b b" + by (simp add: assms(2) assms(3) frac_one) + then have "frac a \\<^bsub>Z\<^sub>p\<^esub> = frac b \\<^bsub>Z\<^sub>p\<^esub>" + using assms + by (metis (no_types, lifting) Zp.nonzero_closed + Zp.nonzero_one_closed frac_eqE frac_eqI') + then show ?thesis + using \_def assms(1) assms(2) inc_inj2 local.inc_def + by (simp add: Zp.nonzero_closed) +qed + +lemma fract_cancel_right: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "b \ (a \
b) = a" + by (simp add: Qp.Units_closed Qp.m_lcomm Units_eq_nonzero assms(1) + assms(2)) + +lemma fract_cancel_left: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "(a \
b) \ b = a" + by (simp add: Qp.m_comm Qp.nonzero_closed assms(1) assms(2) + local.fract_cancel_right nonzero_inverse_Qp) + +lemma fract_mult: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ nonzero Q\<^sub>p" + shows "(a \
b) \ (c \
d) = ((a \ c)\
(b \ d))" + using Q\<^sub>p_def assms(1) assms(2) assms(3) assms(4) + by (simp add: fract_mult) + +lemma Qp_nat_pow_nonzero: + assumes "x \ nonzero Q\<^sub>p" + shows "x[^](n::nat) \ nonzero Q\<^sub>p" + using Qp.Units_pow_closed Units_eq_nonzero assms by auto + +lemma Qp_nonzero_nat_pow: + assumes "x \ carrier Q\<^sub>p" + assumes "n > 0" + assumes "x[^](n::nat) \ nonzero Q\<^sub>p" + shows "x \ nonzero Q\<^sub>p" + using Frac_nonzero_nat_pow Q\<^sub>p_def assms(1) assms(2) assms(3) by blast + +lemma Qp_int_pow_nonzero: + assumes "x \ nonzero Q\<^sub>p" + shows "x[^](n::int) \ nonzero Q\<^sub>p" + using Frac_int_pow_nonzero Q\<^sub>p_def assms by blast + +lemma Qp_nonzero_int_pow: + assumes "x \ carrier Q\<^sub>p" + assumes "n > 0" + assumes "x[^](n::int) \ nonzero Q\<^sub>p" + shows "x \ nonzero Q\<^sub>p" + using Frac_nonzero_int_pow Q\<^sub>p_def assms + by auto + +lemma pow_p_frac_0: + assumes "(m::int) \ n" + assumes "n \0" + shows "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))" +proof- + have 0: "\

\carrier Z\<^sub>p" + by (simp add: Zp.nat_inc_closed) + have 1: "m - n \0" + using assms(1) by auto + have 2: "nat (m - n) + (nat n) = nat m" + using "1" assms(2) by linarith + have 3: "m \0" + using assms by auto + then have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(nat (m-n))) \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>(nat n)) = (\

[^]\<^bsub>Z\<^sub>p\<^esub>(nat m))" + by (simp add: "2" p_natpow_prod) + then have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) = (\

[^]\<^bsub>Z\<^sub>p\<^esub>m)" + using int_pow_int 1 3 assms(2) int_nat_eq by metis + then have P0: "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) = frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))\\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)" + by simp + have "\

\carrier Z\<^sub>p" + by (simp add: "0") + have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(nat n)) = [(p^(nat n))] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: p_pow_rep0) + then have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(nat n)) \ carrier Z\<^sub>p" + by (simp add: Zp_nat_inc_closed) + then have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \ carrier Z\<^sub>p" + using assms(2) by (metis int_nat_eq int_pow_int) + then have P1: "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) = frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))\\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) ((\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)))" + by (simp add: \[p] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> [^]\<^bsub>Z\<^sub>p\<^esub> (m - n) \\<^bsub>Z\<^sub>p\<^esub> [p] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> [^]\<^bsub>Z\<^sub>p\<^esub> n = [p] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> [^]\<^bsub>Z\<^sub>p\<^esub> m\) + have P2: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \ carrier Z\<^sub>p" + using "1" p_pow_car by blast + have P3: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \ carrier Z\<^sub>p" + using \\

[^]\<^bsub>Z\<^sub>p\<^esub> n \ carrier Z\<^sub>p\ + by blast + have P4: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \ nonzero Z\<^sub>p" + using assms(2) p_int_pow_nonzero + by blast + have P5: "\\<^bsub>Z\<^sub>p\<^esub> \ nonzero Z\<^sub>p" + using nonzero_def + by (simp add: Zp.nonzero_one_closed) + have "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) \ (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) + = frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))\\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) ((\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)))" + by (simp add: P2 P3 P4 Zp.nonzero_one_closed local.frac_mult) + then have "frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))\\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) ((\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>n))) = (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) " + by (simp add: P2 P4 Zp.nonzero_one_closed frac_cancel_lr mult_comm) + then have P6: "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) = (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) " + using P1 by blast + have "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n))" + using inc_def by (simp add: P2) + then show ?thesis + using P6 by blast +qed + +lemma pow_p_frac: + assumes "(m::int) \ n" + assumes "m \0" + shows "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) = frac \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>(n-m))" +proof- + have "(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) (\

[^]\<^bsub>Z\<^sub>p\<^esub>m)) = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub>(n-m))" + by (simp add: assms(1) assms(2) pow_p_frac_0) + then have P0:"(frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) (\

[^]\<^bsub>Z\<^sub>p\<^esub>m)) = frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(n-m)) \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: assms(1) local.inc_def p_pow_car) + have P1: "\\<^bsub>Z\<^sub>p\<^esub> \nonzero Z\<^sub>p" + by (simp add: Zp.nonzero_one_closed) + have P2: "\

[^]\<^bsub>Z\<^sub>p\<^esub>n \ nonzero Z\<^sub>p" + using assms(1) assms(2) p_int_pow_nonzero by auto + have P3: "\

[^]\<^bsub>Z\<^sub>p\<^esub>m \ nonzero Z\<^sub>p" + by (simp add: assms(2) p_int_pow_nonzero) + have P4: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(n-m)) \ nonzero Z\<^sub>p" + by (simp add: assms(1) p_int_pow_nonzero) + show " frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) = frac \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>(n-m))" + using P0 P1 P2 P3 P4 p_pow_nonzero + by (meson local.frac_inv_id) +qed + +text\The copy of the prime \p\ living in $\mathbb{Q}_p$:\ + +abbreviation \ where +"\ \ [p] \\<^bsub>Q\<^sub>p\<^esub> \" + +lemma(in domain_frac) frac_inc_of_nat: +"Frac_inc R ([(n::nat)]\ \) = [n]\\<^bsub>Frac R\<^esub>\\<^bsub>Frac R\<^esub>" + by (simp add: inc_equation nat_inc_rep) + +lemma inc_of_nat: +"(\ ([(n::nat)]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>)) = [n]\\<^bsub>Q\<^sub>p\<^esub>\" + unfolding Q\<^sub>p_def \_def + using frac_inc_of_nat[of n] + by auto + +lemma(in domain_frac) frac_inc_of_int: +"Frac_inc R ([(n::int)]\ \) = [n]\\<^bsub>Frac R\<^esub>\\<^bsub>Frac R\<^esub>" + apply(induction n) + apply (simp add: add_pow_int_ge inc_equation nat_inc_rep) + by (simp add: add_pow_int_lt frac_uminus inc_equation nat_inc_rep nonzero_one_closed) + +lemma inc_of_int: +"(\ ([(n::int)]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>)) = [n]\\<^bsub>Q\<^sub>p\<^esub>\" + unfolding Q\<^sub>p_def \_def + using frac_inc_of_int[of n] + by auto + +lemma p_inc: +"\ = \ \

" + by (simp add: inc_of_int) + +lemma p_nonzero: +"\ \ nonzero Q\<^sub>p" + using Z\<^sub>p_def Zp_nat_inc_closed inc_of_nonzero ord_Zp_p p_inc + p_nonzero by auto + +lemma p_natpow_inc: + fixes n::nat + shows "\[^]n = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub> n)" + by (simp add: Qp.int_nat_pow_rep inc_of_int p_pow_rep0) + +lemma p_intpow_inc: + fixes n::int + assumes "n \0" + shows "\[^]n = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub> n)" + using p_natpow_inc + by (metis assms int_nat_eq int_pow_int) + +lemma p_intpow: + fixes n::int + assumes "n < 0" + shows "\[^]n = (frac \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub> (-n)))" +proof- + have U0: "(\ [^] (nat (-n))) \ Units Q\<^sub>p" + using Qp.Units_pow_closed Units_eq_nonzero p_nonzero by blast + have E0: "(\ [^] (nat (-n))) = (\ [^] (-n))" + using assms by (simp add: int_pow_def nat_pow_def) + then have U1: "(\ [^] (-n)) \ Units Q\<^sub>p" using U0 + by simp + have "(\[^]n) = inv \<^bsub>Q\<^sub>p\<^esub> (\ [^] (nat (-n)))" + using assms by (simp add: int_pow_def nat_pow_def) + then have "(\[^]n) = inv \<^bsub>Q\<^sub>p\<^esub> (\ [^] (-n))" + using E0 by simp + then have "(\[^]n) = inv \<^bsub>Q\<^sub>p\<^esub> \ (\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n))" + using assms p_intpow_inc by auto + then have E1: "(\[^]n) = inv \<^bsub>Q\<^sub>p\<^esub> frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n)) \\<^bsub>Z\<^sub>p\<^esub>" + using assms local.inc_def p_pow_car by auto + have A: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n)) \ nonzero Z\<^sub>p" + using assms p_pow_nonzero p_int_pow_nonzero + by auto + then show ?thesis + using A frac_inv inc_def E1 + by (simp add: Q\<^sub>p_def Zp.nonzero_one_closed) +qed + +lemma p_natpow_closed[simp]: + fixes n::nat + shows "(\[^]n) \ (carrier Q\<^sub>p)" + "(\[^]n) \ (nonzero Q\<^sub>p)" + apply blast + using Qp_nat_pow_nonzero p_nonzero by blast + +lemma nonzero_int_pow_distrib: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "(a \ b) [^](k::int) = a[^]k \ b[^]k" +proof(induction k) + case (nonneg n) + then show ?case using pow_nat[of n _ Q\<^sub>p] + by (smt Qp.nat_pow_distrib Qp.nonzero_closed assms(1) assms(2) int_pow_int) +next + case N: (neg n) + have "a \ b \ Units Q\<^sub>p" + using assms Units_eq_nonzero by blast + hence "(a \ b) [^] - int (Suc n) = inv ((a \ b) [^] (Suc n))" + by (metis Qp.int_pow_inv' int_pow_int) + then show ?case using + Qp.int_pow_inv' Qp.int_pow_unit_closed Qp.inv_of_prod[of "a[^]Suc n" "b[^]Suc n"] + Qp.nat_pow_distrib Qp.nonzero_closed Units_eq_nonzero assms(1) assms(2) int_pow_int by metis +qed + +lemma val_ring_subring: +"subring \\<^sub>p Q\<^sub>p" + using Q\<^sub>p_def \_def inc_im_is_subring by blast + +lemma val_ring_closed: +"\\<^sub>p \ carrier Q\<^sub>p" + by (simp add: subringE(1) val_ring_subring) + +lemma p_pow_diff: + fixes n::int + fixes m::int + assumes "n \0" + assumes "m \0" + shows "\ [^] (n - m) = frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) (\

[^]\<^bsub>Z\<^sub>p\<^esub>m)" +proof- + have 0: "comm_monoid Q\<^sub>p" + by (simp add: Qp.comm_monoid_axioms) + have 1: "\ \ Units Q\<^sub>p" + using Units_eq_nonzero p_nonzero + by blast + have 2: "\ [^] (n - m) = (\[^]n) \ (\ [^] -m)" + by (metis "1" Qp.int_pow_add diff_conv_add_uminus) + have 3: "\ [^] (n - m) = (\[^]n) \ inv\<^bsub>Q\<^sub>p\<^esub>(\ [^] m)" + by (simp add: "1" "2" Qp.int_pow_inv') + then show ?thesis using assms + using fract_frac p_int_pow_nonzero p_intpow_inc p_pow_car by presburger +qed + +lemma Qp_int_pow_add: + fixes n::int + fixes m::int + assumes "a \ nonzero Q\<^sub>p" + shows "a [^] (n + m) = (a [^] n) \ (a [^] m)" + using monoid.int_pow_add[of Q\<^sub>p a n m] Units_eq_nonzero assms + by (simp add: Qp.monoid_axioms) + +lemma Qp_nat_pow_pow: + fixes n::nat + fixes m::nat + assumes "a \ carrier Q\<^sub>p" + shows "(a[^](n*m)) = ((a[^]n)[^]m)" + by (simp add: Qp.nat_pow_pow assms) + +lemma Qp_p_nat_pow_pow: + fixes n::nat + fixes m::nat + shows "(\ [^] (n*m)) = ((\[^]n)[^]m)" + using Qp_nat_pow_pow + by simp + +lemma Qp_units_int_pow: + fixes n::int + assumes "a \ nonzero Q\<^sub>p" + shows "a[^]n = a[^]\<^bsub>units_of Q\<^sub>p\<^esub>n" + apply(cases "n \ 0") + using monoid.units_of_pow[of Q\<^sub>p] + apply (metis int_pow_def2 mult_of_is_Units nat_pow_mult_of not_le) + by (metis Qp.Units_pow_closed Qp.units_of_inv Units_eq_nonzero assms int_pow_def2 mult_of_is_Units nat_pow_mult_of) + +lemma Qp_int_pow_pow: + fixes n::int + fixes m::int + assumes "a \ nonzero Q\<^sub>p" + shows "(a[^](n*m)) = ((a[^]n)[^]m)" +proof- + have 0: "a \ carrier (units_of Q\<^sub>p)" + by (simp add: Units_eq_nonzero assms units_of_carrier) + have "group (units_of Q\<^sub>p)" + using monoid.units_group Qp.units_group + by blast + then show ?thesis + using 0 group.int_pow_pow[of "units_of Q\<^sub>p"] Qp_int_pow_nonzero Qp_units_int_pow assms + by auto +qed + +lemma Qp_p_int_pow_pow: + fixes n::int + fixes m::int + shows "(\ [^] (n*m)) = ((\[^]n)[^]m)" + using Qp_int_pow_pow p_nonzero by blast + +lemma Qp_int_nat_pow_pow: + fixes n::int + fixes m::nat + assumes "a \ nonzero Q\<^sub>p" + shows "(a[^](n*m)) = ((a[^]n)[^]m)" + by (simp add: Qp_int_pow_pow assms int_pow_int) + +lemma Qp_p_int_nat_pow_pow: + fixes n::int + fixes m::nat + shows "(\ [^] (n*m)) = ((\[^]n)[^]m)" + by (simp add: Qp_int_nat_pow_pow p_nonzero) + +lemma Qp_nat_int_pow_pow: + fixes n::nat + fixes m::int + assumes "a \ nonzero Q\<^sub>p" + shows "(a[^](n*m)) = ((a[^]n)[^]m)" + by (simp add: Qp_int_pow_pow assms int_pow_int) + +lemma Qp_p_nat_int_pow_pow: + fixes n::nat + fixes m::int + shows "(\ [^] (n*m)) = ((\[^]n)[^]m)" + by (simp add: Qp_nat_int_pow_pow p_nonzero) + +lemma p_intpow_closed: + fixes n::int + shows "(\[^]n) \ (carrier Q\<^sub>p)" + "(\[^]n) \ (nonzero Q\<^sub>p)" + apply (simp add: Qp.nonzero_closed Qp_int_pow_nonzero p_nonzero) + by (simp add: Qp_int_pow_nonzero p_nonzero) + +lemma p_intpow_add: + fixes n::int + fixes m::int + shows "\ [^] (n + m) = (\ [^] n) \ (\ [^] m)" + using Qp_int_pow_add p_nonzero by blast + +lemma p_intpow_inv: + fixes n::int + shows "(\ [^] n) \ (\ [^] -n) = \" + using Units_eq_nonzero monoid.int_pow_inv'[of Q\<^sub>p \ n] + by (metis add.right_inverse int_pow_0 p_intpow_add) + +lemma p_intpow_inv': + fixes n::int + shows "(\ [^] -n) \ (\ [^] n) = \" + using p_intpow_inv + by (metis add.commute p_intpow_add) + +lemma p_intpow_inv'': + fixes n::int + shows "(\ [^] -n) = inv\<^bsub>Q\<^sub>p\<^esub> (\ [^] n)" + by (simp add: Qp.int_pow_inv' Units_eq_nonzero p_nonzero) + +lemma p_int_pow_factor_int_pow: + assumes "a \ nonzero Q\<^sub>p" + shows "(\[^](n::int) \ a)[^](k::int) = \[^](n*k) \ a[^]k" + using assms nonzero_int_pow_distrib p_intpow_closed(2) Qp_p_int_pow_pow by presburger + +lemma p_nat_pow_factor_int_pow: + assumes "a \ nonzero Q\<^sub>p" + shows "(\[^](n::nat) \ a)[^](k::int) = \[^](n*k) \ a[^]k" + using assms Qp_p_int_nat_pow_pow p_natpow_closed(1) + by (metis int_pow_int p_int_pow_factor_int_pow) + +lemma p_pow_factor: +"\[^]((int N)*l + k) = (\[^]l)[^](N::nat) \ \[^] k" + by (metis Qp_p_int_nat_pow_pow mult_of_nat_commute p_intpow_add) + +lemma p_nat_pow_factor_nat_pow: + assumes "a \ carrier Q\<^sub>p" + shows "(\[^](n::nat) \ a)[^](k::nat) = \[^](n*k) \ a[^]k" + using Qp.nat_pow_distrib Qp_p_nat_pow_pow assms p_natpow_closed(1) by presburger + +lemma p_int_pow_factor_nat_pow: + assumes "a \ carrier Q\<^sub>p" + shows "(\[^](n::int) \ a)[^](k::nat) = \[^](n*k) \ a[^]k" + using assms Qp.nat_pow_distrib Qp_p_int_nat_pow_pow p_intpow_closed(1) by presburger + +lemma(in ring) r_minus_distr: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c \ carrier R" + shows "a \ b \ a \ c = a \ (b \ c)" + using assms + unfolding a_minus_def + by (simp add: r_distr r_minus) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\The Valuation on $\mathbb{Q}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Extending the Valuation from $\mathbb{Z}_p$ to $\mathbb{Q}_p$\ + (********************************************************************************************) + (********************************************************************************************) + +text\ + The valuation of a $p$-adic number can be defined as the difference of the valuations of an + arbitrary choice of numerator and denominator.\ +definition ord where +"ord x = (ord_Zp (numer x)) - (ord_Zp (denom x))" + +definition val where +"val x = (if x = \ then (\::eint) else eint (ord x))" + +lemma val_ord[simp]: + assumes "a \ nonzero Q\<^sub>p" + shows "val a = ord a" + using assms nonzero_def val_def by force + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Properties of the Valuation\ + (********************************************************************************************) + (********************************************************************************************) + +lemma ord_of_frac: + assumes "a \ nonzero Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "ord (frac a b) = (ord_Zp a) - (ord_Zp b)" +proof- + have "frac a b = frac (numer (frac a b)) (denom (frac a b))" + by (simp add: assms(1) assms(2)) + then have "a \\<^bsub>Z\<^sub>p\<^esub> (denom (frac a b)) = b \\<^bsub>Z\<^sub>p\<^esub> (numer (frac a b))" + by (simp add: assms(1) assms(2) numer_denom_swap) + then have "(ord_Zp a) - (ord_Zp b) = (ord_Zp (numer (frac a b))) - (ord_Zp (denom (frac a b)))" + using ord_Zp_eq_frac Q\<^sub>p_def Z\<^sub>p_def + by (metis Zp.frac_closed Zp.nonzero_closed Zp.numer_denom_facts(4) assms(1) assms(2) local.numer_denom_facts(1) local.numer_denom_facts(2) nonzero_fraction ord_of_nonzero(2) ord_pos) + then show ?thesis + using ord_def + by presburger +qed + +lemma val_zero: +"val \ = \" by (simp add: val_def) + +lemma ord_one[simp]: +"ord \ = 0" + using Zp.nonzero_one_closed local.frac_one ord_of_frac by fastforce + +lemma val_one[simp]: +"val (\) = 0" + using ord_one + by (simp add: Qp.one_nonzero zero_eint_def) + +lemma val_of_frac: + assumes "a \ carrier Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + shows "val (frac a b) = (val_Zp a) - (val_Zp b)" +proof(cases "a = \\<^bsub>Z\<^sub>p\<^esub>") + case True + then show ?thesis + using assms(1) assms(2) local.val_zero + by (simp add: Q\<^sub>p_def val_Zp_def) +next + case False + then have "a \ nonzero Z\<^sub>p" + by (simp add: assms(1) nonzero_def) + then show ?thesis + using ord_of_frac[of a b] assms(2) val_def val_ord_Zp + nonzero_numer_imp_nonzero_fraction + by (simp add: Zp.nonzero_memE(2)) +qed + +lemma Z\<^sub>p_division_Qp_0[simp]: + assumes "u \ Units Z\<^sub>p" + assumes "v \ Units Z\<^sub>p" + shows "frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub>= frac u v" +proof- + have 0: "frac v v = \" + using frac_one + by (simp add: Q\<^sub>p_def Zp.Units_nonzero assms(2)) + have 1:"(inv\<^bsub>Z\<^sub>p\<^esub> v) \ carrier Z\<^sub>p" + using assms by blast + have 2:"frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> \ carrier Q\<^sub>p" + by (simp add: "1" Zp.Units_closed Zp.nonzero_one_closed assms(1) local.frac_closed) + have 3: "frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> = (frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub>) \ frac v v" + by (simp add: "0" "2") + then have 4: "frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> = (frac ((u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> v) v)" + by (simp add: Zp.Units_nonzero Zp.nonzero_closed assms(1) assms(2) frac_eqI') + then have 4: "frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> = (frac (u \\<^bsub>Z\<^sub>p\<^esub> ((inv\<^bsub>Z\<^sub>p\<^esub> v) \\<^bsub>Z\<^sub>p\<^esub> v)) v)" + by (simp add: mult_assoc) + have 5:"(inv\<^bsub>Z\<^sub>p\<^esub> v) \\<^bsub>Z\<^sub>p\<^esub> v = \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: assms(2)) + then show "frac (u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \\<^bsub>Z\<^sub>p\<^esub> = (frac u v)" + by (simp add: "4" Zp.Units_closed assms(1)) +qed + +lemma Z\<^sub>p_division_Qp_1: + assumes "u \ Units Z\<^sub>p" + assumes "v \ Units Z\<^sub>p" + obtains w where "w \ Units Z\<^sub>p" + "\ w = frac u v" +proof- + have " (inv\<^bsub>Z\<^sub>p\<^esub> v) \ Units Z\<^sub>p" + by (simp add: assms(2)) + then have "(u \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> v)) \ Units Z\<^sub>p" + using assms + by blast + then show ?thesis + using Z\<^sub>p_division_Qp_0 Zp.Units_closed assms(1) assms(2) + local.inc_def that by presburger +qed + +lemma val_ring_ord_criterion: + assumes "a \ carrier Q\<^sub>p" + assumes "a \ \" + assumes "ord a \ 0" + shows "a \ \\<^sub>p" +proof- + obtain c d where P0: "a = frac c d" and P1: "c \ nonzero Z\<^sub>p" and P2: "d \ nonzero Z\<^sub>p" + by (metis assms(1) assms(2) get_common_denominator nonzero_fraction_imp_nonzero_numer) + obtain m n where P3: "m = ord_Zp c" and P4:"n = ord_Zp d" + by metis + obtain u where "u = ac_Zp c" + by simp + then have P5:"c = (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) \\<^bsub>Z\<^sub>p\<^esub> u" and P6:"u \ Units Z\<^sub>p" + apply (simp add: P1 P3 \u = ac_Zp c\ ac_Zp_factors') + using P1 Zp.nonzero_memE + by (simp add: \u = ac_Zp c\ ac_Zp_is_Unit) + obtain v where "v = ac_Zp d" by simp + have P7:"d = (\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \\<^bsub>Z\<^sub>p\<^esub> v" and P8:"v \ Units Z\<^sub>p" + using P2 P4 \v = ac_Zp d\ ac_Zp_factors' apply blast + using P2 Zp.nonzero_memE + by (simp add: \v = ac_Zp d\ ac_Zp_is_Unit) + have P9: "a = frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>m) \\<^bsub>Z\<^sub>p\<^esub> u) ((\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \\<^bsub>Z\<^sub>p\<^esub> v)" + by (simp add: P0 P5 P7) + have P10: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>m) \ carrier Z\<^sub>p" + using P1 P3 Z\<^sub>p_def ord_pos Zp.nonzero_closed Zp.nonzero_memE(2) p_pow_car + by auto + have P11: "(\

[^]\<^bsub>Z\<^sub>p\<^esub>n) \ nonzero Z\<^sub>p" + by (simp add: P2 P4 Zp.nonzero_closed Zp.nonzero_memE(2) ord_pos p_int_pow_nonzero) + have P12: "u \ carrier Z\<^sub>p" + using P6 Units_def + by blast + have P13: "v \ nonzero Z\<^sub>p" + using P8 Units_def ord_of_nonzero(2) + by (simp add: Zp.Units_nonzero) + have P14: "a = (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>m) (\

[^]\<^bsub>Z\<^sub>p\<^esub>n)) \ (frac u v)" + using P12 P13 P10 P9 P11 Q\<^sub>p_def frac_mult by metis + have P15: "m \ n" + proof- + have "ord_Zp c \ ord_Zp d" + using P0 P1 P2 assms(3) ord_of_frac[of c d] + by (metis P3 P4 antisym eq_iff eq_iff_diff_eq_0 le_cases le_iff_diff_le_0 ord_Zp_def) + then show ?thesis + using P3 P4 by blast + qed + have P16: "n \ 0" + by (simp add: P2 P4 Zp.nonzero_closed Zp.nonzero_memE(2) ord_pos) + have P17: "a = (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) \ (frac u v)" + using P14 P15 P16 local.inc_def[of "(\

[^]\<^bsub>Z\<^sub>p\<^esub> (n - m))"] pow_p_frac_0[of n m] + by (simp add: local.inc_def p_pow_car) + obtain w where P18: "w \ Units Z\<^sub>p" and P19: "\ w = frac u v " + using Z\<^sub>p_division_Qp_1 P6 P8 by blast + have P20: "w \ carrier Z\<^sub>p" + using P18 Units_def by blast + have P21: "a = \ (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \ \ w" + using P15 P17 P19 \_def inc_equation p_pow_car by auto + have P22: "a = (frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub>) \ (frac w \\<^bsub>Z\<^sub>p\<^esub>)" + using P17 P19 P20 local.inc_def by auto + have P23: "\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n) \ carrier Z\<^sub>p" + by (simp add: P15 p_pow_car) + have P24: "a = (frac ((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub> w) \\<^bsub>Z\<^sub>p\<^esub>)" + using P20 P22 P23 frac_mult + by (simp add: Zp.nonzero_one_closed) + have P24: "a = \((\

[^]\<^bsub>Z\<^sub>p\<^esub>(m-n)) \\<^bsub>Z\<^sub>p\<^esub> w)" + by (simp add: P20 P23 P24 cring.cring_simprules(5) domain.axioms(1) local.inc_def) + then show ?thesis + using P20 P23 by blast +qed + +lemma val_ring_val_criterion: + assumes "a \ carrier Q\<^sub>p" + assumes "val a \ 0" + shows "a \ \\<^sub>p" + apply(cases "a = \") + using Qp.int_inc_zero Q\<^sub>p_def inc_of_int apply blast + using assms unfolding val_def + by (simp add: val_ring_ord_criterion zero_eint_def) + +lemma ord_of_inv: + assumes "a \ carrier Q\<^sub>p" + assumes "a \ \" + shows "ord (inv\<^bsub>Q\<^sub>p\<^esub> a) = - (ord a)" +proof- + obtain b c where + Frac: "a = frac b c" and + Car: "b \ carrier Z\<^sub>p" and + Nz_c: "c \ nonzero Z\<^sub>p" + using assms(1) local.numer_denom_facts(1) local.numer_denom_facts(2) + local.numer_denom_facts(5) by blast + have Nz_b: "b \ nonzero Z\<^sub>p" + using Frac Car Nz_c assms(2) nonzero_fraction_imp_nonzero_numer by metis + then have "(inv\<^bsub>Q\<^sub>p\<^esub> a) = frac c b" + using Frac Nz_c frac_inv Q\<^sub>p_def + by auto + then show ?thesis using Frac Nz_b Nz_c ord_of_frac[of b c] ord_of_frac[of c b] + by (simp add: nonzero_def ord_Zp_def) +qed + +lemma val_of_inv: + assumes "a \ carrier Q\<^sub>p" + assumes "a \ \" + shows "val (inv\<^bsub>Q\<^sub>p\<^esub> a) = - (val a)" + using ord_of_inv unfolding uminus_eint_def + by (simp add: assms(1) assms(2) inv_in_frac(2) val_def) + +text\Zp is a valuation ring in Qp\ + +lemma Z\<^sub>p_mem: + assumes "a \ carrier Q\<^sub>p" + shows "a \ \\<^sub>p \ (inv\<^bsub>Q\<^sub>p\<^esub> a \ \\<^sub>p)" +proof(cases "inv\<^bsub>Q\<^sub>p\<^esub>a \ \\<^sub>p \ a = \") + case True + then show ?thesis + using val_ring_subring subringE(2) by auto +next + case False + then have Nz: "a \ \" by auto + have "\ (ord a < 0)" + proof + assume "ord a < 0" + then have "ord (inv\<^bsub>Q\<^sub>p\<^esub> a) >0" + by (simp add: assms(1) Nz ord_of_inv) + then have 0: "ord (inv\<^bsub>Q\<^sub>p\<^esub> a) \0" + by auto + have 1: "(inv\<^bsub>Q\<^sub>p\<^esub> a) \ carrier Q\<^sub>p" + by (simp add: assms(1) Nz inv_in_frac) + have 2: "(inv\<^bsub>Q\<^sub>p\<^esub> a) \\" + by (simp add: assms(1) Nz inv_in_frac(2)) + then have "(inv\<^bsub>Q\<^sub>p\<^esub> a) \ \\<^sub>p" + using val_ring_ord_criterion by (simp add: "0" "1") + then show False + using False by blast + qed + then show ?thesis + using val_ring_ord_criterion assms(1) Nz by auto +qed + +lemma Qp_val_ringI: + assumes "a \ carrier Q\<^sub>p" + assumes "val a \ 0" + shows "a \ \\<^sub>p" +using assms val_ring_val_criterion by blast + + +text\Criterion for determining when an element in Qp is zero\ +lemma val_nonzero: + assumes "a \ carrier Q\<^sub>p" + assumes "s > val a" + shows "a \ nonzero Q\<^sub>p" +proof- + have "val a \ \" + by (metis assms(2) eint_ord_simps(6)) + then show ?thesis + using assms + by (metis (mono_tags, opaque_lifting) local.val_zero not_nonzero_Qp) +qed + +lemma val_ineq: + assumes "a \ carrier Q\<^sub>p" + assumes "val \ \ val a" + shows "a = \" + using assms unfolding val_def + by (metis (mono_tags, lifting) eint_ord_code(5)) + +lemma ord_minus: + assumes "a \ nonzero Q\<^sub>p" + shows "ord a = ord (\a)" +proof- + have "\ a = \ (frac (numer a) (denom a))" + using assms Qp.nonzero_closed local.numer_denom_facts(5) by auto + then have "\ a = (frac (\\<^bsub>Z\<^sub>p\<^esub> (numer a)) (denom a))" + by (simp add: Qp.nonzero_closed assms local.frac_uminus local.numer_denom_facts(1) local.numer_denom_facts(2)) + then show ?thesis + by (metis Q\<^sub>p_def Qp.add.inv_eq_1_iff Qp.nonzero_closed Zp.add.inv_closed assms + local.numer_denom_facts(1) local.numer_denom_facts(2) nonzero_fraction_imp_nonzero_numer + numer_nonzero ord_Zp_of_a_inv ord_def ord_of_frac) +qed + +lemma val_minus: + assumes "a \ carrier Q\<^sub>p" + shows "val a = val (\a)" +proof(cases "a = \") + case True + then show ?thesis + using Qp.minus_zero by presburger +next + case False + then show ?thesis using Qp.domain_axioms assms cring.cring_simprules(21) + cring.cring_simprules(22) domain.axioms(1) not_nonzero_Qp + ord_minus val_def + by metis +qed + +text\The valuation is multiplicative:\ + +lemma ord_mult: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + shows "(ord (x \ y)) = (ord x) + (ord y)" +proof- + have 0:"x \ carrier Q\<^sub>p" using assms by(simp add:nonzero_def) + have 1:"y \carrier Q\<^sub>p" using assms by(simp add:nonzero_def) + obtain a b c where + A: "a \ carrier Z\<^sub>p" and + B: "b \ carrier Z\<^sub>p" and + C: "c \ nonzero Z\<^sub>p" and + Fx: "x = frac a c" and + Fy: "y = frac b c" + using get_common_denominator 0 1 by blast + have An: "a \ nonzero Z\<^sub>p" + using A C Fx assms(1) nonzero_def nonzero_fraction_imp_nonzero_numer + Qp.nonzero_memE(2) by auto + have Bn: " b \ nonzero Z\<^sub>p" + using B C Fy assms(2) nonzero_def nonzero_fraction_imp_nonzero_numer + Qp.nonzero_memE(2) by auto + have Fxy: "x \ y = frac (a \\<^bsub>Z\<^sub>p\<^esub> b) (c \\<^bsub>Z\<^sub>p\<^esub> c)" + by (simp add: A B C Fx Fy frac_mult) + have Cn: "c \\<^bsub>Z\<^sub>p\<^esub> c \ nonzero Z\<^sub>p" + using C Localization.submonoid.m_closed Zp.domain_axioms domain.nonzero_is_submonoid + by metis + have Ordxy0: "ord (x \ y) = ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> b) - ord_Zp (c \\<^bsub>Z\<^sub>p\<^esub> c)" + by (metis "0" "1" An Bn C Cn Fx Fxy Fy Qp.integral + Zp.nonzero_mult_closed Zp.zero_closed fraction_zero + nonzero_fraction nonzero_fraction_imp_nonzero_numer ord_of_frac) + have Ordxy1: "ord (x \ y) = (ord_Zp a) + (ord_Zp b) - ((ord_Zp c) + (ord_Zp c))" + using An Bn C + by (simp add: Ordxy0 ord_Zp_mult) + show ?thesis + proof- + have "ord x + ord y = (ord_Zp a) - (ord_Zp c) + ((ord_Zp b) - (ord_Zp c))" + using An Bn C Fx Fy ord_of_frac[of a c] ord_of_frac[of b c] by presburger + then show ?thesis + using Ordxy1 + by presburger + qed +qed + +lemma val_mult0: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + shows "(val (x \ y)) = (val x) + (val y)" +proof- + have 0: "val x = ord x" + using assms(1) val_ord by metis + have 1: "val y = ord y" + using assms(2) val_ord by metis + have "x \ y \ \" + using field_axioms assms(1) assms(2) integral Qp.integral_iff + Qp.nonzero_closed Qp.nonzero_memE(2) by presburger + then have 2: "val (x \ y) = ord (x \ y)" + by (simp add: val_def) + have 3: "val x + val y = ord x + ord y " + by (simp add: "0" "1") + have 4: "val (x \ y) = ord (x \ y)" + using "2" by auto + then show ?thesis using 3 4 ord_mult assms nonzero_def + by (simp add: nonzero_def) +qed + +text\val is multiplicative everywhere\ + +lemma val_mult: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + shows "(val (x \ y)) = (val x) + (val y)" + apply(cases "x = \ \ y = \") + using assms local.val_zero apply auto[1] + by (meson assms(1) assms(2) not_nonzero_Qp val_mult0) + +text\val and ord are compatible with inclusion\ + +lemma ord_of_inc: +assumes "x \ nonzero Z\<^sub>p" +shows "ord_Zp x = ord(\ x)" +proof- + have "\ x = frac x \\<^bsub>Z\<^sub>p\<^esub>" + using assms(1) + by (simp add: Zp.nonzero_closed local.inc_def) + then have "ord ( \ x) = ord_Zp x - ord_Zp \\<^bsub>Z\<^sub>p\<^esub>" + using assms(1) ord_of_frac + by (simp add: Zp.nonzero_one_closed) + then show ?thesis + using ord_Zp_one + by (simp add: ord_Zp_def) +qed + +lemma val_of_inc: +assumes "x \ carrier Z\<^sub>p" +shows "val_Zp x = val (\ x)" +proof(cases "x \ nonzero Z\<^sub>p") + case True + then show ?thesis + using inc_of_nonzero nonzero_def ord_Zp_def ord_of_inc val_Zp_def val_ord + by (simp add: nonzero_def) +next + case False + then show ?thesis + by (metis Zp.nonzero_memI Zp.nonzero_one_closed assms local.inc_def nonzero_fraction_imp_numer_not_zero val_Zp_def val_def) +qed + +lemma Qp_inc_id: + assumes "a \ nonzero Q\<^sub>p" + assumes "ord a \0" + obtains b where "b \ nonzero Z\<^sub>p" and "a = \ b" + using assms + by (metis (no_types, opaque_lifting) Qp.nonzero_closed Qp.nonzero_memE(2) + Zp.nonzero_one_closed Zp.zero_closed Zp_defs(2) val_ring_ord_criterion image_iff local.inc_def + nonzero_fraction_imp_numer_not_zero not_nonzero_Zp that) + +lemma val_ring_memI: + assumes "a \ carrier Q\<^sub>p" + assumes "val a \ 0" + shows "a \ \\<^sub>p" + using assms Qp_val_ringI by blast + +lemma val_ring_memE: + assumes "a \ \\<^sub>p" + shows "val a \ 0" "a \ carrier Q\<^sub>p" + using assms val_of_inc val_pos apply auto[1] + using assms inc_closed by auto + +lemma val_ring_add_closed: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "a \ b \ \\<^sub>p" + using val_ring_subring subringE(7) by (metis assms(1) assms(2)) + +lemma val_ring_times_closed: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "a \ b \ \\<^sub>p" + using val_ring_subring subringE(6) by (metis assms(1) assms(2)) + +lemma val_ring_ainv_closed: + assumes "a \ \\<^sub>p" + shows "\ a \ \\<^sub>p" + using val_ring_subring subringE(5) by (metis assms) + +lemma val_ring_minus_closed: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "a \ b \ \\<^sub>p" + using assms val_ring_subring val_ring_ainv_closed val_ring_add_closed + unfolding a_minus_def by blast + +lemma one_in_val_ring: +"\ \ \\<^sub>p" + apply(rule val_ring_memI) + apply blast + unfolding val_one by blast + +lemma zero_in_val_ring: +"\ \ \\<^sub>p" + apply(rule val_ring_memI) + apply blast + unfolding val_zero + by simp + +lemma ord_p: +"ord \ = 1" + using p_nonzero ord_Zp_p ord_of_inc p_inc + by (smt Zp_int_inc_closed ord_of_nonzero(2)) + +lemma ord_p_pow_nat: +"ord (\ [^] (n::nat)) = n" + using p_pow_nonzero ord_Zp_p ord_of_inc p_inc ord_Zp_p_pow p_natpow_inc p_pow_nonzero' + by auto + +lemma ord_p_pow_int: +"ord (\ [^] (n::int)) = n" +proof(cases "n \0") + case True + then show ?thesis + by (metis int_nat_eq int_pow_int ord_p_pow_nat) +next + case False + then have Neg: "n <0" by auto + then have 0: "\[^]n = frac \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n))" + using p_intpow by auto + have "(\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n)) \ nonzero Z\<^sub>p" + using False p_int_pow_nonzero + by (simp add: nonzero_def) + then have "ord (\ [^] (n::int)) = ord_Zp \\<^bsub>Z\<^sub>p\<^esub> - ord_Zp (\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n))" + using "0" ord_of_frac + by (simp add: Zp.nonzero_one_closed) + then have "ord (\ [^] (n::int)) = - ord_Zp (\

[^]\<^bsub>Z\<^sub>p\<^esub>(-n))" + using ord_Zp_one by linarith + then have "ord (\ [^] (n::int)) = -(-n)" + using Neg ord_Zp_p_int_pow + by (metis int.lless_eq neg_0_le_iff_le) + then show ?thesis + by auto +qed + +lemma ord_nonneg: + assumes "x \ \\<^sub>p" + assumes "x \ \" + shows "ord x \0" +proof- + obtain a where A: "x = \ a \ a \ carrier Z\<^sub>p" + using assms(1) by blast + then have "a \ nonzero Z\<^sub>p" using assms(2) + local.inc_def nonzero_fraction_imp_numer_not_zero not_nonzero_Zp + using Zp.nonzero_one_closed by blast + then have "ord_Zp a \0" + using A + by (simp add: Zp.nonzero_memE(2) ord_pos) + then show ?thesis + using A \a \ nonzero Z\<^sub>p\ ord_of_inc by metis +qed + +lemma val_p: +"val \ = 1" + using p_inc val_Zp_p val_of_inc val_def ord_p p_nonzero val_ord + by simp + +lemma val_p_int_pow: +"val (\[^](k::int)) = k" + using ord_p_pow_int p_intpow_closed(2) val_ord by presburger + +lemma val_p_int_pow_neg: +"val (\[^](-k::int)) = - eint k" + by (metis eint.distinct(2) local.val_zero p_intpow_closed(1) p_intpow_inv'' val_of_inv val_p_int_pow) + +lemma nonzero_nat_pow_ord: + assumes "a \ nonzero Q\<^sub>p" + shows "ord (a [^] (n::nat)) = n * ord a" + apply(induction n) + apply simp + using Qp_nat_pow_nonzero assms semiring_normalization_rules(2) + by (simp add: semiring_normalization_rules(2) ord_mult) + + +lemma add_cancel_eint_geq: + assumes "(eint a) + x \ (eint a ) + y" + shows "x \y" + using assms eint_add_left_cancel_le by blast + +lemma(in padic_fields) prod_equal_val_imp_equal_val: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val (a \ b) = val (a \ c)" + shows "val b = val c" +proof(cases "b = \") + case True + then have "val (a \ b) = \" + using Qp.nonzero_closed Qp.r_null assms(1) local.val_zero by presburger + then have "c = \" + using assms True + by (metis Qp.integral Qp.nonzero_closed Qp.nonzero_mult_in_car Qp.not_nonzero_memI eint_ord_simps(3) not_nonzero_Qp val_ineq) + then show ?thesis + using True by blast +next + case False + obtain n where n_def: "val a = eint n" + using assms val_ord by blast + then show ?thesis using val_mult[of a b] val_mult[of a c] unfolding n_def + by (simp add: Qp.nonzero_closed assms(1) assms(2) assms(3) assms(4)) +qed + + +lemma two_times_eint: + shows "2*(x::eint) = x + x" + by (metis eint_2_minus_1_mult eint_add_cancel_fact plus_eq_infty_iff_eint times_eint_simps(4)) + +lemma times_cfs_val_mono: + assumes "u \ Units Q\<^sub>p" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val (u \ a) \ val (u \ b)" + shows "val a \ val b" +proof- + have "eint (ord u) + val a \ eint (ord u) + val b" + using assms val_ord val_mult Qp.Units_nonzero Qp.nonzero_closed Units_eq_nonzero by metis + thus ?thesis + apply(induction "val a") + using add_cancel_eint_geq apply blast + using add_cancel_eint_geq by blast +qed + +lemma times_cfs_val_mono': + assumes "u \ Units Q\<^sub>p" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val (u \ a) \ val (u \ b) + \" + shows "val a \ val b + \" +proof- + obtain c where c_def: "c \ carrier Q\<^sub>p \ val c = \" + by (metis Qp.zero_closed eint.exhaust local.val_zero p_intpow_closed(1) val_p_int_pow) + have 0: "val (u \ a) \ val (u \ (b \ c))" + using assms val_mult[of "u \ b" c] Qp.m_assoc[of u b c] c_def Qp.Units_closed Qp.cring_simprules(5) + by metis + have 1: "val a \ val (b \ c)" + apply(rule times_cfs_val_mono[of u]) + using assms apply blast using assms apply blast + apply(rule Qp.ring_simprules) using assms apply blast using c_def apply blast + by(rule 0) + show ?thesis + using 1 val_mult[of b c] assms c_def by smt +qed + +lemma times_cfs_val_mono'': + assumes "u \ Units Q\<^sub>p" + assumes "a \ carrier Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val a \ val b + \" + shows "val (u \ a) \ val (u \ b) + \" + apply(rule times_cfs_val_mono'[of "inv u" "u \ a" "u \ b" \]) + using assms apply blast + apply(rule Qp.ring_simprules) + using assms apply blast using assms apply blast + apply(rule Qp.ring_simprules) + using assms apply blast using assms apply blast + using m_assoc assms + by (metis Qp.Units_closed Qp.cring_simprules(5) Qp.inv_cancelR(1)) + +lemma val_ineq_cancel_leq: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val (a \ b) \ val (a \ c)" + shows "val b \ val c" + using Units_eq_nonzero assms(1) assms(2) assms(3) assms(4) times_cfs_val_mono by blast + +lemma val_ineq_cancel_leq': + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val b \ val c" + shows "val (a \ b) \ val (a \ c)" + apply(rule val_ineq_cancel_leq[of "inv a" "a \ b" "a \ c"]) + using assms(1) nonzero_inverse_Qp apply blast + using Qp.nonzero_closed assms(1) assms(2) apply blast + using Qp.nonzero_closed assms assms(2) apply blast + by (metis Qp.inv_cancelR(1) Qp.m_closed Qp.nonzero_closed Units_eq_nonzero assms(1) assms(2) assms(3) assms(4)) + +lemma val_ineq_cancel_le: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val (a \ b) < val (a \ c)" + shows "val b < val c" + using Qp.nonzero_closed assms(1) assms(2) assms(3) assms(4) val_mult by auto + +lemma val_ineq_cancel_le': + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val b < val c" + shows "val (a \ b) < val (a \ c)" + apply(rule val_ineq_cancel_le[of "inv a" "a \ b" "a \ c"]) + using assms(1) nonzero_inverse_Qp apply blast + using Qp.nonzero_closed assms(1) assms(2) apply blast + using Qp.nonzero_closed assms assms(2) apply blast + by (metis Qp.inv_cancelR(1) Qp.m_closed Qp.nonzero_closed Units_eq_nonzero assms(1) assms(2) assms(3) assms(4)) + +lemma finite_val_imp_nonzero: + assumes "a \ carrier Q\<^sub>p" + assumes "val a \ \" + shows "a \ nonzero Q\<^sub>p" + using assms unfolding val_def nonzero_def + by (metis (mono_tags, lifting) mem_Collect_eq) + +lemma val_ineq_cancel_leq'': + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "val b \ val c + eint N" + shows "val (a \ b) \ val (a \ c) + eint N" +proof- + obtain d where d_def: "d = \[^]N \ c" + by blast + show ?thesis + proof(cases "c = \") + case True + then show ?thesis unfolding True + using True Units_eq_nonzero assms(1) assms(2) assms(4) times_cfs_val_mono'' by blast + next + case False + have F0: "c \ nonzero Q\<^sub>p" + using False assms Qp.not_nonzero_memE by blast + have F1: "val (a \ c) < \" + using F0 assms + by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.nonzero_memE(2) eint.distinct(1) eint_ord_simps(4) val_def) + have F2: "b \ nonzero Q\<^sub>p" + using F0 assms + by (metis False Groups.add_ac(2) add_cancel_eint_geq finite_val_imp_nonzero local.val_zero plus_eint_simps(2) val_ineq) + have F3: "ord b \ ord c + N" + using assms F0 + by (metis F2 eint_ord_simps(1) plus_eint_simps(1) val_ord) + have F4: "ord a + ord b \ ord a + ord c+ N" + using assms F0 F1 F2 ord_mult F3 by presburger + have F5: "ord (a \ b) \ ord (a \ c) + N" + using F4 F2 F0 assms ord_mult by presburger + have F6: "a \ b \ nonzero Q\<^sub>p" + by (metis F2 Qp.integral Qp.nonzero_memE(1) Qp.nonzero_mult_closed Qp.not_nonzero_memE assms(1)) + have F7: "a \ c \ nonzero Q\<^sub>p" + by (metis False Qp.integral Qp.nonzero_closed Qp.nonzero_memI Qp.nonzero_mult_closed assms(1) assms(3)) + show "val (a \ b) \ val (a \ c) + eint N" + using val_ord[of "a \ b" ] val_ord[of "a \c"] F7 F7 F4 + Units_eq_nonzero assms(1) assms(2) assms(3) assms(4) times_cfs_val_mono'' by blast + qed +qed + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\The Ultrametric Inequality on $\mathbb{Q}_p$\ + (********************************************************************************************) + (********************************************************************************************) + + +lemma ord_ultrametric: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "x \ y \ nonzero Q\<^sub>p" + shows "ord (x \ y) \ min (ord x) (ord y)" +proof- + have 0:"x \ carrier Q\<^sub>p" using assms by(simp add:nonzero_def) + have 1:"y \carrier Q\<^sub>p" using assms by(simp add:nonzero_def) + obtain a b c where + A: "a \ carrier Z\<^sub>p" and + B: "b \ carrier Z\<^sub>p" and + C: "c \ nonzero Z\<^sub>p" and + Fx: "x = frac a c" and + Fy: "y = frac b c" + using 0 1 get_common_denominator by blast + have An: "a \ nonzero Z\<^sub>p" + using A C Fx assms(1) nonzero_fraction_imp_nonzero_numer + Qp.nonzero_memE(2) by auto + have Bn: " b \ nonzero Z\<^sub>p" + using B C Fy assms(2) nonzero_fraction_imp_nonzero_numer Qp.nonzero_memE(2) by auto + have Fxy: "x \ y = frac (a \\<^bsub>Z\<^sub>p\<^esub> b) c" using 0 1 + by (simp add: A B C Fx Fy frac_add_common_denom) + have ABn: " a \\<^bsub>Z\<^sub>p\<^esub> b \ nonzero Z\<^sub>p" + proof- + have "a \\<^bsub>Z\<^sub>p\<^esub> b \ carrier Z\<^sub>p" + using A B Z\<^sub>p_def padic_add_closed prime by blast + then show ?thesis + using Fxy C assms(3) Qp.nonzero_memE(2) + nonzero_fraction_imp_nonzero_numer by blast + qed + have Ordx: "ord x = ord_Zp a - ord_Zp c" + using Fx An C ord_of_frac by metis + have Ordy: "ord y = ord_Zp b - ord_Zp c" + using Fy Bn C ord_of_frac by metis + have Ordxy: "ord (x \ y) = ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> b) - ord_Zp c" + using Fxy ABn C ord_of_frac by metis + then show ?thesis + using Ordx Ordy Ordxy ord_Zp_ultrametric[of a b] ABn An Bn + by linarith +qed + +lemma ord_ultrametric': + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "x \\<^bsub>Q\<^sub>p\<^esub> y \ nonzero Q\<^sub>p" + shows "ord (x \\<^bsub>Q\<^sub>p\<^esub> y) \ min (ord x) (ord y)" +proof- + have "ord y = ord (\y)" + using assms(2) ord_minus by blast + then show ?thesis + using assms ord_ultrametric[of x "\y"] + unfolding a_minus_def + using Qp.add.inv_closed Qp.add.inv_eq_1_iff Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_memI + by metis +qed + +lemma val_ultrametric0: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "x \ y \ nonzero Q\<^sub>p" + shows " min (val x) (val y) \ val (x \ y) " +proof- + have 0: "val (x \ y) = ord (x \ y)" + using assms(3) nonzero_def val_def[of "(x \ y)"] by fastforce + have 1: "val x = ord x" + using assms(1) nonzero_def val_def + by (simp add: nonzero_def) + have 2: "val y = ord y" + using assms(2) nonzero_def val_def val_ord by auto + have 3: "ord (x \ y) \ min (ord x) (ord y)" + by (simp add: assms(1) assms(2) assms(3) ord_ultrametric) + then show ?thesis + by (simp add: "0" "1" "2") +qed + +lemma val_ultrametric: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + shows " min (val x) (val y) \ val (x \ y)" + apply(cases "x = \ \ y = \") + using assms(1) assms(2) apply auto[1] + by (metis Qp.add.m_closed assms(1) assms(2) eint_ord_code(3) local.val_zero not_nonzero_Qp val_ultrametric0) + +lemma val_ultrametric': + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + shows " min (val x) (val y) \ val (x \ y)" + using val_ultrametric[of x "\y"] + val_minus[of y] + assms + by (metis Qp.domain_axioms a_minus_def cring.cring_simprules(3) domain.axioms(1)) + +lemma diff_ord_nonzero: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "ord x \ ord y" + shows "x \ y \ nonzero Q\<^sub>p" +proof(rule ccontr) + assume " x \ y \ nonzero Q\<^sub>p" + then have "x \ y = \" + using Qp.add.m_closed Qp.nonzero_closed Qp.nonzero_memI assms(1) assms(2) by blast + then have "x = \ y" + using Qp.minus_equality Qp.nonzero_closed assms(1) assms(2) by blast + then have "ord x = ord y" + using assms(2) ord_minus by presburger + then show False + using assms(3) by blast +qed + +lemma ord_ultrametric_noteq: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "ord x > ord y" + shows "ord (x \ y) = (ord y)" +proof(rule ccontr) + assume "ord (x \ y) \ ord y" + have 00:"x \ y \ nonzero Q\<^sub>p" + proof(rule ccontr) + assume "x \ y \ nonzero Q\<^sub>p" + then have "x \ y = \" + using Qp.add.m_closed Qp.nonzero_closed Qp.nonzero_memI assms(1) assms(2) by blast + then have "x = \ y" + by (smt \x \ y \ nonzero Q\<^sub>p\ assms(1) assms(2) assms(3) diff_ord_nonzero) + then have "ord x = ord y" + using assms(2) ord_minus by presburger + then show False + using assms(3) by linarith + qed + then have 0: "ord (x \ y) > ord y" + using ord_ultrametric[of x y] \ord (x \ y) \ ord y\ assms(1) assms(2) assms(3) + by linarith + have 1: "((y \ x) \ x) = y" + using "00" Qp.add.inv_solve_right' Qp.add.m_comm Qp.minus_eq Qp.nonzero_closed assms(1) assms(2) by presburger + have 2: "ord ((y \ x) \ x) \ min (ord (y \ x)) (ord x) " + using 1 assms ord_ultrametric'[of "(y \ x)" x] diff_ord_nonzero by auto + have 3: "ord y \ min (ord x) (ord (x \ y))" + using 2 1 Q\<^sub>p_def Qp.domain_axioms Z\<^sub>p_def assms(1) assms(2) cring.cring_simprules(10) + domain.axioms(1) Qp.nonzero_closed by fastforce + show False + using 3 0 assms + by linarith +qed + +lemma ord_ultrametric_noteq': + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "ord x > ord y" + shows "ord (x \ y) = (ord y)" + using assms ord_ultrametric_noteq[of x "\y"] + by (metis Qp.add.inv_closed Qp.minus_eq Qp.nonzero_closed Qp.nonzero_memI Qp.r_neg Qp.r_zero ord_minus) + +lemma ord_ultrametric_noteq'': + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + assumes "ord y > ord x" + shows "ord (x \ y) = (ord x)" + using assms ord_ultrametric_noteq'[of y x] + by (metis Qp.minus_a_inv Qp.nonzero_closed Qp.not_eq_diff_nonzero ord_minus) + +lemma val_ultrametric_noteq: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + assumes "val x > val y" + shows "val (x \ y) = val y" + apply(cases "x = \") + apply (simp add: assms(2)) + using assms unfolding val_def + by (smt Qp.not_nonzero_memI diff_ord_nonzero eint_ord_simps(2) not_nonzero_Qp ord_ultrametric_noteq val_def val_nonzero) + +lemma val_ultrametric_noteq': + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + assumes "val x > val y" + shows "val (x \ y) = val y" + using assms val_ultrametric_noteq[of x "\y"] + by (metis Qp.domain_axioms a_minus_def cring.cring_simprules(3) domain.axioms(1) val_minus) + +lemma ultrametric_equal_eq: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + assumes "val (y \ x) > val x" + shows "val x = val y" + using assms + by (metis (no_types, lifting) Qp.add.inv_closed Qp.add.m_assoc Qp.l_neg Qp.minus_closed Qp.minus_eq Qp.r_zero val_ultrametric_noteq) + +lemma ultrametric_equal_eq': + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + assumes "val (x \ y) > val x" + shows "val x = val y" + using assms ultrametric_equal_eq[of x y] + by (metis Qp.minus_a_inv Qp.minus_closed val_minus) + +lemma val_ultrametric_noteq'': + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + assumes "val x > val y" + shows "val (y \ x) = val y" + by (metis Qp.minus_a_inv Qp.minus_closed assms(1) assms(2) assms(3) val_minus val_ultrametric_noteq') + +text\Ultrametric over finite sums:\ + +lemma Min_mono: + assumes "finite A" + assumes "A \ {}" + assumes "\ a. a \ A \ f a \ a" + shows "Min (f`A) \ Min A" + by (meson Min_in Min_le_iff assms(1) assms(2) assms(3) finite_imageI image_eqI image_is_empty) + +lemma Min_mono': + assumes "finite A" + assumes "\ (a::'a). a \ A \ (f::'a \ eint) a \ g a" + shows "Min (f`A) \ Min (g `A)" +proof- + have "(\a \ A. f a \ g a) \ Min (f`A) \ Min (g` A)" + apply(rule finite.induct[of A]) + apply (simp add: assms(1)) + apply simp + proof fix A a assume F: "finite A" "(\a\A. f a \ g a) \ Min (f ` A) \ Min (g ` A)" +"\a\insert a A. f a \ g a" + show "Min (f ` insert a A) \ Min (g ` insert a A)" + proof- + obtain k where k_def: "k \ insert a A \ g k = Min (g ` insert a A)" + using assms + by (metis (mono_tags, opaque_lifting) F(1) Min_eq_iff finite.insertI finite_imageI image_iff image_is_empty insert_not_empty) + then have 0: "f k \ g k" + using F(3) by blast + thus ?thesis using k_def + by (metis F(1) Min_le dual_order.trans finite.insertI finite_imageI image_eqI) + qed + qed + thus ?thesis using assms by blast +qed + +lemma eint_ord_trans: + assumes "(a::eint) \ b" + assumes "b \ c" + shows "a \ c" + using assms by auto + +lemma eint_Min_geq: + assumes "finite (A::eint set)" + assumes "\x. x \ A \ x \ c" + assumes "A \ {}" + shows "Min A \ c" + using Min_in[of A] assms(2)[of "Min A"] assms by blast + +lemma eint_Min_gr: + assumes "finite (A::eint set)" + assumes "\x. x \ A \ x > c" + assumes "A \ {}" + shows "Min A > c" + using Min_in[of A] assms(2)[of "Min A"] assms by blast + +lemma finsum_val_ultrametric: + assumes "g \ A \ carrier Q\<^sub>p" + assumes "finite A" + assumes "A \ {}" + shows " val (finsum Q\<^sub>p g A) \ Min (val ` (g`A))" +proof- + have "A \ {} \ g \ A \ carrier Q\<^sub>p \ val (finsum Q\<^sub>p g A) \ Min (val ` (g`A))" + apply(rule finite.induct[of A]) + apply (simp add: assms(2)) + apply blast + proof fix A a assume A: "finite A" "A \ {}\ g \ A \ carrier Q\<^sub>p \ val (finsum Q\<^sub>p g A) \ Min (val ` g ` A)" + " insert a A \ {} \ g \ insert a A \ carrier Q\<^sub>p " + show "val (finsum Q\<^sub>p g (insert a A)) \ Min (val ` g ` insert a A)" + apply(cases "a \ A") + apply (metis A(2) A(3) insert_absorb) + proof(cases "A = {}") + have g_closed: "g \ A \ carrier Q\<^sub>p" + using A(3) by blast + have g_closed': "g \ insert a A \ carrier Q\<^sub>p" + using A(3) by linarith + then have ga: "g a \ carrier Q\<^sub>p" + by blast + assume a_notin: "a \ A" + case True + have "g a \ carrier Q\<^sub>p" using A(3) + by blast + then have 0: "(finsum Q\<^sub>p g (insert a A)) = g a" + using A True abelian_monoid.finsum_insert[of Q\<^sub>p A a g] abelian_monoid.finsum_empty[of Q\<^sub>p g] + Qp.abelian_monoid_axioms Qp.add.l_cancel_one Qp.zero_closed a_notin g_closed + by metis + have 1: "Min (val ` g ` insert a A) = val (g a)" + by (metis A(1) True Min_in finite.insertI finite_imageI image_empty + image_insert insert_not_empty singletonD) + show ?thesis + using 0 1 + by simp + next + case False + assume a_notin: "a \ A" + have g_closed: "g \ A \ carrier Q\<^sub>p" + using A(3) by blast + have g_closed': "g \ insert a A \ carrier Q\<^sub>p" + using A(3) by linarith + then have ga: "g a \ carrier Q\<^sub>p" + by blast + have 0: "finsum Q\<^sub>p g (insert a A) = g a \ finsum Q\<^sub>p g A" + by (simp add: A(1) a_notin g_closed ga) + have 1: "min (val (g a)) (Min (val ` g ` A)) = Min (insert (val (g a)) (val ` g ` A))" + proof- + have 10: "finite (val ` g ` A) " using A + by blast + then have 11: "val ` g ` A \ {}" + using False + by blast + show ?thesis + using 10 11 Min_insert + by simp + qed + have 2: " val (g a \ finsum Q\<^sub>p g A) \ min (val (g a)) (val (finsum Q\<^sub>p g A))" + using val_ultrametric[of "g a" "finsum Q\<^sub>p g A" ] abelian_monoid.finsum_closed[of Q\<^sub>p g A] + g_closed ga Qp.abelian_monoid_axioms by blast + have 3: " val (finsum Q\<^sub>p g A) \ Min (val ` g ` A)" + using A False + by blast + have 4: " val (finsum Q\<^sub>p g (insert a A)) \ min (val (g a)) (val (finsum Q\<^sub>p g A))" + using 2 "0" + by presburger + have 5: " val (finsum Q\<^sub>p g A) \ Min (val ` g ` A)" + using False "3" by blast + show " val (finsum Q\<^sub>p g (insert a A)) \ Min (val ` g ` insert a A)" + using 5 4 1 + by (smt image_insert min.cobounded1 min_def order_trans) + qed + qed + then show ?thesis + using assms + by blast +qed + +lemma (in padic_fields) finsum_val_ultrametric': + assumes "g \ A \ carrier Q\<^sub>p" + assumes "finite A" + assumes "\i. i \ A \ val (g i) \ c" + shows " val (finsum Q\<^sub>p g A) \ c" +proof(cases "A = {}") + case True + then show ?thesis using assms + unfolding True Qp.finsum_empty + using eint_ord_code(3) local.val_zero by presburger +next + case False + show ?thesis + proof(rule eint_ord_trans[of _ "Min (val ` g ` A)"]) + have 0: "\s. s \ val ` g ` A \ s \ c" + proof- fix s assume A: "s \ val ` g ` A" + then obtain a where a_def: "a \ A \ s = val (g a)" + by blast + have s_eq: "s = val (g a)" + using a_def by blast + show "c \ s" + using assms(3)[of a] A a_def unfolding s_eq by blast + qed + show "c \ Min (val ` g ` A)" + apply(rule eint_Min_geq) + using assms apply blast using assms apply blast + using False by blast + show "Min (val ` g ` A) \ val (finsum Q\<^sub>p g A)" + by(rule finsum_val_ultrametric[of g A], rule assms, rule assms, rule False) + qed +qed + +lemma (in padic_fields) finsum_val_ultrametric'': + assumes "g \ A \ carrier Q\<^sub>p" + assumes "finite A" + assumes "\i. i \ A \ val (g i) > c" + assumes "c < \" + shows " val (finsum Q\<^sub>p g A) > c" +proof(cases "A = {}") + case True + then show ?thesis using assms + unfolding True Qp.finsum_empty val_zero + using eint_ord_code(3) local.val_zero by blast +next + case False + show ?thesis + proof(rule less_le_trans[of _ "Min (val ` g ` A)"]) + have 0: "\s. s \ val ` g ` A \ s > c" + proof- fix s assume A: "s \ val ` g ` A" + then obtain a where a_def: "a \ A \ s = val (g a)" + by blast + have s_eq: "s = val (g a)" + using a_def by blast + show "c < s" + using assms(3)[of a] A a_def unfolding s_eq by blast + qed + show "c < Min (val ` g ` A)" + apply(rule eint_Min_gr) + using assms apply blast using assms + using "0" apply blast + using False by blast + show "Min (val ` g ` A) \ val (finsum Q\<^sub>p g A)" + by(rule finsum_val_ultrametric[of g A], rule assms, rule assms, rule False) + qed +qed + +lemma Qp_diff_diff: + assumes "x \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + shows "(x \ c) \ (d \ c) = x \ d " + by (smt Qp.domain_axioms a_minus_def assms(1) assms(2) assms(3) cring.cring_simprules(10) + cring.cring_simprules(19) cring.cring_simprules(3) cring.cring_simprules(4) + cring.cring_simprules(7) domain.axioms(1)) + +text\This variant of the ultrametric identity formalizes the common saying that "all triangles in $\mathbb{Q}_p$ are isosceles":\ +lemma Qp_isosceles: + assumes "x \ carrier Q\<^sub>p" + assumes "c \ carrier Q\<^sub>p" + assumes "d \ carrier Q\<^sub>p" + assumes "val (x \ c) \ v" + assumes "val (d \ c) \ v" + shows "val (x \ d) \ v" +proof- + have "val (x \ d) \ min (val (x \ c)) (val (d \ c))" + using assms Qp_diff_diff[of x c d] + by (metis Qp.domain_axioms cring.cring_simprules(4) domain.axioms(1) val_ultrametric') + then show ?thesis using assms + by (meson dual_order.trans min_le_iff_disj) +qed + +text\More variants on the ultrametric inequality\ + +lemma MinE: + assumes "finite (A::eint set)" + assumes "a = Min A" + assumes "b \ A" + shows "a \ b" + using assms by (simp add: assms(3)) + +lemma MinE': + assumes "finite (A::eint set)" + assumes "a = Min A" + assumes "b \ A - {a}" + shows "a < b" +proof- + have 0: "a \ b" + using assms MinE by blast + have 1: "a \ b" using assms by blast + show ?thesis using 0 1 + by simp +qed + +lemma MinE'': + assumes "finite A" + assumes "f \ A \ (UNIV :: eint set)" + assumes "a = Min (f ` A)" + assumes "b \ A" + shows "a \ f b" + apply(rule MinE[of "f` A"]) using assms apply blast using assms apply blast using assms by blast + +lemma finsum_val_ultrametric_diff: + assumes "g \ A \ carrier Q\<^sub>p" + assumes "finite A" + assumes "A \ {}" + assumes "\a b. a \ A \ b \ A \ a \ b \ val (g a) \ val (g b)" + shows " val (finsum Q\<^sub>p g A) = Min (val ` g`A)" +proof- + have "Min (val ` g ` A) \ val ` g ` A" + using Min_in[of "val ` g ` A"] assms by blast + then obtain a where a_def: "a \ A \ Min (val ` g`A) = val (g a)" + by blast + have 0: "\b. b \ A \ b \ a \ val (g b) > val (g a)" + apply(rule MinE'[of "val ` g ` A"]) using assms apply blast + using a_def assms apply presburger + using assms(4)[of a] a_def + by (metis image_eqI insert_Diff insert_iff) + have 1: "g a \ finsum Q\<^sub>p g (A - {a}) = finsum Q\<^sub>p g (insert a (A - {a}))" + using assms Qp.finsum_insert[of "A - {a}" a g] a_def finite_subset[of "A - {a}" A] by blast + have 2: "finsum Q\<^sub>p g A = g a \ finsum Q\<^sub>p g (A - {a})" + unfolding 1 using a_def + by (metis insert_Diff_single insert_absorb) + have 3: "g a \ carrier Q\<^sub>p" + using a_def assms by blast + show "val (finsum Q\<^sub>p g A) = Min (val ` g`A)" + proof(cases "A = {a}") + case True + show ?thesis using 3 Qp.finsum_empty[of g] a_def unfolding 2 unfolding True + by auto + next + case False + then have F0: "A - {a} \ {}" + using assms by blast + have F1: "finsum Q\<^sub>p g (A - {a}) \ carrier Q\<^sub>p" + apply(rule Qp.finsum_closed) + using assms by blast + have F2: "val (finsum Q\<^sub>p g (A - {a})) \ Min (val ` g` (A- {a}))" + apply(rule finsum_val_ultrametric) + using assms apply blast using assms apply blast using False a_def by blast + have F3: "Min (val ` g` (A- {a})) \ (val ` g` (A- {a}))" + apply(rule Min_in) + using assms apply blast using False a_def by blast + obtain b where b_def: "b \ A - {a} \ Min (val ` g` (A- {a})) = val (g b)" + using F3 by blast + have F4: "Min (val ` g` (A- {a})) > val (g a)" + using a_def assms b_def by (metis "0" Diff_iff singletonI) + have F5: "val (finsum Q\<^sub>p g (A - {a})) > val (g a)" + using F4 F2 less_le_trans by blast + then show ?thesis using 2 F1 3 + by (metis Qp.a_ac(2) a_def val_ultrametric_noteq) + qed +qed + +lemma finsum_val_ultrametric_diff': + assumes "g \ A \ carrier Q\<^sub>p" + assumes "finite A" + assumes "A \ {}" + assumes "\a b. a \ A \ b \ A \ a \ b \ val (g a) \ val (g b)" + shows " val (finsum Q\<^sub>p g A) = (MIN a \ A. (val (g a)))" +proof- + have 0: "(\ a. val (g a)) ` A = (val ` g`A)" + by blast + show ?thesis using assms finsum_val_ultrametric_diff[of g A] unfolding 0 by blast +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Constructing the Angular Component Maps on $\mathbb{Q}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Unreduced Angular Component Map\ + (********************************************************************************************) + (********************************************************************************************) + +text\While one can compute the residue of a $p$-adic integer mod $p^n$, this operation does not generalize to the $p$-adic field unless we restrict our attention to the valuation ring. However, we can still define the angular component maps on the field $\mathbb{Q}_p$, which allows us to take a sort of residue for any element $x \in \mathbb{Q}_p$. Given a nonzero element $x \in \mathbb{Q}_p^{\times}$, we can normalize it to obtain $p^{-ord( x)}x$ which has of valuation zero, and then computes its residue (viewed as an element of $\mathbb{Z}_p$). The resulting map agrees with the standard residue map on elements of $\mathbb{Q}_p$ of valuation zero, but not on terms of positive or negative valuation. For example, the element $p^2$ has an order $1$ residue of $0$, but its order $1$ angular component is $1$. In the formalism below, we will use the term "\texttt{angular\_component}" to refer to the unreduced normalization map $x \mapsto p^{-ord( x)}x$, and use the notation "\texttt{ac n}" to refer to the angular component which has been reduced mod $p^n$. This is line with the terminology used in \cite{denef1986}. \ +definition angular_component where +"angular_component a = (ac_Zp (numer a)) \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> ac_Zp (denom a))" + +lemma ac_fract: + assumes "c \ carrier Q\<^sub>p" + assumes "a \ nonzero Z\<^sub>p" + assumes "b \ nonzero Z\<^sub>p" + assumes "c = frac a b" + shows "angular_component c = (ac_Zp a)\\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub>(ac_Zp b)" +proof- + have "(numer c) \\<^bsub>Z\<^sub>p\<^esub> b = (denom c) \\<^bsub>Z\<^sub>p\<^esub> a" + by (simp add: assms(2) assms(3) assms(4) mult_comm numer_denom_swap) + then have "ac_Zp ((numer c) \\<^bsub>Z\<^sub>p\<^esub> b) = ac_Zp ((denom c) \\<^bsub>Z\<^sub>p\<^esub> a)" + by simp + then have "(ac_Zp (numer c)) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp b) = (ac_Zp (denom c)) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp a)" + by (metis Q\<^sub>p_def Zp.nonzero_closed Zp.numer_denom_facts(3) ac_Zp_mult assms(2) assms(3) assms(4) local.frac_closed local.numer_denom_facts(1) nonzero_fraction_imp_nonzero_numer nonzero_numer_imp_nonzero_fraction numer_denom_frac) + then have "(inv \<^bsub>Z\<^sub>p\<^esub> (ac_Zp (denom c))) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (numer c)) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp b) = (ac_Zp a)" + using ac_Zp_is_Unit[of "ac_Zp (denom c)"] Zp.domain_axioms inv_cancelR(1) + Q\<^sub>p_def Zp.Units_closed Zp.inv_cancelR(1) Zp.m_closed Zp.nonzero_closed Zp.nonzero_memE(2) Z\<^sub>p_def ac_Zp_is_Unit assms(1) assms(2) mult_assoc padic_fields.numer_denom_facts(2) padic_fields_axioms by auto + then have "(inv \<^bsub>Z\<^sub>p\<^esub> (ac_Zp (denom c))) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (numer c)) = (ac_Zp a) \\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub> (ac_Zp b)" + using ac_Zp_is_Unit[of "ac_Zp b"] Zp.domain_axioms inv_cancelL(2) + by (smt Q\<^sub>p_def Zp.Units_inv_closed Zp.Units_r_inv Zp.nonzero_closed Zp.nonzero_memE(2) Zp.r_one Z\<^sub>p_def ac_Zp_is_Unit assms(1) assms(3) mult_assoc mult_comm padic_fields.numer_denom_facts(2) padic_fields_axioms) + then show ?thesis + by (simp add: angular_component_def mult_comm) +qed + +lemma angular_component_closed: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component a \ carrier Z\<^sub>p" + using ac_fract assms Q\<^sub>p_def Qp.nonzero_closed Qp.nonzero_memE(2) Zp.Units_inv_closed Zp.m_closed + Zp.nonzero_closed Zp.nonzero_memE(2) Z\<^sub>p_def ac_Zp_is_Unit angular_component_def local.numer_denom_facts(3) + padic_fields.numer_denom_facts(1) padic_fields.numer_denom_facts(2) padic_fields_axioms padic_integers.ac_Zp_in_Zp padic_integers_def prime by auto + +lemma angular_component_unit: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component a \ Units Z\<^sub>p" + using ac_fract[of a "numer a" "denom a"] Q\<^sub>p_def Qp.nonzero_closed + Qp.nonzero_memE(2) Zp.Units_inv_Units Zp.Units_m_closed + Zp.nonzero_closed Zp.nonzero_memE(2) Z\<^sub>p_def ac_Zp_is_Unit + angular_component_def assms local.numer_denom_facts(1) + local.numer_denom_facts(2) padic_fields.numer_denom_facts(3) + padic_fields_axioms by auto + +lemma angular_component_factors_x: + assumes "x \ nonzero Q\<^sub>p" + shows "x = (\[^](ord x)) \ \ (angular_component x)" +proof- + have 0: "angular_component x = (ac_Zp (numer x)) \\<^bsub>Z\<^sub>p\<^esub> (inv\<^bsub>Z\<^sub>p\<^esub> ac_Zp (denom x))" + by (simp add: angular_component_def) + have 1: "(numer x) = (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (numer x))) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (numer x))" + proof- + have "numer x \ nonzero Z\<^sub>p" + using assms unfolding Q\<^sub>p_def + by (simp add: numer_nonzero) + then show ?thesis using ac_Zp_factors_x[of "numer x"] + by (simp add: ac_Zp_factors') + qed + have 2: "(denom x) = (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (denom x))) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (denom x))" + proof- + have "denom x \ nonzero Z\<^sub>p" + using nonzero_memE assms numer_denom_facts(2) + by (simp add: Qp.nonzero_closed) + then show ?thesis + using ac_Zp_factors_x[of "denom x"] Zp.nonzero_closed Zp.nonzero_memE(2) by auto + qed + have 3: "\ (angular_component x) = frac (ac_Zp (numer x)) (ac_Zp (denom x))" + by (metis "0" Q\<^sub>p_def Qp.nonzero_closed Qp.nonzero_memE(2) Zp.nonzero_closed Zp.nonzero_memE(2) Zp.numer_denom_facts(2) Z\<^sub>p_def Z\<^sub>p_division_Qp_0 ac_Zp_is_Unit angular_component_closed assms local.inc_def padic_fields.numer_denom_facts(2) padic_fields.numer_denom_facts(3) padic_fields_axioms) + have 4: "(\[^]((ord x))) \ \ (angular_component x) = + (\[^]((ord_Zp (numer x)) - (ord_Zp (denom x)))) \ frac (ac_Zp (numer x)) (ac_Zp (denom x))" + using "3" ord_def by presburger + have 5: "(\[^]((ord x))) \ \ (angular_component x) = + frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>((ord_Zp (numer x)))) (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (denom x))) \ frac (ac_Zp (numer x)) (ac_Zp (denom x))" + proof- + have "(\[^]((ord_Zp (numer x)) - (ord_Zp (denom x)))) + = frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>((ord_Zp (numer x)))) (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (denom x)))" + by (metis Q\<^sub>p_def Qp.nonzero_closed Z\<^sub>p_def assms domain.nonzero_memE(1) domain.nonzero_memE(2) numer_nonzero ord_pos p_pow_diff padic_fields.numer_denom_facts(2) padic_fields_axioms padic_int_is_domain prime) + then show ?thesis using 4 by metis + qed + have 6: "(\[^]((ord x))) \ \ (angular_component x) = + frac (\

[^]\<^bsub>Z\<^sub>p\<^esub>((ord_Zp (numer x))) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (numer x))) (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (denom x)) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (denom x)))" + using 5 frac_mult[of "\

[^]\<^bsub>Z\<^sub>p\<^esub>((ord_Zp (numer x)))" " (\

[^]\<^bsub>Z\<^sub>p\<^esub>(ord_Zp (denom x)))" "(ac_Zp (numer x)) " " (ac_Zp (denom x))"] mult_comm + by (metis "2" Q\<^sub>p_def Qp.nonzero_closed Zp.integral_iff Zp.nonzero_closed Zp.nonzero_memE(2) ac_Zp_in_Zp assms local.numer_denom_facts(2) not_nonzero_Zp numer_nonzero ord_pos p_int_pow_nonzero) + then show ?thesis + using "1" "2" nonzero_memE assms numer_denom_facts(5) Qp.nonzero_closed by auto +qed + +lemma angular_component_mult: + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + shows "angular_component (x \ y) = (angular_component x) \\<^bsub>Z\<^sub>p\<^esub> (angular_component y)" +proof- + obtain a b where "a = numer x" and + "b = denom x" and + a_nz: "a \ nonzero Z\<^sub>p" and + b_nz: "b \ nonzero Z\<^sub>p" and + x_frac: "x = frac a b" + using assms Qp.nonzero_memE[of x] + by (meson local.numer_denom_facts(1) local.numer_denom_facts(2) + local.numer_denom_facts(3) local.numer_denom_facts(5) not_nonzero_Zp) + obtain c d where "c = numer y" and + "d = denom y" and + c_nz: "c \ nonzero Z\<^sub>p" and + d_nz: "d \ nonzero Z\<^sub>p" and + y_frac: "y = frac c d" + using assms Qp.nonzero_memE[of y] + by (meson local.numer_denom_facts(1) local.numer_denom_facts(2) local.numer_denom_facts(3) local.numer_denom_facts(5) not_nonzero_Zp) + have 0: "(x \ y) = frac (a \\<^bsub>Z\<^sub>p\<^esub> c) (b \\<^bsub>Z\<^sub>p\<^esub> d)" + using a_nz b_nz c_nz d_nz frac_mult x_frac y_frac + by (simp add: Zp.nonzero_closed) + have 1: "angular_component (x \ y) = (ac_Zp (a \\<^bsub>Z\<^sub>p\<^esub> c)) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp (b \\<^bsub>Z\<^sub>p\<^esub> d))" + by (metis (mono_tags, lifting) "0" Qp.nonzero_mult_closed Zp.integral_iff Zp.nonzero_closed Zp.nonzero_mult_closed a_nz ac_fract assms(1) assms(2) b_nz c_nz d_nz not_nonzero_Zp) + have 2: "angular_component (x \ y) = (ac_Zp a) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp c) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> ((ac_Zp b) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp d))" + by (simp add: "1" a_nz ac_Zp_mult b_nz c_nz d_nz) + have 3: "angular_component (x \ y) = (ac_Zp a) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp c) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp b) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp d)" + using 2 + by (simp add: Zp.inv_of_prod Zp.nonzero_closed Zp.nonzero_memE(2) ac_Zp_is_Unit b_nz d_nz mult_assoc) + + have 4: "angular_component (x \ y) = (ac_Zp a) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp b) \\<^bsub>Z\<^sub>p\<^esub> (ac_Zp c) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp d)" + using "3" a_nz ac_Zp_in_Zp ac_Zp_is_Unit b_nz c_nz m_assoc mult_comm + mult_assoc by auto + have 5: "angular_component (x \ y) = ((ac_Zp a) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp b)) \\<^bsub>Z\<^sub>p\<^esub> ( (ac_Zp c) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp d))" + using 4 by (simp add: mult_assoc) + then show ?thesis + by (simp add: Z\<^sub>p_def \a = numer x\ \b = denom x\ \c = numer y\ \d = denom y\ + padic_fields.angular_component_def padic_fields_axioms) +qed + +lemma angular_component_inv: + assumes "x \ nonzero Q\<^sub>p" + shows "angular_component (inv\<^bsub>Q\<^sub>p\<^esub> x) = inv\<^bsub>Z\<^sub>p\<^esub> (angular_component x)" + by (metis Q\<^sub>p_def Qp.one_closed Zp.Units_r_inv Zp.inv_unique' Zp.nonzero_closed Zp.nonzero_one_closed Zp.zero_not_one Z\<^sub>p_def \_def ac_Zp_is_Unit angular_component_closed angular_component_mult assms frac_nonzero_inv(1) frac_nonzero_inv(2) inc_equation inc_of_one nonzero_inverse_Qp padic_fields.ac_fract padic_fields_axioms) + +lemma angular_component_one: +"angular_component \ = \\<^bsub>Z\<^sub>p\<^esub>" + using \_def ac_Zp_one ac_fract inc_equation inc_of_one + Zp.nonzero_one_closed by fastforce + +lemma angular_component_ord_zero: + assumes "ord x = 0" + assumes "x \ nonzero Q\<^sub>p" + shows "\ (angular_component x) = x" +proof- + have 0: "ord x \0" + using assms by auto + have 1: "x \ nonzero Q\<^sub>p" + proof- + have "x \ \" + using nonzero_def assms(2) Qp.nonzero_memE(2) by auto + then show ?thesis + by (simp add: assms(2)) + qed + obtain a where a_def: "a \ nonzero Z\<^sub>p \ \ a = x" + using 0 1 assms Qp_inc_id[of x] + by metis + then have "x = frac a \\<^bsub>Z\<^sub>p\<^esub>" + using local.inc_def Zp.nonzero_closed by blast + then have "angular_component x = ac_Zp a \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp \\<^bsub>Z\<^sub>p\<^esub>)" + using ac_fract[of x a ] "1" nonzero_memE Q\<^sub>p_def Zp.nonzero_closed Zp.nonzero_one_closed Z\<^sub>p_def a_def + local.frac_closed padic_fields.ac_fract padic_fields_axioms by auto + then have "angular_component x = ac_Zp a \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: ac_Zp_one) + then have 0: "angular_component x = ac_Zp a \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + by simp + have "ac_Zp a \ nonzero Z\<^sub>p" + using Zp.Units_nonzero Zp.nonzero_closed Zp.not_nonzero_memI a_def ac_Zp_is_Unit by auto + thus ?thesis using 0 + by (metis Zp.l_one Zp.nonzero_closed Zp_defs(1) Zp_defs(4) a_def ac_Zp_factors' + assms(1) int_pow_0 mult_comm ord_of_inc) +qed + +lemma angular_component_of_inclusion: + assumes "x \ nonzero Z\<^sub>p" + assumes "y = \ x" + shows "angular_component y = ac_Zp x" +proof- + have "y = local.frac x \\<^bsub>Z\<^sub>p\<^esub>" + using assms + by (simp add: Zp.nonzero_closed local.inc_def) + then have 0: "angular_component y = (ac_Zp x) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> ac_Zp (\\<^bsub>Z\<^sub>p\<^esub>)" + using assms ac_fract[of y x ] + by (simp add: \y = \ x\ Zp.nonzero_closed Zp.nonzero_one_closed local.frac_closed) + have 1: "ac_Zp \\<^bsub>Z\<^sub>p\<^esub> = \\<^bsub>Z\<^sub>p\<^esub>" + using ac_Zp_one by blast + have 2: "(ac_Zp x) \ carrier Z\<^sub>p" + using Zp.Units_closed Zp.nonzero_closed ac_Zp_is_Unit assms(1) nonzero_fraction_imp_numer_not_zero nonzero_numer_imp_nonzero_fraction by auto + have 3: "inv \<^bsub>Z\<^sub>p\<^esub> (ac_Zp \\<^bsub>Z\<^sub>p\<^esub>) = \\<^bsub>Z\<^sub>p\<^esub>" + using 1 + by simp + have 4: "(ac_Zp x) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> ac_Zp (\\<^bsub>Z\<^sub>p\<^esub>) = (ac_Zp x) \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + using "3" by auto + have 5: "(ac_Zp x) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> ac_Zp (\\<^bsub>Z\<^sub>p\<^esub>) = (ac_Zp x)" + using 4 2 Zp.domain_axioms by simp + then show ?thesis + by (simp add: "0") +qed + +lemma res_uminus: + assumes "k > 0" + assumes "f \ carrier Z\<^sub>p" + assumes "c \ carrier (Zp_res_ring k)" + assumes "c = \\<^bsub>Zp_res_ring k\<^esub> (f k)" + shows "c = ((\\<^bsub>Z\<^sub>p\<^esub> f) k)" + using Z\<^sub>p_def assms(2) assms(4) prime Zp_residue_a_inv(1) by auto + +lemma ord_fract: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "ord (a \

b) = ord a - ord b" + using assms nonzero_memE nonzero_def nonzero_inverse_Qp ord_mult ord_of_inv + Qp.nonzero_closed Qp.nonzero_memE(2) by presburger + +lemma val_fract: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "val (a \
b) = val a - val b" + apply(cases "a = \") + apply (simp add: Units_eq_nonzero assms(2) val_def) +proof- + assume "a \ \" + then have 0: "a \ nonzero Q\<^sub>p" + by (simp add: Qp.nonzero_memI assms(1)) + have 1: "(a \
b) \ nonzero Q\<^sub>p" + by (simp add: "0" Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) nonzero_inverse_Qp) + show ?thesis using 0 1 assms + by (simp add: ord_fract) +qed + +lemma zero_fract: + assumes "a \ nonzero Q\<^sub>p" + shows "\ \
a = \" + by (simp add: Units_eq_nonzero assms) + +lemma fract_closed: + assumes "a \ carrier Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "(a \
b) \ carrier Q\<^sub>p" + by (simp add: Units_eq_nonzero assms(1) assms(2)) + +lemma val_of_power: + assumes "a \ nonzero Q\<^sub>p" + shows "val (a[^](n::nat)) = n*(val a)" + unfolding val_def using assms + by (simp add: Qp.nonzero_memE(2) Qp_nat_pow_nonzero nonzero_nat_pow_ord) + +lemma val_zero_imp_val_pow_zero: + assumes "a \ carrier Q\<^sub>p" + assumes "val a = 0" + shows "val (a[^](n::nat)) = 0" + apply(induction n) + using assms apply simp +proof- + fix n ::nat + assume A: "val (a [^] n) = 0" + have 0: "a [^] Suc n = a [^] n \ a" + by simp + show "val (a [^] Suc n) = 0 " + unfolding 0 using A assms + by (simp add: val_mult) +qed + +text\val and ord of powers of p\ + +lemma val_p_nat_pow: +"val (\[^](k::nat)) = eint k" + by (simp add: ord_p_pow_nat) + +lemma ord_p_int_pow: +"ord (\[^](k::int)) = k" + by (simp add: ord_p_pow_int) + +lemma ord_p_nat_pow: +"ord (\[^](k::nat)) = k" + by (simp add: ord_p_pow_nat) + +lemma val_nonzero_frac: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "val (a \
b) = c" + shows "val a = val b + c" +proof- + obtain n where n_def: "c = eint n" + using assms val_ord by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed inv_in_frac(1) inv_in_frac(2) val_def) + have 0: "ord (a \
b) = n" + by (metis assms(3) eint.distinct(2) eint.inject n_def val_def) + have 1: "ord a - ord b = n" + using 0 assms by (metis ord_fract) + have 2: "ord a = ord b + n" + using 1 by presburger + show "val a = val b + c" + unfolding n_def using 2 assms + by (metis Qp.nonzero_memE(1) fract_closed local.fract_cancel_right n_def val_mult) +qed + +lemma val_nonzero_frac': + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "val (a \
b) = 0" + shows "val a = val b" + using val_nonzero_frac[of a b 0] + by (metis arith_simps(50) assms(1) assms(2) assms(3)) + +lemma equal_val_imp_equal_ord: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val a = val b" + shows "ord a = ord b" "b \ nonzero Q\<^sub>p" + using assms + apply (metis eint.inject eint.simps(3) eint_ord_simps(4) val_nonzero val_ord) + using assms unfolding nonzero_def + using val_def by auto + +lemma int_pow_ord: + assumes "a \ nonzero Q\<^sub>p" + shows "ord (a[^](i::int)) = i* (ord a)" +proof(induction i) + show "\n. ord (a [^] int n) = int n * ord a" + using assms + by (metis int_pow_int mult_of_nat_commute nonzero_nat_pow_ord) + show "\n. ord (a [^] - int (Suc n)) = - int (Suc n) * ord a" + proof- fix n + have "(a [^] - int (Suc n)) \ (a [^] int (Suc n)) = a [^] (- int (Suc n) + int (Suc n))" + using assms Qp_int_pow_add by blast + hence "(a [^] - int (Suc n)) \ (a [^] int (Suc n)) = \" + using assms + by (metis ab_group_add_class.ab_left_minus int_pow_0) + hence "ord (a [^] - int (Suc n)) + ord (a [^] int (Suc n)) = 0" + by (metis Qp_int_pow_nonzero assms ord_mult ord_one) + thus "ord (a [^] - int (Suc n)) = - int (Suc n) * ord a" + using nonzero_nat_pow_ord assms + by (metis \\n. ord (a [^] int n) = int n * ord a\ add.inverse_inverse add_left_cancel more_arith_simps(4) more_arith_simps(7)) + qed +qed + +lemma int_pow_val: + assumes "a \ nonzero Q\<^sub>p" + shows "val (a[^](i::int)) = i* (val a)" + using int_pow_ord Qp_int_pow_nonzero assms times_eint_simps(1) val_ord by presburger + +lemma neg_int_pow_val: + assumes "a \ nonzero Q\<^sub>p" + shows "val (a[^]-(i::int)) = - (val (a[^]i))" + using int_pow_val[of a i] + by (metis Qp_int_pow_nonzero assms int_pow_ord mult_minus_left val_ord val_p_int_pow val_p_int_pow_neg) + +lemma int_pow_sum_val: + assumes "a \ nonzero Q\<^sub>p" + shows "val (a[^]((i::int) + j)) = (val (a[^]i)) + val (a[^]j)" + using assms Qp.int_pow_add Qp_int_pow_nonzero Units_eq_nonzero val_mult0 by presburger + +lemma int_pow_diff_val: + assumes "a \ nonzero Q\<^sub>p" + shows "val (a[^]((i::int) - j)) = (val (a[^]i)) - val (a[^]j)" +proof- + obtain k where k_def: "val a = eint k" + using assms val_ord by blast + have 0: "val (a[^]((i::int) - j)) = eint ((i-j)*k)" + using assms k_def int_pow_val times_eint_simps(1) by presburger + show ?thesis unfolding 0 using k_def + by (metis "0" Qp_int_pow_nonzero Rings.ring_distribs(3) assms eint.simps(1) idiff_eint_eint int_pow_ord val_ord) +qed + +lemma nat_add_pow_mult_assoc: + assumes "a \ carrier Q\<^sub>p" + shows "[(n::nat)]\a = [n]\\ \ a" + using assms Qp.add_pow_ldistr Qp.l_one Qp.one_closed by presburger + +lemma(in padic_integers) equal_res_imp_equal_ord_Zp: + assumes "N > 0" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a N = b N" + assumes "a N \ 0" + shows "ord_Zp a = ord_Zp b" +proof- + have "(a \ b) N = 0" + using assms(2) assms(3) assms(4) res_diff_zero_fact'' by blast + have "ord_Zp a < N" + using assms zero_below_ord by force + then show ?thesis + by (smt R.minus_closed Zp_def \(a \ b) N = 0\ above_ord_nonzero assms(2) assms(3) assms(4) assms(5) ord_Zp_ultrametric_eq' ord_Zp_ultrametric_eq'' padic_integers.ord_Zp_not_equal_imp_notequal(2) padic_integers.ord_of_nonzero(2) padic_integers.ord_pos padic_integers_axioms residue_of_zero(2)) +qed + +lemma(in padic_integers) equal_res_mod: + assumes "N > k" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a N = b N" + shows "a k = b k" +proof- + have 0: "a k = a N mod p^k" + using assms + by (metis Zp_int_inc_rep Zp_int_inc_res less_or_eq_imp_le p_residue_padic_int) + have 1: "b k = b N mod p^k" + using assms + by (metis Zp_int_inc_rep Zp_int_inc_res less_or_eq_imp_le p_residue_padic_int) + show ?thesis unfolding 0 1 assms using assms by blast +qed + +lemma Qp_char_0: + assumes "(n::nat) \ 0" + shows "[n]\\ \ \" +proof- + have "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \ \\<^bsub>Z\<^sub>p\<^esub>" + using Zp_char_0' assms by blast + thus ?thesis using inc_of_nat + by (metis Qp.nat_inc_zero Zp.nat_inc_zero Zp_nat_inc_closed \_def inc_inj2) +qed + +lemma Qp_char_0_int: + assumes "(n::int) \ 0" + shows "[n]\\ \ \" +proof- + have "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \ \\<^bsub>Z\<^sub>p\<^esub>" + using Zp_char_0 assms + by (metis int_inc_mult linorder_neqE_linordered_idom mult_zero_l zero_less_mult_iff) + thus ?thesis using inc_of_int + by (metis Q\<^sub>p_def Zp_int_inc_closed \_def inc_inj1) +qed + +lemma add_int_pow_inject: + assumes "[(k::int)]\\ = [(j::int)]\\" + shows "k = j" +proof(rule ccontr) + assume A: "k \ j" + then have 0: "k - j \0 " + by presburger + hence 1: "[(k-j)]\\ \ \ " + using Qp_char_0_int by blast + hence 2: "[k]\\ \ [(-j)]\ \ \ \ " + by (metis (no_types, opaque_lifting) Qp.add.int_pow_mult Qp.one_closed diff_minus_eq_add diff_zero minus_diff_eq) + hence 2: "[k]\\ \ [j]\ \ \ \ " + by (metis Qp.add.int_pow_mult Qp.int_inc_zero Qp.one_closed assms more_arith_simps(4)) + thus False using assms + using Qp.int_inc_closed Qp.r_right_minus_eq by blast +qed + +lemma val_ord_nat_inc: + assumes "(n::nat) > 0" + shows "ord ([n]\\) = val([n]\\)" + using val_ord assms Qp_char_0 + by (metis less_irrefl_nat val_def) + +lemma val_ord_int_inc: + assumes "(n::int) \ 0" + shows "ord ([n]\\) = val([n]\\)" + using val_ord assms Qp_char_0_int val_def by presburger + + + (********************************************************************************************) + (********************************************************************************************) + subsubsection\Reduced Angular Component Maps\ + (********************************************************************************************) + (********************************************************************************************) + +definition ac :: "nat \ padic_number \ int" where +"ac n x = (if x = \ then 0 else (angular_component x) n)" + +lemma ac_in_res_ring: + assumes "x \ nonzero Q\<^sub>p" + shows "ac n x \ carrier (Zp_res_ring n)" + unfolding ac_def + using assms angular_component_closed[of x] + by (simp add: Qp.nonzero_memE(2) residues_closed) + +lemma ac_in_res_ring'[simp]: + assumes "x \ carrier Q\<^sub>p" + shows "ac n x \ carrier (Zp_res_ring n)" + apply(cases "x \ nonzero Q\<^sub>p") + using ac_in_res_ring apply blast + by (metis Qp.nonzero_memI ac_def assms mod_0 mod_in_carrier p_res_ring_car p_residue_alt_def) + +lemma ac_mult': + assumes "x \ nonzero Q\<^sub>p" + assumes "y \ nonzero Q\<^sub>p" + shows "ac n (x \ y) = (ac n x) \\<^bsub>Zp_res_ring n\<^esub> (ac n y)" + unfolding ac_def +proof- + have 0: "angular_component (x \ y) = angular_component x \\<^bsub>Z\<^sub>p\<^esub> angular_component y" + using assms angular_component_mult[of x y] + by auto + show "(if x \ y = \ then 0 else angular_component (x \ y) n) = + (if x = \ then 0 else angular_component x n) \\<^bsub>Zp_res_ring n\<^esub> (if y = \ then 0 else angular_component y n)" + using assms angular_component_closed[of x] angular_component_closed[of y] + by (simp add: "0" Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) residue_of_prod) +qed + +lemma ac_mult: + assumes "x \ carrier Q\<^sub>p" + assumes "y \ carrier Q\<^sub>p" + shows "ac n (x \ y) = (ac n x) \\<^bsub>Zp_res_ring n\<^esub> (ac n y)" + apply(cases "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p") + apply (simp add: ac_mult') + using assms unfolding ac_def + by (smt Qp.integral_iff Qp.nonzero_memI residue_times_zero_l residue_times_zero_r) + +lemma ac_one[simp]: + assumes "n \ 1" + shows "ac n \ = 1" +proof- + have "ac n \ = \\<^bsub>Z\<^sub>p\<^esub> n" + unfolding ac_def + using angular_component_one + by simp + then show ?thesis + using assms residue_of_one(2) by auto +qed + +lemma ac_one': + assumes "n > 0" + shows "ac n \ = \\<^bsub>Zp_res_ring n\<^esub>" + using assms residue_ring_def + by auto + +lemma ac_units: + assumes "x \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "ac n x \ Units (Zp_res_ring n)" +proof- + obtain y where y_def: "y = inv x" + by simp + have y_nz: "y \ nonzero Q\<^sub>p" + using assms(1) nonzero_inverse_Qp y_def + by blast + have 0: "ac n (x \ y) = (ac n x) \\<^bsub>Zp_res_ring n\<^esub> (ac n y)" + using ac_mult' assms(1) y_nz by blast + have 1: "x \ y = \" + by (metis Qp.field_axioms Qp.nonzero_memE Qp.nonzero_memE assms(1) field.field_inv(2) y_def) + have 2: "(ac n x) \\<^bsub>Zp_res_ring n\<^esub> (ac n y) = \\<^bsub>Zp_res_ring n\<^esub>" + using "0" "1" ac_one' assms(2) + by auto + show ?thesis + by (metis "2" R_comm_monoid ac_in_res_ring assms(1) assms(2) comm_monoid.UnitsI(1) nonzero_inverse_Qp y_def) + +qed + +lemma ac_inv: + assumes "x \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "ac n (inv x) = inv\<^bsub>Zp_res_ring n\<^esub> (ac n x)" +proof- + have "x \ inv x = \" + by (simp add: Qp.monoid_axioms Units_eq_nonzero assms(1) monoid.Units_r_inv) + then have "ac n (x \ inv x) = \\<^bsub>Zp_res_ring n\<^esub>" + by (simp add: \(x \
x) = \\ ac_one' assms(2)) + then have "ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (inv x) = \\<^bsub>Zp_res_ring n\<^esub>" + using Units_eq_nonzero Units_inverse_Qp ac_mult' assms(1) + by metis + then show ?thesis + by (metis R_comm_monoid ac_in_res_ring assms(1) assms(2) comm_monoid.comm_inv_char nonzero_inverse_Qp) +qed + +lemma ac_inv': + assumes "x \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "ac n (inv x) \\<^bsub>Zp_res_ring n\<^esub> (ac n x) = \\<^bsub>Zp_res_ring n\<^esub>" + using ac_inv[of x n] ac_units[of x n] assms + by (metis (no_types, opaque_lifting) Qp.field_axioms Qp.nonzero_memE + Qp.nonzero_memE Units_eq_nonzero Units_inverse_Qp ac_mult ac_one' field.field_inv(1)) + +lemma ac_inv'': + assumes "x \ nonzero Q\<^sub>p" + assumes "n > 0" + shows " (ac n x) \\<^bsub>Zp_res_ring n\<^esub> ac n (inv x)= \\<^bsub>Zp_res_ring n\<^esub>" + using ac_inv' assms(1) assms(2) residue_mult_comm by auto + +lemma ac_inv''': + assumes "x \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "(ac n x) \\<^bsub>Zp_res_ring n\<^esub> ac n (inv x)= 1" + "ac n (inv x) \\<^bsub>Zp_res_ring n\<^esub> (ac n x) = 1" + by (auto simp add: ac_inv'' ac_inv' assms(1) assms(2) p_res_ring_one) + +lemma ac_val: + assumes "a \ nonzero Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + assumes "val a = val b" + assumes "val (a \ b) \ val a + n" + shows "ac n a = ac n b" +proof- + obtain m where m_def: "m = ord a" + by simp + have 0: "a = (\[^]m) \ \ (angular_component a)" + by (simp add: angular_component_factors_x assms(1) m_def) + have 1: "b = (\[^]m) \ \ (angular_component b)" + proof- + have "ord b = ord a" + using assms(1) assms(2) assms(3) by auto + then show ?thesis + by (metis angular_component_factors_x assms(2) m_def) + qed + have 2: "(a \b) = (\[^]m) \ \ (angular_component a) + \(\[^]m) \ \ (angular_component b)" + using "0" "1" by auto + have 3: "(a \b) = (\[^]m) \( \ (angular_component a) + \ \ (angular_component b))" + using 2 assms angular_component_closed inc_closed Qp.cring_simprules Qp.r_minus_distr + by (metis (no_types, lifting) p_intpow_closed(1)) + have 4: "(a \b) = (\[^]m) \( \ ((angular_component a) + \\<^bsub>Z\<^sub>p\<^esub>(angular_component b)))" + by (simp add: "3" angular_component_closed assms(1) assms(2) inc_of_diff) + have 5: "val (a \b) = m + val (( \ ((angular_component a) + \\<^bsub>Z\<^sub>p\<^esub>(angular_component b))))" + by (metis "4" Zp.nonzero_one_closed angular_component_closed assms(1) assms(2) + cring.cring_simprules(4) frac_closed local.inc_def ord_p_pow_int p_intpow_closed(1) + p_intpow_closed(2) Zp.domain_axioms domain_def val_mult val_ord) + have 6: "m = val a" + using Q\<^sub>p_def assms(1) m_def by auto + have 7: "m + val (\ (angular_component a \\<^bsub>Z\<^sub>p\<^esub> angular_component b)) \ val a + n" + using "5" assms(4) by presburger + have 8: "m + val (\ (angular_component a \\<^bsub>Z\<^sub>p\<^esub> angular_component b)) \ m + n" + using "6" "7" + by (metis plus_eint_simps(1)) + have 9: "val (\ (angular_component a \\<^bsub>Z\<^sub>p\<^esub> angular_component b)) \ n" + using 8 add_le_cancel_left eint_ile eint_ord_simps(1) linear plus_eint_simps(1) + by metis + have 10: "val_Zp (angular_component a \\<^bsub>Z\<^sub>p\<^esub> angular_component b) \ n" + using 9 + by (metis angular_component_closed assms(1) assms(2) cring.cring_simprules(4) + Zp.domain_axioms domain_def val_of_inc) + have 11: "(angular_component a \\<^bsub>Z\<^sub>p\<^esub> angular_component b) n = \\<^bsub>Zp_res_ring n\<^esub>" + using 9 + by (simp add: "10" angular_component_closed assms(1) assms(2) p_res_ring_zero zero_below_val_Zp) + have 12: "(angular_component a n) \\<^bsub>Zp_res_ring n\<^esub> (angular_component b n) = \\<^bsub>Zp_res_ring n\<^esub>" + using "11" angular_component_closed assms(2) residue_of_diff by auto + have 13: "(angular_component a n) = (angular_component b n)" + apply(cases "n = 0") + apply (metis angular_component_closed assms(1) assms(2) p_res_ring_0' residues_closed) + using 12 angular_component_closed residue_closed ring.ring_simprules[of "Zp_res_ring n"] + by (meson assms(1) assms(2) cring_def prime residues.cring residues_n ring.r_right_minus_eq) + then show ?thesis + by (simp add: Qp.nonzero_memE ac_def assms(1) assms(2)) +qed + +lemma angular_component_nat_pow: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component (a [^] (k::nat)) = (angular_component a) [^]\<^bsub>Z\<^sub>p\<^esub> k" + apply(induction k) + apply (simp add: angular_component_one) + by (simp add: Qp_nat_pow_nonzero angular_component_mult assms) + +lemma angular_component_int_pow: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component (a [^] (k::int)) = (angular_component a) [^]\<^bsub>Z\<^sub>p\<^esub> k" + apply(cases "k \ 0") + using angular_component_nat_pow assms + apply (metis int_pow_int nonneg_int_cases) +proof- + assume "\ 0 \k" + show "angular_component (a [^] k) = angular_component a [^] \<^bsub>Z\<^sub>p\<^esub> k" + using angular_component_nat_pow[of a "nat (-k)"] assms + Qp_nat_pow_nonzero[of a] \\ 0 \ k\ angular_component_inv[of "(a [^] k)"] int_pow_def2 + by (metis angular_component_nat_pow angular_component_inv) +qed + +lemma ac_nat_pow: + assumes "a \ nonzero Q\<^sub>p" + shows "ac n (a [^] (k::nat)) = (ac n a)^ k mod (p^n)" +proof(cases "k = 0") + case True + show ?thesis apply(cases "n = 0") + apply (metis Group.nat_pow_0 Qp.one_closed True ac_in_res_ring' mod_self p_res_ring_0' power_0) + using True prime residues.one_cong residues.res_one_eq residues_n by auto +next + case False + show ?thesis + apply(cases "n = 0") + apply (metis Qp_nat_pow_nonzero ac_in_res_ring assms mod_by_1 p_res_ring_0' power_0) + using assms angular_component_nat_pow + by (metis Qp.nonzero_closed Qp.nonzero_pow_nonzero Qp.not_nonzero_memI ac_def angular_component_closed neq0_conv power_residue) +qed + +lemma ac_nat_pow': + assumes "a \ nonzero Q\<^sub>p" + assumes "n \0" + shows "ac n (a [^] (k::nat)) = (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" +proof- + have "(ac n a)^ k mod (p^n) = (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" + apply(induction k) + apply (metis Group.nat_pow_0 assms(2) prime residues.pow_cong residues_n) + by (metis Group.nat_pow_Suc Qp_nat_pow_nonzero ac_mult' ac_nat_pow assms(1)) + then show ?thesis + by (simp add: ac_nat_pow assms(1)) +qed + +lemma ac_int_pow: + assumes "a \ nonzero Q\<^sub>p" + assumes "n > 0" + shows "ac n (a [^] (k::int)) = (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" + apply(cases "k \0") + using assms ac_nat_pow' + apply (metis int.lless_eq int.nat_pow_0 p_residues pow_nat residues_def) + using assms ac_nat_pow' ac_inv[of "a [^] k"] ac_units[of "a [^] k"] + by (metis Qp_nat_pow_nonzero ac_inv int_pow_def2 neq0_conv) + +lemma angular_component_p: +"angular_component \ = \\<^bsub>Z\<^sub>p\<^esub>" +proof- + have "\ = frac \

\\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: Zp_nat_inc_closed local.inc_def p_inc) + then have 0: "angular_component \ = (ac_Zp \

) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> (ac_Zp \\<^bsub>Z\<^sub>p\<^esub>)" + using ac_Zp_one unfolding angular_component_def + by (metis Q\<^sub>p_def Qp.int_inc_closed Zp.one_nonzero Zp_int_inc_closed ac_fract + local.numer_denom_facts(2) local.numer_denom_facts(5) nonzero_fraction_imp_nonzero_numer + nonzero_numer_imp_nonzero_fraction numer_nonzero p_nonzero) + then have "angular_component \ = (ac_Zp \

) \\<^bsub>Z\<^sub>p\<^esub> inv\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: ac_Zp_one) + then have "angular_component \ = (ac_Zp \

)" + by (simp add: ac_Zp_p) + then show ?thesis + using ac_Zp_p + by simp +qed + +lemma angular_component_p_nat_pow: +"angular_component (\ [^] (n::nat)) = \\<^bsub>Z\<^sub>p\<^esub>" + apply(induction n) + apply (simp add: angular_component_one) + using angular_component_nat_pow angular_component_p nat_pow_one p_nonzero Zp.nat_pow_one + by presburger + +lemma angular_component_p_int_pow: +"angular_component (\ [^] (n::int)) = \\<^bsub>Z\<^sub>p\<^esub>" + apply(cases "n \ 0") + apply (metis angular_component_p_nat_pow int_pow_int nonneg_int_cases) + using angular_component_p_nat_pow[of "nat (-n)"] angular_component_inv[of "(\ [^] n)"] angular_component_inv[of "\ [^] (nat (-n))"] + by (metis (mono_tags, opaque_lifting) Group.nat_pow_0 Zp.inv_one add.inverse_neutral angular_component_one int_nat_eq + int_pow_def2 nat_0_iff nat_int nat_zminus_int p_natpow_closed(2) ) + +lemma ac_p_nat_pow: + assumes "k > 0" + shows "ac k (\ [^] (n::nat)) = 1" +proof- + have "\ (\ [^] n) = \" + by (simp add: Qp.nonzero_memE) + have " angular_component (\ [^] n) k = 1" + using assms angular_component_p_nat_pow[of n] ac_def ac_one angular_component_one + by auto + then show ?thesis + unfolding ac_def using angular_component_p_nat_pow[of n] + by (simp add: \(\ [^] n) \ \\) +qed + +lemma ac_p: + assumes "k > 0" + shows "ac k \ = 1" + by (metis Qp.int_inc_closed Qp.nat_pow_eone ac_p_nat_pow assms) + +lemma ac_p_int_pow: + assumes "k > 0" + shows "ac k (\ [^] (n::int)) = 1" + using Qp.nonzero_memE(2) ac_def ac_one' angular_component_one angular_component_p_int_pow + assms p_intpow_closed(2) p_res_ring_one by auto + +lemma angular_component_p_nat_pow_factor: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component ((\ [^] (n::nat)) \ a) = angular_component a" +proof- + have 0: "angular_component ((\ [^] n) \ a) = angular_component (\ [^] n) \\<^bsub>Z\<^sub>p\<^esub> angular_component a" + using assms angular_component_mult[of "(\ [^] (n::nat))" a] p_natpow_closed(2) + by blast + have 1: "angular_component (\ [^] n) = \\<^bsub>Z\<^sub>p\<^esub>" + by (simp add: angular_component_p_nat_pow) + have 2: "angular_component ((\ [^] n) \ a) = \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> angular_component a" + by (simp add: "0" "1") + then show ?thesis using angular_component_closed[of a] assms Zp.domain_axioms + by (simp add: assms cring.cring_simprules(12) domain.axioms(1)) +qed + +lemma ac_p_nat_pow_factor: + assumes "m > 0" + assumes "a \ nonzero Q\<^sub>p" + shows "ac m ((\ [^] (n::nat)) \ a) = ac m a" + using angular_component_p_nat_pow_factor assms ac_def + by (metis (no_types, lifting) Qp.field_axioms Qp_nat_pow_nonzero Qp.nonzero_memE + Qp.nonzero_memE Ring.integral p_nonzero) + +lemma angular_component_p_nat_pow_factor_right: + assumes "a \ nonzero Q\<^sub>p" + shows "angular_component (a \ (\ [^] (n::nat))) = angular_component a" +proof- + have "((\ [^] (n::nat)) \ a) = (a \ (\ [^] (n::nat)))" + using assms Qp.domain_axioms domain_def + by (simp add: assms domain_def Qp.nonzero_memE cring.cring_simprules(14)) + then show ?thesis + using angular_component_p_nat_pow_factor[of a n] + by (simp add: assms) +qed + +lemma ac_p_nat_pow_factor_right: + assumes "m > 0" + assumes "a \ carrier Q\<^sub>p" + shows "ac m (a \ (\ [^] (n::nat))) = ac m a" + using assms angular_component_p_nat_pow_factor_right[of a n] + unfolding ac_def + by (metis Qp.integral Qp.l_null Qp.not_nonzero_memE Qp.not_nonzero_memI + angular_component_p_nat_pow_factor_right p_natpow_closed(1) p_natpow_closed(2)) + +lemma angular_component_p_int_pow_factor: + assumes "a \ carrier Q\<^sub>p" + shows "angular_component ((\ [^] (n::int)) \ a) = angular_component a" + by (metis Qp.integral_iff Qp.l_one Qp.nonzero_memI Qp.one_nonzero angular_component_mult + angular_component_one angular_component_p_int_pow assms p_intpow_closed(1) p_intpow_closed(2)) + +lemma ac_p_int_pow_factor: + assumes "a \ nonzero Q\<^sub>p" + shows "ac m ((\ [^] (n::int)) \ a) = ac m a" + apply(cases "m = 0") + apply (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.nonzero_memE(2) + ac_def angular_component_p_int_pow_factor assms p_intpow_closed(2)) + using assms angular_component_p_int_pow_factor[of a n] + by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.not_nonzero_memI ac_def + p_intpow_closed(2)) + +lemma angular_component_p_int_pow_factor_right: + assumes "a \ carrier Q\<^sub>p" + shows "angular_component (a \ (\ [^] (n::int))) = angular_component a" + using Qp.m_comm angular_component_p_int_pow_factor assms p_intpow_closed(1) by auto + +lemma ac_p_int_pow_factor_right: + assumes "a \ carrier Q\<^sub>p" + shows "ac m (a \ (\ [^] (n::int))) = ac m a" + using assms angular_component_p_int_pow_factor_right unfolding ac_def + using Qp.integral_iff Qp.nonzero_memE(2) p_intpow_closed(1) p_intpow_closed(2) by presburger + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\An Inverse for the inclusion map $\iota$\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition to_Zp where +"to_Zp a = (if (a \ \\<^sub>p) then (SOME x. x \ carrier Z\<^sub>p \ \ x = a) else \\<^bsub>Z\<^sub>p\<^esub>)" + +lemma to_Zp_closed: + assumes "a \ carrier Q\<^sub>p" + shows "to_Zp a \ carrier Z\<^sub>p" + apply(cases "a \ \\<^sub>p") + using assms unfolding to_Zp_def \_def + apply (smt image_iff tfl_some) + by simp + +lemma to_Zp_inc: + assumes "a \ \\<^sub>p" + shows "\ (to_Zp a) = a" +proof- + obtain c where c_def: "c = (SOME x. x \ carrier Z\<^sub>p \ \ x = a)" + by simp + have "(\ x. x \ carrier Z\<^sub>p \ \ x = a)" + using assms(1) + by blast + then have "c \ carrier Z\<^sub>p \ \ c = a" + using c_def + by (metis (mono_tags, lifting) tfl_some) + then show "\ (to_Zp a) = a" + using to_Zp_def c_def assms(1) + by auto +qed + +lemma inc_to_Zp: + assumes "b \ carrier Z\<^sub>p" + shows "to_Zp (\ b) = b" +proof- + have "\ (to_Zp (\ b)) = (\ b)" + using assms to_Zp_inc[of "\ b"] + by blast + then show ?thesis + using inc_inj2[of b "to_Zp (\ b)"] assms to_Zp_closed inc_closed + unfolding \_def Q\<^sub>p_def + by auto +qed + +lemma to_Zp_add: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "to_Zp (a \ b) = to_Zp a \\<^bsub>Z\<^sub>p\<^esub> (to_Zp b)" + by (metis (no_types, lifting) Zp.domain_axioms assms(1) assms(2) + cring.cring_simprules(1) domain.axioms(1) imageE inc_of_sum inc_to_Zp) + +lemma to_Zp_mult: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "to_Zp (a \ b) = to_Zp a \\<^bsub>Z\<^sub>p\<^esub> (to_Zp b)" +proof- + have "(a \ b) \ \\<^sub>p" + by (simp add: val_ring_subring assms(1) assms(2) subringE(6)) + then have + "\ ((to_Zp a) \\<^bsub>Z\<^sub>p\<^esub> (to_Zp) b) = ((\ (to_Zp a)) \ (\ (to_Zp b)))" + using assms(1) assms(2) inc_of_prod inc_to_Zp + by auto + then have "\ ((to_Zp a) \\<^bsub>Z\<^sub>p\<^esub> (to_Zp) b) = a \b" + by (simp add: assms(1) assms(2) to_Zp_inc) + then have "to_Zp ( \ ((to_Zp a) \\<^bsub>Z\<^sub>p\<^esub> (to_Zp) b)) = to_Zp (a \b)" + by simp + then show ?thesis + by (metis (no_types, opaque_lifting) Zp.domain_axioms val_ring_subring assms(1) assms(2) + cring.cring_simprules(5) domain.axioms(1) inc_to_Zp subringE(1) subset_iff to_Zp_closed) +qed + +lemma to_Zp_minus: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "to_Zp (a \ b) = to_Zp a \\<^bsub>Z\<^sub>p\<^esub> (to_Zp b)" + by (metis (no_types, lifting) Zp.domain_axioms assms(1) assms(2) cring_def domain.axioms(1) + image_iff inc_of_diff inc_to_Zp ring.ring_simprules(4)) + +lemma to_Zp_one: + shows "to_Zp \ = \\<^bsub>Z\<^sub>p\<^esub>" + using Zp_def Zp.one_closed \_def inc_of_one inc_to_Zp padic_integers_axioms + by fastforce + +lemma to_Zp_zero: + shows "to_Zp \ = \\<^bsub>Z\<^sub>p\<^esub>" + using Q\<^sub>p_def Zp_def Zp.domain_axioms \_def domain_frac.inc_inj1 inc_to_Zp + padic_integers_axioms to_Zp_def domain_frac_axioms by fastforce + +lemma to_Zp_ominus: + assumes "a \ \\<^sub>p" + shows "to_Zp (\ a) = \\<^bsub>Z\<^sub>p\<^esub> (to_Zp a)" +proof- + have "\a \ \\<^sub>p" + by (simp add: val_ring_subring assms subringE(5)) + then show ?thesis + by (metis (no_types, lifting) Zp.domain_axioms Zp.nonzero_one_closed assms + cring.cring_simprules(3) domain.axioms(1) frac_uminus image_iff inc_to_Zp local.inc_def) +qed + +lemma to_Zp_val: + assumes "a \ \\<^sub>p" + shows "val_Zp (to_Zp a) = val a" + by (metis assms imageE inc_to_Zp val_of_inc) + +lemma val_of_nat_inc: +"val ([(k::nat)]\\) \ 0" +proof- + have "[(k::nat)]\\ \ \\<^sub>p" + by (metis Zp.one_closed Zp_nat_mult_closed image_eqI inc_of_nat) + thus ?thesis + using val_ring_memE(1) by blast +qed + +lemma val_of_int_inc: +"val ([(k::int)]\\) \ 0" +proof- + have "\ ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = [k]\\" + using inc_of_int by blast + hence "[k]\\ \ \\<^sub>p" + using Zp.one_closed Zp_nat_mult_closed image_eqI inc_of_int + by blast + thus ?thesis + using val_ring_memE(1) by blast +qed + +lemma to_Zp_nat_inc: +"to_Zp ([(a::nat)]\\) = [a]\\<^bsub>Z\<^sub>p\<^esub>\ \<^bsub>Z\<^sub>p\<^esub>" + apply(induction a) + using Qp.nat_inc_zero Zp.nat_inc_zero to_Zp_zero apply presburger +proof- + fix a::nat assume A: "to_Zp ([a] \ \) = [a] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + have 0: "[(Suc a)]\\ = [a]\\ \ \" + using Qp.add.nat_pow_Suc by blast + have 1: "[a]\\ \ \\<^sub>p" + using Qp.nat_inc_closed val_of_nat_inc val_ring_memI by blast + have 2: "to_Zp ([Suc a] \ \) = (to_Zp ([a]\\)) \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + using A 1 unfolding 0 + by (metis Zp.cring_simprules(6) Zp.nat_inc_closed inc_of_nat inc_of_one inc_of_sum inc_to_Zp sum_closed) + show "to_Zp ([Suc a] \ \) = [Suc a] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" + unfolding 2 A + using add_comm nat_suc by presburger +qed + +lemma to_Zp_int_neg: +"to_Zp ([(-int (a::nat))]\\) = \\<^bsub>Z\<^sub>p\<^esub>([int a]\\<^bsub>Z\<^sub>p\<^esub>\ \<^bsub>Z\<^sub>p\<^esub>)" +proof- + have 0: "[(-int (a::nat))]\\ = \ ([int a]\\)" + using Qp.add.int_pow_neg by blast + have 1: "[int a]\\ \ \\<^sub>p" + using Qp.int_inc_closed val_of_int_inc val_ring_memI by blast + show ?thesis using 1 unfolding 0 + by (metis Zp.int_inc_closed inc_of_int inc_to_Zp to_Zp_ominus) +qed + +lemma(in ring) int_add_pow: +"[int n] \ \ = [n]\\" + unfolding add_pow_def + by (simp add: int_pow_int) + +lemma int_add_pow: +"[int n] \ \ = [n]\\" + unfolding add_pow_def + by (simp add: int_pow_int) + +lemma Zp_int_add_pow: +"[int n] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = [n]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>" + unfolding add_pow_def + by (simp add: int_pow_int) + +lemma to_Zp_int_inc: +"to_Zp ([(a::int)]\\) = ([a]\\<^bsub>Z\<^sub>p\<^esub>\ \<^bsub>Z\<^sub>p\<^esub>)" + apply(induction a) + unfolding int_add_pow Zp_int_add_pow + using to_Zp_nat_inc apply blast + unfolding to_Zp_int_neg using Zp.add.int_pow_neg Zp.one_closed + by presburger + +lemma to_Zp_nat_add_pow: + assumes "a \ \\<^sub>p" + shows "to_Zp ([(n::nat)]\a) = [n]\\<^bsub>Z\<^sub>p\<^esub> to_Zp a" + apply( induction n) + using Qp.nat_mult_zero Zp_nat_inc_zero to_Zp_zero apply presburger +proof- fix n::nat assume A: "to_Zp ([n] \ a) = [n] \\<^bsub>Z\<^sub>p\<^esub> to_Zp a" + have 0: "[Suc n] \ a = [n]\a \ a" + using Qp.add.nat_pow_Suc by blast + have 1: "to_Zp ([n] \ a \ a) = to_Zp ([n] \ a) \\<^bsub>Z\<^sub>p\<^esub> to_Zp a" + apply(rule to_Zp_add[of "[n]\a" a] ) + apply( induction n) + using Qp.nat_mult_zero zero_in_val_ring apply blast + unfolding Qp.add.nat_pow_Suc by(rule val_ring_add_closed, blast, rule assms, rule assms) + show "to_Zp ([Suc n] \ a) = [Suc n] \\<^bsub>Z\<^sub>p\<^esub> to_Zp a " + unfolding Qp.add.nat_pow_Suc Zp.add.nat_pow_Suc 1 A by blast +qed + +lemma val_ring_res: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + shows "to_Zp (a \ b) N = to_Zp a N \\<^bsub>Zp_res_ring N\<^esub> to_Zp b N" +proof- + have "to_Zp (a \ b) N = (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) N" + using assms to_Zp_minus by presburger + then show ?thesis using assms residue_of_diff to_Zp_closed val_ring_memE(2) + by (simp add: val_ring_memE(1)) +qed + +lemma res_diff_in_val_ring_imp_in_val_ring: + assumes "a \ \\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "a \ b \ \\<^sub>p" + shows "b \ \\<^sub>p" +proof- + have "b = a \ (a \ b)" + unfolding a_minus_def + using assms val_ring_memE(2)[of a] Qp.r_neg2 Q\<^sub>p_def Qp.minus_sum by auto + thus ?thesis using assms val_ring_minus_closed[of a "a \ b"] + by presburger +qed + +lemma(in padic_fields) equal_res_imp_res_diff_zero: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "to_Zp a N = to_Zp b N" + shows "to_Zp (a \ b) N = 0" + using val_ring_res[of a b] assms + by (metis res_diff_zero_fact' to_Zp_closed val_ring_memE(2)) + +lemma(in padic_fields) equal_res_imp_val_diff_bound: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "to_Zp a N = to_Zp b N" + shows "val (a \ b) \ N" + using assms equal_res_imp_res_diff_zero[of a b N] + by (metis to_Zp_closed to_Zp_minus to_Zp_val val_Zp_dist_def val_Zp_dist_res_eq2 val_ring_memE(2) val_ring_minus_closed) + +lemma(in padic_fields) equal_res_equal_val: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "val a < N" + assumes "to_Zp a N = to_Zp b N" + shows "val a = val b" +proof- + have "val (a \ b) \ N" + using assms equal_res_imp_val_diff_bound by blast + then have "val (a \ b) > val a" + using assms less_le_trans by blast + then show "val a = val b" + using assms by (meson ultrametric_equal_eq' val_ring_memE) +qed + +lemma(in padic_fields) val_ring_equal_res_imp_equal_val: + assumes "a \ \\<^sub>p" + assumes "b \ \\<^sub>p" + assumes "val a < eint N" + assumes "val b < eint N" + assumes "to_Zp a N = to_Zp b N" + shows "val a = val b" +proof- + have "val_Zp (to_Zp (a \ b)) \ N" + using assms val_ring_memE to_Zp_closed to_Zp_minus val_Zp_dist_def val_Zp_dist_res_eq2 by presburger + thus ?thesis + by (meson assms(1) assms(2) assms(3) assms(5) equal_res_equal_val) +qed + +end +end diff --git a/thys/Padic_Field/Padic_Semialgebraic_Function_Ring.thy b/thys/Padic_Field/Padic_Semialgebraic_Function_Ring.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Padic_Semialgebraic_Function_Ring.thy @@ -0,0 +1,6503 @@ +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\<^bsub>n\<^esub> Q\<^sub>p)" + using F.cring by blast + +lemma Qp_funs_is_monoid: +"monoid (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + using F.is_monoid by blast + +lemma Qp_funs_car_memE: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + shows "f \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ (carrier Q\<^sub>p)" + by (simp add: assms ring_pow_function_ring_car_memE(2)) + +lemma Qp_funs_car_memI: + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "\x. x \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ g x = undefined" + shows "g \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" +apply(rule Qp.function_ring_car_memI) + using assms apply blast + using assms by blast + +lemma Qp_funs_car_memI': + assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = g" + shows "g \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + assumes "g = (\ x \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)). f x)" + shows "g \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + apply(rule Qp_funs_car_memI) + using assms + apply (meson restrict_Pi_cancel) + by (metis assms(2) restrict_def) + +lemma Qp_funs_one: +"\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). \)" + unfolding function_ring_def function_one_def + by (meson monoid.select_convs(2)) + +lemma Qp_funs_zero: +"\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). \\<^bsub>Q\<^sub>p\<^esub>)" + unfolding function_ring_def function_zero_def + by (meson ring_record_simps(11)) + +lemma Qp_funs_add: + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier Q\<^sub>p" + assumes "g \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier Q\<^sub>p" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \\<^bsub>Q\<^sub>p\<^esub> g x" + using assms function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "Q\<^sub>p"] + unfolding function_add_def + by (metis (mono_tags, lifting) restrict_apply' ring_record_simps(12)) + +lemma Qp_funs_add': + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "g \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \\<^bsub>Q\<^sub>p\<^esub> g x" + using assms Qp_funs_add Qp_funs_car_memE + by blast + +lemma Qp_funs_add'': + assumes "f \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "g \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x \\<^bsub>Q\<^sub>p\<^esub> g x)" + unfolding function_ring_def function_add_def using ring_record_simps(12) + by metis + +lemma Qp_funs_add''': + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \\<^bsub>Q\<^sub>p\<^esub> g x" + using assms function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "Q\<^sub>p"] + unfolding function_add_def + by (metis (mono_tags, lifting) restrict_apply' ring_record_simps(12)) + +lemma Qp_funs_mult: + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier Q\<^sub>p" + assumes "g \ (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier Q\<^sub>p" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \ g x" + using assms function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "Q\<^sub>p"] + unfolding function_mult_def + by (metis (no_types, lifting) monoid.select_convs(1) restrict_apply') + +lemma Qp_funs_mult': + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "g \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \ g x" + using assms Qp_funs_mult Qp_funs_car_memE + by blast + +lemma Qp_funs_mult'': + assumes "f \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "g \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = f x \ g x" + using assms function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "Q\<^sub>p"] + unfolding function_mult_def + by (metis (mono_tags, lifting) monoid.select_convs(1) restrict_apply') + +lemma Qp_funs_a_inv: + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) 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\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). \ (f x))" +proof fix x + show "(\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x = (\x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). \ f x) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + 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 ("\\<^bsub>_\<^esub>_") where +"Qp_const n c \ constant_function (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) c" + +lemma Qp_constE: + assumes "c \ carrier Q\<^sub>p" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "\ x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + shows "f \ (Units (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). inv\<^bsub>Q\<^sub>p\<^esub> (f x))" +proof- + obtain g where g_def: "g = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). inv\<^bsub>Q\<^sub>p\<^esub> (f x))" + by blast + have g_closed: "g \ (carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) \ 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 \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>" + 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\<^bsub>n\<^esub> Q\<^sub>p))" + using comm_monoid.UnitsI[of "Fun\<^bsub>n\<^esub> Q\<^sub>p"] assms(1) g_closed local.F.comm_monoid_axioms by presburger + have "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = g" + using g_def g_closed 0 cring.invI[of "Fun\<^bsub>n\<^esub> Q\<^sub>p"] Qp_funs_is_cring assms(1) + by presburger + show "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). inv\<^bsub>Q\<^sub>p\<^esub> (f x))" + using assms g_def 0 \inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = g\ + by blast +qed + +lemma Qp_funs_Units_memE: + assumes "f \ (Units (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>" + "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>" + "\ x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + using monoid.Units_r_inv[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" f ] assms Qp_funs_is_monoid + apply blast + using monoid.Units_l_inv[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" f ] assms Qp_funs_is_monoid + apply blast +proof- + obtain g where g_def: "g = inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f" + by blast + show "\ x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + have "f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>" + using assms g_def Qp_funs_is_monoid + \\Group.monoid (Fun\<^bsub>n\<^esub> Q\<^sub>p); f \ Units (Fun\<^bsub>n\<^esub> Q\<^sub>p)\ \ f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>\ + 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 Q\<^sub>p" + using g_def A assms local.function_ring_car_closed by auto + then show "f x \ \\<^bsub>Q\<^sub>p\<^esub>" + using 0 + by (metis Qp.l_null local.one_neq_zero) + qed +qed + +lemma Qp_funs_m_inv: + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ (Units (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "(inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x = inv\<^bsub>Q\<^sub>p\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier Q\<^sub>p. is_semialg_function n f \ f = restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))}" + +lemma semialg_functions_memE: + assumes "f \ semialg_functions n" + shows "is_semialg_function n f" + "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + 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\<^bsub>n\<^esub> Q\<^sub>p)" + using semialg_functions_memE + by blast + +lemma semialg_functions_memI: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)))" +proof(rule is_semialg_functionI) + show 0: "restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>))) k S)" + proof- fix k S assume A: "S \ semialg_sets (1 + k)" + have "(partial_pullback n (restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) 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 (Q\<^sub>p\<^bsup>n\<^esup>))) 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 (Q\<^sub>p\<^bsup>n\<^esup>))) \ 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 Q\<^sub>p" + shows "is_semialg_function n (constant_function (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) a)" +proof- + have "(constant_function (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) a) = restrict (Qp_ev (Qp.indexed_const a)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + 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 Q\<^sub>p" + 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: +"\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> = Qp_const n \" + unfolding constant_function_def function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] function_one_def + by simp + +lemma function_zero_as_constant: +"\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> = Qp_const n \\<^bsub>Q\<^sub>p\<^esub>" + unfolding constant_function_def function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] function_zero_def + by simp + +lemma sum_in_semialg_functions: + assumes "f \ semialg_functions n" + assumes "g \ semialg_functions n" + shows "f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g \ semialg_functions n" +proof- + have 0:"f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g = restrict (function_tuple_comp Q\<^sub>p [f,g] Qp_add_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + proof fix x + show "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = restrict (function_tuple_comp Q\<^sub>p [f, g] Qp_add_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True + have " restrict (function_tuple_comp Q\<^sub>p [f, g] Qp_add_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) 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 Q\<^sub>p [f, g] Qp_add_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x = f x \\<^bsub>Q\<^sub>p\<^esub> g 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 (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] + 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 \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = undefined" + using function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] 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 Q\<^sub>p [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 \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g \ semialg_functions n" +proof- + have 0:"f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g = restrict (function_tuple_comp Q\<^sub>p [f,g] Qp_mult_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + proof fix x + show "(f \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = restrict (function_tuple_comp Q\<^sub>p [f, g] Qp_mult_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True + have " restrict (function_tuple_comp Q\<^sub>p [f, g] Qp_mult_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) 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 Q\<^sub>p [f, g] Qp_mult_fun) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) 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 (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] + 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 \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> g) x = undefined" + using function_ring_def[of "carrier (Q\<^sub>p\<^bsup>n\<^esup>)" Q\<^sub>p] 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 Q\<^sub>p [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 (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + shows "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \ semialg_functions n " +proof- + have 0: "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = restrict (function_tuple_comp Q\<^sub>p [f] Qp_invert) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + proof fix x + show "(inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x = restrict (function_tuple_comp Q\<^sub>p [f] Qp_invert) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True + have "(function_tuple_comp Q\<^sub>p [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 Q\<^sub>p [f] Qp_invert) x = inv\<^bsub>Q\<^sub>p\<^esub> (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\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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 Q\<^sub>p [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 "\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \ semialg_functions n" +proof- + have "\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f = (Qp_const n (\ \)) \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f" + proof fix x + show "(\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x = (Qp_const n (\ \) \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True + have 0: "(\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) x = \ (f x)" + using Qp_funs_a_inv semialg_functions_memE True assms(1) by blast + have 1: "(Qp_const n (\ \) \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f) 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 Q\<^sub>p" + 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 "(\\<^bsub>function_ring (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) Q\<^sub>p\<^esub> f) 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\<^bsub>n\<^esub> Q\<^sub>p)" + 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\<^bsub>n\<^esub> Q\<^sub>p)" + using semialg_functions_subring cring.subcringI' + using Qp_funs_is_cring by blast + +definition SA where +"SA n = (Fun\<^bsub>n\<^esub> Q\<^sub>p)\ carrier := semialg_functions n\" + +lemma SA_is_ring: + shows "ring (SA n)" +proof- + have "ring (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + by (simp add: Qp_funs_is_cring cring.axioms(1)) + then show ?thesis + unfolding SA_def + using ring.subring_is_ring[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" "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\<^bsub>n\<^esub> Q\<^sub>p" "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\<^bsub>n\<^esub> Q\<^sub>p)" + by (simp add: SA_car semialg_functions_in_Qp_funs) + +lemma SA_car_memI: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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\<^bsub>n\<^esub> Q\<^sub>p)" + "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + 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: +"(\\<^bsub>SA n\<^esub>) = (\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>)" + unfolding SA_def + by simp + +lemma SA_times: +"(\\<^bsub>SA n\<^esub>) = (\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>)" + unfolding SA_def + by simp + +lemma SA_one: +"(\\<^bsub>SA n\<^esub>) = (\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>)" + unfolding SA_def + by simp + +lemma SA_zero: +"(\\<^bsub>SA n\<^esub>) = (\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>)" + unfolding SA_def + by simp + +lemma SA_zero_is_function_ring: +"(Fun\<^bsub>0\<^esub> Q\<^sub>p) = SA 0" +proof- + have 0: "carrier (Fun\<^bsub>0\<^esub> Q\<^sub>p) = carrier (SA 0)" + proof + show "carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p) \ carrier (SA 0)" + proof fix f assume A0: "f \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)" + show "f \ carrier (SA 0)" + proof(rule SA_car_memI) + show "f \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)" + using A0 by blast + show "is_semialg_function 0 f" + proof(rule is_semialg_functionI) + show "f \ carrier (Q\<^sub>p\<^bsup>0\<^esup>) \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" + using Qp_zero_carrier by blast + have 1: "(partial_pullback 0 f k S) = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). a#x \ S}" + proof + show "partial_pullback 0 f k S \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>))) \ S" + proof + show "cartesian_product {[a]} (partial_pullback 0 f k S) \ cartesian_product {[a]} (carrier (Q\<^sub>p\<^bsup>k\<^esup>)) \ 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 (Q\<^sub>p\<^bsup>k\<^esup>))" + unfolding 21 apply(rule cartesian_product_memI'[of _ Q\<^sub>p 1 _ k]) + using a_def apply (metis function_ring_car_closed Qp_zero_carrier \f \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)\ 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 (Q\<^sub>p\<^bsup>k\<^esup>)) \ S" + using 20 by blast + qed + show " cartesian_product {[a]} (carrier (Q\<^sub>p\<^bsup>k\<^esup>)) \ S \ cartesian_product {[a]} (partial_pullback 0 f k S)" + proof fix x assume A: "x \ cartesian_product {[a]} (carrier (Q\<^sub>p\<^bsup>k\<^esup>)) \ S" + then obtain y where y_def: "x = a#y \ y \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>k\<^esup>)" + 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]}" Q\<^sub>p 1 "partial_pullback 0 f k S" k "[a]" y ]) + apply (metis function_ring_car_closed Qp_zero_carrier \f \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)\ 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 Q\<^sub>p" + using a_def Qp.function_ring_car_memE 0 \f \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)\ by blast + hence "[a] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" + 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 (Q\<^sub>p\<^bsup>k\<^esup>)))" + 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 (Q\<^sub>p\<^bsup>1\<^esup>)" + 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 (Q\<^sub>p\<^bsup>0\<^esup>)) Q\<^sub>p)" + using SA_car_in_Qp_funs_car by blast + qed + then have 1: "semialg_functions 0 = carrier (Fun\<^bsub>0\<^esub> Q\<^sub>p)" + unfolding 0 SA_def by auto + show ?thesis unfolding SA_def 1 by auto +qed + +lemma constant_fun_closed: + assumes "c \ carrier Q\<^sub>p" + shows "constant_function (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) c \ carrier (SA m)" + using constant_function_in_semialg_functions SA_car assms by blast + +lemma SA_0_car_memI: + assumes "\ \ carrier (Q\<^sub>p\<^bsup>0\<^esup>) \ carrier Q\<^sub>p" + assumes "\x. x \ carrier (Q\<^sub>p\<^bsup>0\<^esup>) \ \ x = undefined" + shows "\ \ carrier (SA 0)" +proof- + have 0: "carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" + by (simp add: Qp_zero_carrier) + obtain c where c_def: "\ [] = c" + by blast + have 1: "\ = constant_function (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) c" + unfolding constant_function_def restrict_def + using assms c_def unfolding 0 + by (metis empty_iff insert_iff) + have 2: "c \ carrier Q\<^sub>p" + 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 Q\<^sub>p. a = Qp_const 0 c" +proof- + obtain c where c_def: "c = a []" + by blast + have car_zero: "carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" + using Qp_zero_carrier by blast + have 0: "a = constant_function (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) c" + proof fix x + show " a x = constant_function (carrier (Q\<^sub>p\<^bsup>0\<^esup>)) c x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>0\<^esup>)") + 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 Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "\ \<^bsub>SA n\<^esub> a = \" + using function_zero_eval SA_zero assms by presburger + +lemma SA_oneE: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "\ \<^bsub>SA n\<^esub> a = \" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>SA n\<^esub> g) x = f x \\<^bsub>Q\<^sub>p\<^esub> g x" + using Qp_funs_add''' SA_plus assms by presburger + +lemma SA_add': + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>SA n\<^esub> g) x = undefined" +proof- + have "(f \\<^bsub>SA n\<^esub> g) x = function_add (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>SA n\<^esub> g) x = f x \ g x" + using Qp_funs_mult''' SA_times assms by presburger + +lemma SA_mult': + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>SA n\<^esub> g) x = undefined" +proof- + have "(f \\<^bsub>SA n\<^esub> g) x = function_mult (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(\\<^bsub>SA n\<^esub> f) x = \ (f x)" +proof- + have "f \\<^bsub>SA n\<^esub> (\\<^bsub>SA n\<^esub> f) = \ \<^bsub>SA n\<^esub>" + using assms SA_is_cring cring.cring_simprules(17) by metis + have "(f \\<^bsub>SA n\<^esub> (\\<^bsub>SA n\<^esub> f)) x = \ \<^bsub>SA n\<^esub> x" + using assms \f \\<^bsub>SA n\<^esub> \\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>\ by presburger + then have "(f x) \ (\\<^bsub>SA n\<^esub> f) 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(\\<^bsub>SA n\<^esub> f) x = \ (f x)" +proof- + have "f \\<^bsub>SA n\<^esub> (\\<^bsub>SA n\<^esub> f) = \ \<^bsub>SA n\<^esub>" + using assms SA_is_cring cring.cring_simprules(17) by metis + have "(f \\<^bsub>SA n\<^esub> (\\<^bsub>SA n\<^esub> f)) x = \ \<^bsub>SA n\<^esub> x" + using assms \f \\<^bsub>SA n\<^esub> \\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>\ by presburger + then have "(f x) \ (\\<^bsub>SA n\<^esub> f) x = \" + by (metis function_zero_eval SA_add SA_zero assms) + then show ?thesis + by (metis (no_types, lifting) PiE Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f [^]\<^bsub>SA n\<^esub> (k::nat)) x = (f x) [^]\<^bsub>Q\<^sub>p\<^esub> k" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f [^]\<^bsub>SA n\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)) \\<^bsub>SA n\<^esub> restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = f \\<^bsub>SA n\<^esub> g " +proof fix x + show "(restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \\<^bsub>SA n\<^esub> restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) x = (f \\<^bsub>SA n\<^esub> g) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + 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 (Q\<^sub>p\<^bsup>n\<^esup>)) \\<^bsub>SA n\<^esub> restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = f \\<^bsub>SA n\<^esub> g " +proof fix x + show "(restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \\<^bsub>SA n\<^esub> restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) x = (f \\<^bsub>SA n\<^esub> g) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + 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 \\<^bsub>SA n\<^esub> g \ 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 \\<^bsub>SA n\<^esub> g \ 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 \\<^bsub>SA n\<^esub> g \ 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 \\<^bsub>SA n\<^esub> g \ 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 \\<^bsub>SA n\<^esub> g \ 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 \\<^bsub>SA n\<^esub> g \ 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 [^]\<^bsub>SA n\<^esub> (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 \\<^bsub>SA n\<^esub> g) \ 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)]\\<^bsub>R\<^esub>b) \ 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\<^bsub>n\<^esub> Q\<^sub>p)" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "([(m::nat)]\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> b) a = [m]\( b a)" +proof- + have 0: "b a \ carrier Q\<^sub>p" + using Qp.function_ring_car_mem_closed assms by fastforce + have 1: "ring (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + using function_ring_is_ring by blast + show ?thesis + proof(induction m) + case 0 + have "([(0::nat)] \\<^bsub>function_ring (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) Q\<^sub>p\<^esub> b) = \\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>" + using 1 ring.add_pow_zero[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "([(m::nat)]\\<^bsub>SA n\<^esub> b) 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\<^bsub>n\<^esub> Q\<^sub>p" 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\<^bsub>n\<^esub> Q\<^sub>p)" + assumes "is_semialg_function n f" + shows "f \ Units (SA n)" +proof- + have 0: "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + by (meson Qp_funs_is_monoid assms(1) monoid.Units_closed) + have 1: "inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \ semialg_functions n" + using monoid.Units_closed[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" 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 \\<^bsub>SA n\<^esub> (inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f ) = \\<^bsub>SA n\<^esub>" + using Qp_funs_Units_memE(1)[of f n] SA_one SA_times assms(1) + by presburger + then have 3: "(inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f ) \\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>" + 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\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f \ 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 \\<^bsub>SA n\<^esub> inv\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>" + "inv\<^bsub>SA n\<^esub> f \\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>" + 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 "inv\<^bsub>SA n\<^esub> f \ 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\<^bsub>n\<^esub> Q\<^sub>p))" +proof- + have 0: "f \\<^bsub>SA n\<^esub> inv\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>" + "inv\<^bsub>SA n\<^esub> f \\<^bsub>SA n\<^esub> f = \\<^bsub>SA n\<^esub>" + using R.Units_r_inv assms apply blast + using R.Units_l_inv assms by blast + have 1: "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + using assms + by (metis SA_car SA_is_monoid monoid.Units_closed semialg_functions_memE(2)) + have 2: "inv\<^bsub>SA n\<^esub> f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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 "inv\<^bsub>SA n\<^esub> f = inv\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub> f" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) \ f x \ \\<^bsub>Q\<^sub>p\<^esub>" + using assms Qp_funs_Units_memE[of f n] SA_Units_Qp_funs_Units + by blast + +lemma Qp_n_nonempty: + shows "carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ {}" + apply(induction n) + apply (simp add: Qp_zero_carrier) + using cartesian_power_cons[of _ Q\<^sub>p _ \] Qp.one_closed + by (metis Suc_eq_plus1 all_not_in_conv cartesian_power_cons empty_iff) + +lemma SA_one_not_zero: + shows "\\<^bsub>SA n\<^esub> \ \\<^bsub> SA n\<^esub>" +proof- + obtain a where a_def: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using Qp_n_nonempty by blast + have "\\<^bsub>SA n\<^esub> a \ \\<^bsub> SA n\<^esub> a" + 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 \ \\<^bsub> SA n\<^esub>" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "f x \ nonzero Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "f x \ carrier Q\<^sub>p" + using assms SA_car_memE(3) by blast + +lemma SA_Units_closed_fun: + assumes "f \ Units (SA m)" + assumes "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "f x \ carrier Q\<^sub>p" + using SA_Units_closed SA_car_closed assms by blast + +lemma SA_inv_eval: + assumes "f \ Units (SA n)" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(inv\<^bsub>SA n\<^esub> f) x = inv (f x)" +proof- + have "f \\<^bsub>SA n\<^esub> (inv\<^bsub>SA n\<^esub> f) = \ \<^bsub>SA n\<^esub>" + using assms SA_is_cring SA_Units_memE(1) by blast + hence "(f \\<^bsub>SA n\<^esub> (inv\<^bsub>SA n\<^esub> f)) x = \ \<^bsub>SA n\<^esub> x" + using assms by presburger + then have "(f x) \ (inv\<^bsub>SA n\<^esub> f) x = \" + by (metis function_one_eval SA_mult SA_one assms) + then show ?thesis + by (metis Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(h \\<^bsub>SA n\<^esub> (inv\<^bsub>SA n\<^esub> f)) 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "(f[^]\<^bsub>SA m\<^esub>(i::int)) x = (f x)[^]i" +proof(induction i) + case (nonneg n) + have 0: "(f [^]\<^bsub>SA m\<^esub> int n) = (f [^]\<^bsub>SA m\<^esub> n)" + 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 [^]\<^bsub>SA m\<^esub> - int (Suc n)) = inv \<^bsub>SA m\<^esub>(f [^]\<^bsub>SA m\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)) \ carrier (SA n)" + using assms SA_car restrict_in_semialg_functions + by blast + +lemma SA_smult: +"(\\<^bsub>SA n\<^esub>) = (\\<^bsub>Fun\<^bsub>n\<^esub> Q\<^sub>p\<^esub>)" + unfolding SA_def by auto + +lemma SA_smult_formula: + assumes "h \ carrier (SA n)" + assumes "q \ carrier Q\<^sub>p" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(q \\<^bsub>SA n\<^esub> h) 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 Q\<^sub>p" + shows "q \\<^bsub>SA n\<^esub> h \ carrier (SA n)" +proof- + obtain g where g_def: "g = \\<^bsub>n\<^esub> q" + 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 \\<^bsub>SA n\<^esub> h = g \\<^bsub>SA n\<^esub> h" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "val ((\ \\<^bsub>SA m\<^esub>f) x) = val (f x) + 1" +proof- + have 0: "(\\\<^bsub>SA m\<^esub>f) 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 Q\<^sub>p" + 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 \ \\<^bsub>SA m\<^esub>" + assumes "n > 0" + shows "[(n::nat)]\\<^bsub>SA m\<^esub>f \ \\<^bsub>SA m\<^esub>" +proof assume A: "[n] \\<^bsub>SA m\<^esub> f = \\<^bsub>SA m\<^esub>" + obtain x where x_def: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ f x \ \" + using assms + by (metis function_ring_car_eqI R.cring_simprules(2) SA_car_memE(2) SA_zeroE) + have 0: "([(n::nat)]\\<^bsub>SA m\<^esub>f) 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 Q\<^sub>p" + 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) \\<^bsub>m+l\<^esub> S" + +lemma map_partial_pullback_memE: + assumes "as \ map_partial_pullback m f l S" + shows "as \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" "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 (Q\<^sub>p\<^bsup>m+l\<^esup>)" + using map_partial_pullback_memE(1) by blast + +lemma map_partial_pullback_memI: + assumes "as \ carrier (Q\<^sub>p\<^bsup>m+k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "map_partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) - S) = carrier (Q\<^sub>p\<^bsup>m+l\<^esup>) - (map_partial_pullback m f l S) " + apply(rule equalityI) + using map_partial_pullback_def[of m f l "(carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) - 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 (Q\<^sub>p\<^bsup>m+l\<^esup>) - map_partial_pullback m f l S" + show " x \ map_partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) - S) " + apply(rule map_partial_pullback_memI) + using A + apply blast + proof + have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" + by (meson A DiffD1 cartesian_power_drop) + have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using A + by (meson DiffD1 take_closed le_add1) + show "f (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) " + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "map_partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" + apply(rule equalityI) + using map_partial_pullback_memE(1) apply blast +proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" + show "x \ map_partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>n+l\<^esup>))" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + using A take_closed assms + by (meson Pi_mem le_add1) + then show "f (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" + using cartesian_power_drop[of x m l] A cartesian_power_concat(1)[of _ Q\<^sub>p n _ l] + by blast +qed +qed + +definition is_semialg_map where +"is_semialg_map m n f = (f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ + (\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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using is_semialg_map_def assms by blast + +lemma is_semialg_map_closed': + assumes "is_semialg_map m n f" "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "f x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 \\<^bsub>m+k\<^esub> S)" + using assms is_semialg_mapE unfolding map_partial_pullback_def + by blast + +lemma is_semialg_mapI: + assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "\k S. S \ semialg_sets (n + k) \ is_semialgebraic (m + k) (map_partial_image m f \\<^bsub>m+k\<^esub> S)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" "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 (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "\k S. S \ basic_semialgs (n + k) \ is_semialgebraic (m + k) (map_partial_image m f \\<^bsub>m+k\<^esub> S)" + 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 (Q\<^sub>p\<^bsup>n + k\<^esup>) = + 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 (Q\<^sub>p\<^bsup>k+l\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m+l\<^esup>)" + proof- + have 1: "g (take k x) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m + l\<^esup>)\ + 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 (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>k+l\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m+l\<^esup>)" + proof- + have 1: "g (take k x) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m+l\<^esup>)\ + 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 (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" + proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + show " (f \ g) x \ carrier Q\<^sub>p" + 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 \\<^bsub>k\<^esub> S)" +proof- + have "g \\<^bsub>k\<^esub> S = map_partial_pullback k g 0 S" + proof + show "g \\<^bsub>k\<^esub> S \ map_partial_pullback k g 0 S" + proof fix x assume A: "x \ g \\<^bsub>k\<^esub> S" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 \\<^bsub>k\<^esub> S" + 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 (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 \\<^bsub>k\<^esub> S" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + shows "is_semialg_map n m g" +proof(rule is_semialg_mapI) + have 0: "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using assms(1) is_semialg_map_closed + by blast + show "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" then show " g x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. g (take n xs) @ drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + proof fix x assume A: "x \ (\xs. f (take n xs) @ drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + using assms + by (metis (no_types, lifting) A Int_iff vimageE vimageI) + qed + show " (\xs. g (take n xs) @ drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. f (take n xs) @ drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + proof fix x assume A: "x \ (\xs. g (take n xs) @ drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 Q\<^sub>p 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 Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict (Qp_ev (pvar Q\<^sub>p k)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + using assms by (metis (no_types, lifting) eval_pvar restrict_ext) + have 1: "is_semialg_function n (Qp_ev (pvar Q\<^sub>p k))" + using pvar_closed assms poly_is_semialg[of "pvar Q\<^sub>p k"] by blast + show ?thesis + using 0 1 semialg_function_on_carrier[of n "Qp_ev (pvar Q\<^sub>p k)" "(\as. as!k)"] + by presburger +qed + +definition Qp_ith where +"Qp_ith m i = (\ x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>)) Q\<^sub>p)" + apply(rule Qp.function_ring_car_memI[of "carrier(Q\<^sub>p\<^bsup>m\<^esup>)"]) + using assms cartesian_power_car_memE'[of _ Q\<^sub>p m i] unfolding Qp_ith_def + apply (metis restrict_apply) + unfolding restrict_def by meson + have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ Qp_ith m i x = eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 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 Q\<^sub>p (pvar Q\<^sub>p 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 Q\<^sub>p (pvar Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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..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 Q\<^sub>p n fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict (take k) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + unfolding function_tuple_eval_def + proof fix x + show " (\x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). map (\f. f x) fs) x = restrict (take k) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + 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..i. \as\carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 Q\<^sub>p 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 Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n+k\<^esup>). as!i)) [k.. set fs" + then obtain i where i_def: "i \ set [k.. f = (\as \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). 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..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 Q\<^sub>p (n+k) fs) (carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)) = restrict (drop k) (carrier (Q\<^sub>p\<^bsup>n+k\<^esup>))" + unfolding function_tuple_eval_def + proof fix x + show " (\x\carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). map (\f. f x) fs) x = restrict (drop k) (carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)) x" + proof(cases "x\carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)") + 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..f. f x) fs) ! i = (fs!i) x" + using A by (meson nth_map) + have "[k..i. \as\carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). as ! i) [k..i. \as\carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). as ! i) ([k..i. \as\carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). as ! i)"] + by linarith + then have 1: "fs!i = (\as \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>). as!(i+k))" + using fs_def A \[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..p (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 Q\<^sub>p (k + n) fs" "drop k"] + by (metis add.commute) +qed + +lemma project_at_indices_is_semialg_map: + assumes "S \ {..\<^bsub>S\<^esub>" +proof- + obtain k where k_def: "k = card S" + by blast + have 0: "card {.. 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 (Q\<^sub>p\<^bsup>n\<^esup>). as!(nth_elem S i))) [0.. set fs" + then obtain i where i_def: "i \ set [0.. f = (\as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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.. 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 Q\<^sub>p n fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict (\\<^bsub>S\<^esub>) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + unfolding function_tuple_eval_def + proof fix x + show " (\x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). map (\f. f x) fs) x = restrict (\\<^bsub>S\<^esub>) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True + have "(map (\f. f x) fs) = \\<^bsub>S\<^esub> x" + proof- + have "\i. i < k \ (map (\f. f x) fs) ! i = (\\<^bsub>S\<^esub> x) ! 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..as\carrier (Q\<^sub>p\<^bsup>n\<^esup>). as ! nth_elem S ([0..as\carrier (Q\<^sub>p\<^bsup>n\<^esup>). as ! nth_elem S ([0..f. f x) fs ! i = x! (nth_elem S i)" + using fs_def T0 A nth_map[of i "[0..i. \as\carrier (Q\<^sub>p\<^bsup>n\<^esup>). as ! nth_elem S i)"] + by (metis "4" T6 length_map) + show "map (\f. f x) fs ! i = \\<^bsub>S\<^esub> x ! 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 (\\<^bsub>S\<^esub> x)" + 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 (\\<^bsub>S\<^esub> x) = 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 = \\<^bsub>S\<^esub> x ! 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 Q\<^sub>p 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.. {..Coordinate functions are semialgebraic maps.\ + +lemma coord_fun_is_SA: + assumes "is_semialg_map n m g" + assumes "i < m" + shows "coord_fun Q\<^sub>p n g i \ carrier (SA n)" +proof(rule SA_car_memI) + show "coord_fun Q\<^sub>p n g i \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) Q\<^sub>p)" + 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 Q\<^sub>p 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 Q\<^sub>p n g i) = restrict ((\x. x ! i) \ g) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) " + unfolding coord_fun_def using comp_apply + by metis + show ?thesis + using semialg_function_on_carrier[of n "((\x. x ! i) \ g)" "coord_fun Q\<^sub>p 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 Q\<^sub>p n g) [0.. set (map (coord_fun Q\<^sub>p n g) [0.. f = coord_fun Q\<^sub>p n g i" + using set_map[of "coord_fun Q\<^sub>p n g" "[0.. carrier (SA n)" + shows "is_semialg_map n (Suc m) (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x # g x)" +proof- + obtain Fs where Fs_def: "Fs = f # (map (coord_fun Q\<^sub>p n g) [0..p n g) [0..p n g) [0.. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x # g x) = (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). function_tuple_eval Q\<^sub>p n Fs x)" + proof(rule ext) fix x show "(\x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x # g x) x = restrict (function_tuple_eval Q\<^sub>p n Fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + case True then have T0: "(\x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x # g x) x = f x # g x" + by (meson restrict_apply') + have T1: "restrict (function_tuple_eval Q\<^sub>p n Fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x = function_tuple_eval Q\<^sub>p n Fs x" + using True by (meson restrict_apply') + hence T2: "restrict (function_tuple_eval Q\<^sub>p n Fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x = f x # (function_tuple_eval Q\<^sub>p n (map (coord_fun Q\<^sub>p n g) [0.. carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using assms True is_semialg_map_closed by blast + have "length [0..p n g) [0..p n g)" "[0..p n (map (coord_fun Q\<^sub>p n g) [0..p n (map (coord_fun Q\<^sub>p n g) [0..p m] T5 T4 T3 True + nth_map[of _ "(map (\i. \x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). g x ! i) [0..f. f x)"] + nth_map[of _ "[0..i. \x\carrier (Q\<^sub>p\<^bsup>n\<^esup>). g x ! i)"] + unfolding function_tuple_eval_def coord_fun_def restrict_def + by (metis (no_types, lifting) \length [0.. map_nth nth_map) + hence T7: "restrict (function_tuple_eval Q\<^sub>p n Fs) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) 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 Q\<^sub>p n g)" "[0..p n g) [0.. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). function_tuple_eval Q\<^sub>p n Fs x)" + apply(rule semialg_map_on_carrier[of _ _ "function_tuple_eval Q\<^sub>p 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 (Q\<^sub>p\<^bsup>n\<^esup>) (Q\<^sub>p\<^bsup>m\<^esup>)}" + +lemma semialg_mapsE: + assumes "f \ (semialg_maps n m)" + shows "is_semialg_map n m f" + "f \ struct_maps (Q\<^sub>p\<^bsup>n\<^esup>) (Q\<^sub>p\<^bsup>m\<^esup>)" + "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>))" + +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 (Q\<^sub>p\<^bsup>m\<^esup>))" + +definition drop_apply where +"drop_apply m n f = restrict (f \ drop n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" + +lemma take_apply_closed: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + assumes "k \ n" + shows "take_apply k n f \ carrier (Fun\<^bsub>k\<^esub> Q\<^sub>p)" +proof(rule Qp.function_ring_car_memI) + show "\a. a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ take_apply k n f a \ carrier Q\<^sub>p" + proof- fix a assume A: "a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" show "take_apply k n f a \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>k\<^esup>) \ 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\<^bsub>k - n\<^esub> Q\<^sub>p)" + assumes "k \ n" + shows "drop_apply k n f \ carrier (Fun\<^bsub>k\<^esub> Q\<^sub>p)" +proof(rule Qp.function_ring_car_memI) + show " \a. a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ drop_apply k n f a \ carrier Q\<^sub>p" + proof- fix a assume A: "a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" show "drop_apply k n f a \ carrier Q\<^sub>p" + using A assms comp_apply[of f "drop n" a] Qp.function_ring_car_memE[of f ] drop_closed[of n k a Q\<^sub>p] + 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 (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + shows "take_apply (n+k) n f (a@b) = f a" +proof- + have "a@b \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + shows "drop_apply (n+k) n f (a@b) = f b" +proof- + have "a@b \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 \\<^bsub>SA n\<^esub> g) = drop_apply (n+k) k f \\<^bsub>SA (n + k)\<^esub> 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + then obtain b c where bc_def: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ a = b@c" + by (metis (no_types, lifting) add.commute cartesian_power_decomp) + have 0: "drop_apply (n + k) k (f \\<^bsub>SA n\<^esub> g) a = f c \ g c" + using assms bc_def drop_apply_apply[of "f \\<^bsub>SA n\<^esub> g" 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 \\<^bsub>SA n\<^esub> g) a = (drop_apply (n + k) k f \\<^bsub>SA (n + k)\<^esub> 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 \ \<^bsub>SA n\<^esub> g) = drop_apply (n+k) k f \\<^bsub>SA (n + k)\<^esub> 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + then obtain b c where bc_def: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ a = b@c" + by (metis (no_types, lifting) add.commute cartesian_power_decomp) + have 0: "drop_apply (n + k) k (f \\<^bsub>SA n\<^esub> g) a = f c \ g c" + using assms bc_def drop_apply_apply[of "f \\<^bsub>SA n\<^esub> g" 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 \\<^bsub>SA n\<^esub> g) a = (drop_apply (n + k) k f \\<^bsub>SA (n + k)\<^esub> 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 \\<^bsub>SA n\<^esub> = \\<^bsub>SA (n+k)\<^esub>" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + show "drop_apply (n + k) k \\<^bsub>SA n\<^esub> a = \\<^bsub>SA (n + k)\<^esub> a" + unfolding drop_apply_def restrict_def + using SA_one[of "n+k"] SA_one[of n] comp_apply[of "\\<^bsub>SA n\<^esub>" "drop k" a] drop_closed[of k "n+k" a Q\<^sub>p] + 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 \\<^bsub>SA k\<^esub> g) = take_apply (n+k) k f \\<^bsub>SA (n + k)\<^esub> 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + then obtain b c where bc_def: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ a = b@c" + by (metis (no_types, lifting) add.commute cartesian_power_decomp) + hence 0: "take_apply (n + k) k (f \\<^bsub>SA k\<^esub> g) a = f b \ g b" + using assms bc_def take_apply_apply[of "f \\<^bsub>SA k\<^esub> g" 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 \\<^bsub>SA k\<^esub> g) a = (take_apply (n + k) k f \\<^bsub>SA (n + k)\<^esub> 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 \\<^bsub>SA k\<^esub> g) = take_apply (n+k) k f \\<^bsub>SA (n + k)\<^esub> 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + then obtain b c where bc_def: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ a = b@c" + by (metis (no_types, lifting) add.commute cartesian_power_decomp) + hence 0: "take_apply (n + k) k (f \\<^bsub>SA k\<^esub> g) a = f b \ g b" + using assms bc_def take_apply_apply[of "f \\<^bsub>SA k\<^esub> g" 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 \\<^bsub>SA k\<^esub> g) a = (take_apply (n + k) k f \\<^bsub>SA (n + k)\<^esub> 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 \\<^bsub>SA k\<^esub> = \\<^bsub>SA (n+k)\<^esub>" + 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 (Q\<^sub>p\<^bsup>n+k\<^esup>)" + show "take_apply (n + k) k \\<^bsub>SA k\<^esub> a = \\<^bsub>SA (n + k)\<^esub> a" + unfolding take_apply_def restrict_def + using SA_one[of "n+k"] SA_one[of k] comp_apply[of "\\<^bsub>SA k\<^esub>" "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 (Q\<^sub>p\<^bsup>n + k\<^esup>) \ drop_apply (n + k) k f x \ \" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n + k\<^esup>) \ take_apply (n + k) k f x \ \" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 \\<^bsub>n\<^esub> S)" +proof- + have 0: "is_semialgebraic 1 (to_R1 ` S)" + using assms is_univ_semialgebraicE by blast + have 1: "h \\<^bsub>n\<^esub> S = partial_pullback n h (0::nat) (to_R1 ` S)" + proof show "h \\<^bsub>n\<^esub> S \ partial_pullback n h 0 ((\a. [a]) ` S)" + proof fix x assume A: "x \ h \\<^bsub>n\<^esub> S" + then have 0: "h x \ S" by blast + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (meson A evimage_eq) + have 1: "drop n x = []" + using cartesian_power_car_memE[of x Q\<^sub>p n] drop_all x_closed + by blast + have 2: "take n x = x" + using cartesian_power_car_memE[of x Q\<^sub>p 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 \\<^bsub>n\<^esub> S" + proof fix x assume A: "x \ partial_pullback n h 0 ((\a. [a]) ` S)" + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 \\<^bsub>n\<^esub> S" + 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)}" +proof- + obtain F where F_def: "F = function_tuple_eval Q\<^sub>p 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)} = F \\<^bsub>k\<^esub> val_relation_set" + proof + show "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)} \ F \\<^bsub>k\<^esub> val_relation_set" + proof fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)}" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 (Q\<^sub>p\<^bsup>2\<^esup>)" + 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 (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 \\<^bsub>k\<^esub> val_relation_set" + by blast + qed + show "F \\<^bsub>k\<^esub> val_relation_set \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)}" + proof fix x assume A: "x \ F \\<^bsub>k\<^esub> val_relation_set" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ F x \ val_relation_set" + by (meson evimage_eq) + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ [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 (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ C}" +proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ 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 (Q\<^sub>p\<^bsup>k\<^esup>)) 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ C}" +proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ 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 (Q\<^sub>p\<^bsup>k\<^esup>)) 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 Q\<^sub>p" + shows "is_semialgebraic k {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). f x = c} = f \\<^bsub>k\<^esub> {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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C}" +proof(cases "C = \") + case True + then have "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). f x = \}" + using assms unfolding val_def by (meson eint.distinct(1)) + then have "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} = f \\<^bsub>k\<^esub> {\}" + unfolding evimage_def by blast + then show ?thesis + using assms semialg_level_set_is_semialg[of f k \] Qp.zero_closed \{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). f x = \}\ by presburger +next + case False + then obtain N::int where N_def: "C = eint N" + by blast + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ N} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ (eint (N-1))}" + proof + show "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ N} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ (eint (N-1))}" + proof + fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C}" + then have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ N} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ (eint (N-1))}" + using 0 1 2 + by blast + qed + show "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ N} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ (eint (N-1))} \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C}" + proof fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ N} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) \ (eint (N-1))}" + have 0: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) = C} " + using 0 2 by blast + qed + qed + have 1: "is_semialgebraic k {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) = val (f x)}" +proof- + have 0: "is_semialgebraic k {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)}" + using assms semialg_val_ineq_set_is_semialg[of f k g] by blast + have 2: " {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) = val (f x)} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)} \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) < val (f x)} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)} - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) = val (f x)}" + using neq_iff le_less by blast + +lemma semialg_val_ineq_set_complement: +"carrier (Q\<^sub>p\<^bsup>k\<^esup>) - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) \ val (f x)} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) < val (g x)}" + using not_le by blast + +lemma semialg_val_strict_ineq_set_complement: +"carrier (Q\<^sub>p\<^bsup>k\<^esup>) - {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). val (g x) < val (f x)} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) < C}" +proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ 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 (Q\<^sub>p\<^bsup>k\<^esup>)) 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) < C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) > C}" +proof- + obtain a where a_def: "a \ carrier Q\<^sub>p \ 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 (Q\<^sub>p\<^bsup>k\<^esup>)) 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 (Q\<^sub>p\<^bsup>k\<^esup>). val (f x) > C} = {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). 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 (Q\<^sub>p\<^bsup>N\<^esup>). val (c x) \ val (a x) + eint n}" +proof- + obtain b where b_def: "b = \[^]n \\<^bsub>SA N\<^esub> a" + 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 (Q\<^sub>p\<^bsup>N\<^esup>) \ 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 (Q\<^sub>p\<^bsup>N\<^esup>). val (c x) \ val (a x) + eint n} = {x \ carrier (Q\<^sub>p\<^bsup>N\<^esup>). 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 (Q\<^sub>p\<^bsup>N\<^esup>). val (c x) = val (a x) + eint n}" +proof- + obtain b where b_def: "b = \[^]n \\<^bsub>SA N\<^esub> a" + 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 (Q\<^sub>p\<^bsup>N\<^esup>) \ 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 (Q\<^sub>p\<^bsup>N\<^esup>). val (c x) = val (a x) + eint n} = {x \ carrier (Q\<^sub>p\<^bsup>N\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (\\<^bsub>SA m\<^esub>) = carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>). f x \ \}" + +lemma nonzero_evimage_closed: + assumes "f \ carrier (SA n)" + shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x \ \}" +proof- + have "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). f x \ \} = f \\<^bsub>n\<^esub> nonzero Q\<^sub>p" + 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 Q\<^sub>p"] 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (\\<^bsub>SA m\<^esub>) = {}" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ f x \ carrier Q\<^sub>p" + assumes "\x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ f x = undefined" + assumes "\k n P. n > 0 \ P \ carrier (Q\<^sub>p [\\<^bsub>1 + k\<^esub>]) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>). a x \ \}" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). a x \ \} = carrier (Q\<^sub>p\<^bsup>m\<^esup>) - {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). a x = \}" + by blast + show ?thesis using assms SA_zero_set_is_semialg[of a m] complement_is_semialg[of m "{x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). a x = \}"] + unfolding 0 by blast +qed + +lemma zero_set_nonzero_set_covers: +"carrier (Q\<^sub>p\<^bsup>n\<^esup>) = 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ (\ f \ As. {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>). val (a x) \ val (i x)}) ` As" + then obtain f where f_def: "f \ As \ x = {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>) = (\ a \ As. Min_set m As a)" +proof + show "carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ \ (Min_set m As ` As)" + proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) " + 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 (Q\<^sub>p\<^bsup>m\<^esup>). val (f x) \ val (fa x)}) ` As" + then obtain g where g_def: "g \ As \ s= {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). val (f x) \ val (g x)}" + by blast + have s_def: "s= {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a }" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a } = f \\<^bsub>n\<^esub> ord_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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a } = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = k}" + using True by blast + have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a} = {} \ k = 0" + proof(cases "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). ord (f x) mod k = a} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "x \ S" + shows "fun_glue n S f g x = f x" +proof- + have "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) - S" + shows "fun_glue n S f g x = g x" +proof- + have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "fun_glue n S f g \\<^bsub>n\<^esub> T = ((f \\<^bsub>n\<^esub> T) \ S) \ ((g \\<^bsub>n\<^esub> T) - S)" +proof + show "fun_glue n S f g \\<^bsub>n\<^esub> T \ ((f \\<^bsub>n\<^esub> T) \ S) \ ((g \\<^bsub>n\<^esub> T) - S)" + proof fix x assume A: "x \ fun_glue n S f g \\<^bsub>n\<^esub> T " + then have 0: "fun_glue n S f g x \ T" + by blast + have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A by (meson evimage_eq) + show "x \ ((f \\<^bsub>n\<^esub> T) \ S) \ ((g \\<^bsub>n\<^esub> T) - 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 \\<^bsub>n\<^esub> T \ S \ (g \\<^bsub>n\<^esub> T - S) \ fun_glue n S f g \\<^bsub>n\<^esub> T" + proof fix x assume A: "x \ f \\<^bsub>n\<^esub> T \ S \ (g \\<^bsub>n\<^esub> T - S)" + then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + by (metis (no_types, opaque_lifting) Diff_iff Int_iff UnE extensional_vimage_closed subsetD) + show "x \ fun_glue n S f g \\<^bsub>n\<^esub> T" + 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 (Q\<^sub>p\<^bsup>k\<^esup>)" + shows "partial_pullback k (fun_glue k S f g) n T = + ((cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) \ partial_pullback k f n T) \ ((partial_pullback k g n T)- (cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))))" +proof + show "partial_pullback k (fun_glue k S f g) n T \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) \ partial_pullback k f n T \ (partial_pullback k g n T - (cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))))" + proof fix x assume A: "x \ partial_pullback k (fun_glue k S f g) n T " + then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>k+n\<^esup>)" unfolding partial_pullback_def partial_image_def + by (meson evimage_eq) + show " x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ partial_pullback k f n T \ (partial_pullback k g n T - (cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))))" + proof(cases "x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))") + 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 (Q\<^sub>p\<^bsup>k\<^esup>) - 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 (Q\<^sub>p\<^bsup>n\<^esup>)))" + using False by blast + then show ?thesis by blast + qed + qed + show "cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ partial_pullback k f n T \ (partial_pullback k g n T - cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>))) + \ partial_pullback k (fun_glue k S f g) n T" + proof fix x assume A: "x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ partial_pullback k f n T \ (partial_pullback k g n T - cartesian_product S (carrier (Q\<^sub>p\<^bsup>n\<^esup>)))" + then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)) \ 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "fun_glue n S f g x \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) - 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 Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>k\<^esup>)) \ partial_pullback n f k T \ (partial_pullback n g k T - cartesian_product S (carrier (Q\<^sub>p\<^bsup>k\<^esup>)))" + 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 \\<^bsub>SA n\<^esub> \ Units (SA n)" +proof(rule SA_Units_memI) + show "fun_glue n S f \\<^bsub>SA n\<^esub> \ carrier (SA n)" + using fun_glue_closed assms SA_is_cring cring.cring_simprules(6) by blast + show "\x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ fun_glue n S f \\<^bsub>SA n\<^esub> x \ \" + proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + show "fun_glue n S f \\<^bsub>SA n\<^esub> x \ \" + apply(cases "x \ S") + using assms SA_is_cring cring.cring_simprules(6) assms(3)[of x] fun_glueE[of f n "\\<^bsub>SA n\<^esub>" 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 "\\<^bsub>SA n\<^esub>" 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 (Q\<^sub>p\<^bsup>n\<^esup>). let S = (THE S. S \ Xs \ x \ S) in (fs S x))" + +lemma parametric_fun_glue_formula: + assumes "Xs partitions (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 \\<^bsub>SA n\<^esub> \\<^bsub>SA n\<^esub>" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "F \ S \ carrier (SA n) \ finsum (SA n) F S x = (\s\S. 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 = (\s\A. F s x) \ + F \ insert a A \ carrier (SA n) \ finsum (SA n) F (insert a A) x = (\s\insert a A. F s x)" + proof- fix A a assume A: "finite A" "F \ A \ carrier (SA n) \ finsum (SA n) F A x = (\s\A. F s x)" + show " F \ insert a A \ carrier (SA n) \ finsum (SA n) F (insert a A) x = (\s\insert 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 = (\s\A. F s x)" + using A by blast + show "finsum (SA n) F (insert a A) x = (\s\insert 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 Q\<^sub>p" + apply(rule Pi_I, rule SA_car_closed[of _ n] ) + using "0" assms by auto + have F01: "F a x \ carrier Q\<^sub>p" + 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: "(\s\insert a A. F s x) = F a x \ (\s\A. 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 \\<^bsub>SA n\<^esub> finsum (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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "\s. s \ S \ F s x = \" + shows "finsum (SA n) F S x = \" +proof- + have "finsum (SA n) F S x = (\s\S. 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 (Q\<^sub>p\<^bsup>n\<^esup>))" + 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 \\<^bsub>SA n\<^esub> char_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>) \ x \ S \ F S x = \" + proof- fix S x assume A: "S \ Xs" "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "x \ S" + then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)") + 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 \\<^bsub>SA n\<^esub> finsum (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 \\<^bsub>SA n\<^esub> finsum (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 (Q\<^sub>p\<^bsup>n\<^esup>)"] 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 Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>). f x \ \} f \\<^bsub>SA n\<^esub>" + +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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 = inv\<^bsub>SA n\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 Q\<^sub>p" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f x \ \" + shows "one_over_fun n (a \\<^bsub>SA n\<^esub>f) x = inv (a \ (f x))" + using one_over_fun_eq[of "a \\<^bsub>SA n\<^esub> f" 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 Q\<^sub>p" + assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f x \ \" + shows "one_over_fun n (a \\<^bsub>SA n\<^esub>f) x = inv a \ inv (f x)" +proof- + have 0: "one_over_fun n (a \\<^bsub>SA n\<^esub> f) x = inv (a \ f x)" + using assms one_over_fun_smult_eval[of f n a x] + by fastforce + have 1: "f x \ nonzero Q\<^sub>p" + 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)]\\<^bsub>SA n\<^esub>f) \ 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f x \ \" + assumes "(k::nat) > 0" + shows "one_over_fun n ([k]\\<^bsub>SA n\<^esub>f) x = inv ([k] \f x)" +proof- + have 0: "([k] \\<^bsub>SA n\<^esub> f) x = [k] \ f x" + using assms SA_add_pow_apply[of f n x k] by linarith + hence 1: "([k] \\<^bsub>SA n\<^esub> f) 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] \\<^bsub>SA n\<^esub> f) x = inv ([k] \\<^bsub>SA n\<^esub> f) x" + using assms one_over_fun_eq[of "[k]\\<^bsub>SA n\<^esub>f" 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[^]\<^bsub>SA n\<^esub>(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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f x \ \" + shows "one_over_fun n (f[^]\<^bsub>SA n\<^esub>(k::nat)) x = inv ((f x) [^] k)" + using one_over_fun_eq[of "f[^]\<^bsub>SA n\<^esub>k" 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m+l\<^esup>). c x \ nonzero Q\<^sub>p \ ord (c x) mod n = a}" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>). c x \ nonzero Q\<^sub>p \ ord (c x) mod n = a} = c \\<^bsub>m+l\<^esub> (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 (Q\<^sub>p\<^bsup>m+l\<^esup>). c x \ nonzero Q\<^sub>p \ ac n (c x) = k}" +proof- + have "{x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>). c x \ nonzero Q\<^sub>p \ ac n (c x) = k}= (c \\<^bsub>m + l\<^esub> ac_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 (Q\<^sub>p\<^bsup>m\<^esup>). c x \ nonzero Q\<^sub>p \ 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ c x \ \" + shows "is_semialgebraic m {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). ac n (c x) = k}" +proof- + have "{x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). c x \ nonzero Q\<^sub>p \ ac n (c x) = k} = {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 \\<^bsub>m\<^esub>pow_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 (Q\<^sub>p\<^bsup>m\<^esup>) then (THE b. b \ carrier Q\<^sub>p \ 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]\\ \ \\<^sub>p" + by (metis Zp.one_closed Zp_int_mult_closed image_eqI inc_of_int) + hence 2: "angular_component ([t]\\) = ac_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using angular_component_of_inclusion[of "[t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>"] + 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]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) n" + unfolding ac_def using 0 by presburger + hence "val_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = 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]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub> \ Units Z\<^sub>p" + using val_Zp_0_imp_unit by blast + hence 5: "ac_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = [t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>" using + ac_Zp_of_Unit \val_Zp ([t] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0\ by blast + have 6: "ac_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) 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]\\ \ \\<^sub>p" + by (metis Zp.one_closed Zp_int_mult_closed image_eqI inc_of_int) + hence 2: "angular_component ([t]\\) = ac_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" + using angular_component_of_inclusion[of "[t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>"] + 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]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) n" + unfolding ac_def using 0 by presburger + hence "val_Zp ([t]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = 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 Q\<^sub>p" + assumes "b \ nonzero Q\<^sub>p" + shows "ac N (a \

b) = ac N a \\<^bsub>Zp_res_ring N\<^esub> inv \<^bsub>Zp_res_ring N\<^esub> ac 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 Q\<^sub>p" + shows "{x \ carrier Q\<^sub>p. 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 Q\<^sub>p" + shows "is_univ_semialgebraic {x \ carrier Q\<^sub>p. 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 \\<^bsub>n\<^esub> {x \ carrier Q\<^sub>p. P x} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). 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 Q\<^sub>p" + assumes "c \ carrier (SA N)" + shows "is_semialgebraic N {x \ carrier (Q\<^sub>p\<^bsup>N\<^esup>). pow_res n (c x) = pow_res n b}" +proof- + have " c \\<^bsub>N\<^esub> {x \ carrier Q\<^sub>p. pow_res n x = pow_res n b} = {x \ carrier (Q\<^sub>p\<^bsup>N\<^esup>). 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 Q\<^sub>p" + assumes "b \ carrier Q\<^sub>p" + assumes "val a = val b + eint i" + shows "b \ nonzero Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(f \\<^bsub>SA n\<^esub> g) 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 Z\<^sub>p" + shows "is_semialgebraic n (f \\<^bsub>n\<^esub> (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 (Q\<^sub>p\<^bsup>m\<^esup>). f x \ \\<^sub>p \ Qp_res (f x) n = l}" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f x \ \\<^sub>p \ Qp_res (f x) n = l} = f \\<^bsub>m\<^esub> {x \ \\<^sub>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 (Q\<^sub>p\<^bsup>k\<^esup>) \ f x \ \\<^sub>p" + assumes "t \ carrier (Zp_res_ring n)" + shows "is_semialgebraic k {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). to_Zp (f x) n = t}" +proof- + have 0: "[t] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \ carrier Z\<^sub>p " + by blast + have 1: "([t] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n = t" + using assms + unfolding Zp_int_inc_rep p_residue_def residue_def residue_ring_def by simp + have "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). to_Zp (f x) n = t} = f \\<^bsub>k\<^esub> {x \ \\<^sub>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]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" 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 (Q\<^sub>p\<^bsup>N\<^esup>). c x \ \\<^sub>p}" +proof- + have 0: "{x \ carrier (Q\<^sub>p\<^bsup>N\<^esup>). c x \ \\<^sub>p} = c \\<^bsub>N\<^esub> \\<^sub>p" + unfolding evimage_def by blast + have 1: "is_univ_semialgebraic \\<^sub>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 (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>p \ to_Zp (c x) n = s}" +proof- + obtain a where a_def: "a = [s]\\" + by blast + have 0: "a \ \\<^sub>p" + using a_def + by (metis Zp.one_closed Zp_int_mult_closed image_iff inc_of_int) + have 1: "to_Zp a = [s]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>" + using 0 unfolding a_def + by (metis Q\<^sub>p_def Zp_def Zp_int_inc_closed \_def inc_to_Zp padic_fields.inc_of_int padic_fields_axioms) + have 2: "([s]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) 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 (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>p \ to_Zp (c x) n = s} = c \\<^bsub>k\<^esub> B\<^bsub>n\<^esub>[a]" + proof + show "{x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>p \ to_Zp (c x) n = s} \ c \\<^bsub>k\<^esub> B\<^bsub>int n\<^esub>[a]" + proof fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>p \ to_Zp (c x) n = s}" + then have 30: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" by blast + have 31: "c x \ \\<^sub>p" using A by blast + have 32: "to_Zp (c x) n = s" using A by blast + have 33: "to_Zp (c x) \ carrier Z\<^sub>p" + 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)) \\<^bsub>Z\<^sub>p\<^esub> (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)) \\<^bsub>Z\<^sub>p\<^esub> (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)) \\<^bsub>Z\<^sub>p\<^esub> (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\<^bsub>int n\<^esub>[a]" + using 31 c_ballI val_ring_memE by blast + thus "x \ c \\<^bsub>k\<^esub> B\<^bsub>int n\<^esub>[a]" + using 30 by blast + qed + show "c \\<^bsub>k\<^esub> B\<^bsub>int n\<^esub>[a] \ {x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>p \ to_Zp (c x) n = s}" + proof fix x assume A: "x \ c \\<^bsub>k\<^esub> B\<^bsub>int n\<^esub>[a]" + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" + 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 Q\<^sub>p" + using x_closed assms function_ring_car_closed SA_car_memE(2) by blast + hence 11: "c x \ a \ \\<^sub>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 \ \\<^sub>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 \ \\<^sub>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) \\<^bsub>Z\<^sub>p\<^esub> to_Zp a)" + using 11 22 0 + by (metis to_Zp_minus to_Zp_val) + have 44: "val_Zp (to_Zp (c x) \\<^bsub>Z\<^sub>p\<^esub> to_Zp a) \ n" + using 33 00 by presburger + have tzpcx: "to_Zp (c x) \ carrier Z\<^sub>p" + using 22 by (metis image_iff inc_to_Zp) + have tzpa: "to_Zp a \ carrier Z\<^sub>p" + using 0 val_ring_memE to_Zp_closed by blast + have 55: "(to_Zp (c x) \\<^bsub>Z\<^sub>p\<^esub> to_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 (Q\<^sub>p\<^bsup>k\<^esup>). c x \ \\<^sub>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 \\<^bsub>m\<^esub> C)" +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 \\<^bsub>m\<^esub> C = {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f x \ \\<^sub>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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "ring_hom_ring (Fun\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (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\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a)" + +lemma Qp_fpoly_to_Qp_poly_is_hom: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(Qp_fpoly_to_Qp_poly n a) \ ring_hom (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)) (Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + shows "Qp_fpoly_to_Qp_poly n a (to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) f) = to_polynomial Q\<^sub>p (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\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_fpoly_to_Qp_poly n a (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p)) = X_poly Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + shows "Qp_fpoly_to_Qp_poly n a (up_ring.monom (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)) f m) = up_ring.monom Q\<^sub>p_x (f a) m" + unfolding Qp_fpoly_to_Qp_poly_def + using UP_cring.poly_lift_hom_monom[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + shows "Qp_fpoly_to_Qp_poly n a f k = (f k) a" + using assms UP_cring.poly_lift_hom_cf[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "P \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + shows "(UP_cring.to_fun (Fun\<^bsub>n\<^esub> Q\<^sub>p) P f) a = UP_cring.to_fun Q\<^sub>p (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\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "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\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "g \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_fpoly_to_Qp_poly n a (compose (Fun\<^bsub>n\<^esub> Q\<^sub>p) f g) = compose Q\<^sub>p (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\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "eval_hom a" f g] + eval_hom_is_hom[of a] ring_hom_ring.homh[of "Fun\<^bsub>n\<^esub> Q\<^sub>p" Q\<^sub>p "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\<^bsub>n\<^esub> Q\<^sub>p))" + assumes "c \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "Qp_fpoly_to_Qp_poly n a (taylor_expansion (Fun\<^bsub>n\<^esub> Q\<^sub>p) c F) = + taylor_expansion Q\<^sub>p (c a) (Qp_fpoly_to_Qp_poly n a F)" +proof- + have 0: "X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \\<^bsub>UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)\<^esub> to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + 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\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a) (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \\<^bsub>UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)\<^esub> to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c) = X_poly Q\<^sub>p \\<^bsub>Q\<^sub>p_x\<^esub> UPQ.to_poly (c a)" + proof- + have 10: "poly_lift_hom (Fun\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a) \ ring_hom (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)) Q\<^sub>p_x" + using Qp_fpoly_to_Qp_poly_def Qp_fpoly_to_Qp_poly_is_hom assms + by presburger + have 11: " to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p))" + by (meson Qp_funs_is_cring UP_cring.intro UP_cring.to_poly_closed assms) + have 12: "X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \ carrier (UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)) " + using UP_cring.X_closed[of "Fun\<^bsub>n\<^esub> Q\<^sub>p"] unfolding UP_cring_def + using Qp_funs_is_cring + by blast + have "Qp_fpoly_to_Qp_poly n a (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \\<^bsub>UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)\<^esub> to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c) = + Qp_fpoly_to_Qp_poly n a (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p)) \\<^bsub>Q\<^sub>p_x\<^esub> Qp_fpoly_to_Qp_poly n a (to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) 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\<^bsub>n\<^esub> Q\<^sub>p)" Q\<^sub>p_x "X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p)" "to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) 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\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a) (compose (Fun\<^bsub>n\<^esub> Q\<^sub>p) F (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \\<^bsub>UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)\<^esub> to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c)) = + UPQ.sub (poly_lift_hom (Fun\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a) F) + (poly_lift_hom (Fun\<^bsub>n\<^esub> Q\<^sub>p) Q\<^sub>p (eval_hom a) (X_poly (Fun\<^bsub>n\<^esub> Q\<^sub>p) \\<^bsub>UP (Fun\<^bsub>n\<^esub> Q\<^sub>p)\<^esub> to_polynomial (Fun\<^bsub>n\<^esub> Q\<^sub>p) c))" + using 0 1 Qp_fpoly_to_Qp_poly_sub[of F n "X_poly_plus (Fun\<^bsub>n\<^esub> Q\<^sub>p) 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "ring_hom_ring (SA n) Q\<^sub>p (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) Q\<^sub>p (eval_hom a)" + +lemma SA_poly_to_Qp_poly_is_hom: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "(SA_poly_to_Qp_poly n a) \ ring_hom (UP (SA n)) (Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "P \ carrier (UP (SA n))" + shows "SA_poly_to_Qp_poly n a P \ carrier Q\<^sub>p_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)" Q\<^sub>p_x P] + by blast + +lemma SA_poly_to_Qp_poly_add: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (UP (SA n))" + assumes "g \ carrier (UP (SA n))" + shows "SA_poly_to_Qp_poly n a (f \\<^bsub>UP (SA n)\<^esub> g) = SA_poly_to_Qp_poly n a f \\<^bsub>Q\<^sub>p_x\<^esub> SA_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (UP (SA n))" + assumes "g \ carrier (UP (SA n))" + shows "SA_poly_to_Qp_poly n a (f \\<^bsub>UP (SA n)\<^esub> g) = SA_poly_to_Qp_poly n a f \\<^bsub>Q\<^sub>p_x\<^esub> SA_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)" Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (UP (SA n))" + assumes "g \ carrier (UP (SA n))" + shows "SA_poly_to_Qp_poly n a (f \\<^bsub>UP (SA n)\<^esub> g) = SA_poly_to_Qp_poly n a f \\<^bsub>Q\<^sub>p_x\<^esub> SA_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (SA n)" + shows "SA_poly_to_Qp_poly n a (to_polynomial (SA n) f) = to_polynomial Q\<^sub>p (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" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "SA_poly_to_Qp_poly n a (X_poly (SA n)) = X_poly Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "f \ carrier (SA n)" + shows "SA_poly_to_Qp_poly n a (up_ring.monom (UP (SA n)) f m) = up_ring.monom Q\<^sub>p_x (f a) m" + unfolding SA_poly_to_Qp_poly_def + using UP_cring.poly_lift_hom_monom[of "SA n" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 Q\<^sub>p (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" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "SA_poly_to_Qp_poly n a (compose (SA n) f g) = compose Q\<^sub>p (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" Q\<^sub>p "eval_hom a" f g] + eval_hom_is_SA_hom[of a] ring_hom_ring.homh[of "SA n" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "deg Q\<^sub>p (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 = \\<^bsub>SA m\<^esub>" + 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "SA_poly_to_Qp_poly n a (taylor_expansion (SA n) c F) = + taylor_expansion Q\<^sub>p (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" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "SA_poly_to_Qp_poly n a (UP_cring.taylor_term (SA n) c F i) = + UP_cring.taylor_term Q\<^sub>p (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" Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + shows "SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) F) = + UP_cring.pderiv Q\<^sub>p (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 \\<^bsub>UP (SA n)\<^esub> q)) = UPQ.pderiv (SA_poly_to_Qp_poly n a (p \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> q)) = UPQ.pderiv (SA_poly_to_Qp_poly n a (p \\<^bsub>UP (SA n)\<^esub> q))" + proof- + have 0: "SA_poly_to_Qp_poly n a p \ carrier (UP Q\<^sub>p)" + 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 Q\<^sub>p)" + 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 Q\<^sub>p)" + 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 Q\<^sub>p)" + 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 \\<^bsub>UP (SA n)\<^esub> UP_cring.pderiv (SA n) q) = + SA_poly_to_Qp_poly n a (UP_cring.pderiv (SA n) p) \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 \\<^bsub>UP Q\<^sub>p\<^esub> SA_poly_to_Qp_poly n a q) = + UPQ.pderiv (SA_poly_to_Qp_poly n a p) \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.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 \\<^bsub>UP (SA n)\<^esub> q) = UP_cring.pderiv (SA n) p \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> q)) = + UPQ.pderiv (SA_poly_to_Qp_poly n a p) \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.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 \\<^bsub>UP (SA n)\<^esub> q)) = + UPQ.pderiv (SA_poly_to_Qp_poly n a p) \ \<^bsub>UP Q\<^sub>p\<^esub> UPQ.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]\\<^bsub>SA n\<^esub>b) (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 Q\<^sub>p) (b a) m" + using SA_poly_to_Qp_poly_monom \b \ carrier (SA n)\ assms by blast + have 2: "b a \ carrier Q\<^sub>p" + 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 Q\<^sub>p) ([m]\(b a)) (m-1)" + using 1 2 A UPQ.pderiv_monom[of "b a" m] + by presburger + have 4: "[m] \\<^bsub>SA n\<^esub> b \ 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] \\<^bsub>SA n\<^esub> b) (m-1)) = up_ring.monom (UP Q\<^sub>p) (([m] \\<^bsub>SA n\<^esub> b) a) (m-1)" + using SA_poly_to_Qp_poly_monom[of a n "[m]\\<^bsub>SA n\<^esub>b" "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 Q\<^sub>p) (([m] \\<^bsub>SA n\<^esub> b) 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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] \\<^bsub>SA m\<^esub> g (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 \ \\<^bsub>SA m\<^esub>" + unfolding n_def using assms UPSA.deg_ltrm by fastforce + have 1: "(pderiv m f) (n-1) = [n]\\<^bsub>SA m\<^esub> (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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "SA_poly_to_Qp_poly m x (a \\<^bsub>UP (SA m)\<^esub> f) = a x \\<^bsub>UP Q\<^sub>p\<^esub> SA_poly_to_Qp_poly m x f" +proof- + have 0: "a \\<^bsub>UP (SA m)\<^esub> f = to_polynomial (SA m) a \\<^bsub>UP (SA m)\<^esub> f" + using assms UPSA.to_poly_mult_simp(1) by presburger + have 1: "SA_poly_to_Qp_poly m x (a \\<^bsub>UP (SA m)\<^esub> f) = SA_poly_to_Qp_poly m x (to_polynomial (SA m) a) \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 \\<^bsub>UP (SA m)\<^esub> f) = (to_polynomial Q\<^sub>p (a x)) \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 Q\<^sub>p "a x" "SA_poly_to_Qp_poly m x f"]) + unfolding UP_cring_def + apply (simp add: Qp.cring) + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ f i x \ \\<^sub>p" + assumes "deg (SA m) f \ d" + assumes "C \ poly_res_classes n d" + shows "is_semialgebraic m {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>). SA_poly_to_Qp_poly m x f \ C} = + (\ i \ {..d}. {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>) \ SA_poly_to_Qp_poly m x f \ C" + "S \ (\i. {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>p \ Qp_res (f i x) n = Qp_res (g i) n}" + using A by blast + have S_eq: "S = {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>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 Q\<^sub>p (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 \ (\i\d. {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>p \ Qp_res (f i x) n = Qp_res (g i) n}) \ x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ SA_poly_to_Qp_poly m x f \ C" + proof- + fix x assume A: "x \ (\i\d. {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>p \ Qp_res (f i x) n = Qp_res (g i) n})" + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 Q\<^sub>p (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 Q\<^sub>p 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>). f i x \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>). UP_cring.to_fun Q\<^sub>p (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 (Q\<^sub>p\<^bsup>Suc n\<^esup>) \ carrier Q\<^sub>p)" +proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + 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 Q\<^sub>p" + using A t_def cartesian_power_head + by blast + have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using A a_def cartesian_power_tail + by blast + have 0: "SA_poly_to_SA_fun n P x = UP_cring.to_fun Q\<^sub>p (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 Q\<^sub>p" + using assms t_closed a_def 0 UP_cring.to_fun_closed[of Q\<^sub>p "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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "t \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>). f (g a)) \ carrier (SA m)" +proof(rule SA_car_memI) + show "(\a\carrier (Q\<^sub>p\<^bsup>m\<^esup>). f (g a)) \ carrier (Qp_funs m)" + proof(rule Qp_funs_car_memI) + show " (\a\carrier (Q\<^sub>p\<^bsup>m\<^esup>). f (g a)) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" + proof fix a assume A: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + then have "g a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using is_semialg_map_def[of m n g] assms + by blast + then show "f (g a) \ carrier Q\<^sub>p" + using A assms SA_car_memE(3)[of f n] + by blast + qed + show " \x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (\a\carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ (f \ g) a = (\a\carrier (Q\<^sub>p\<^bsup>m\<^esup>). f (g a)) a)" + using assms comp_apply[of f g] unfolding restrict_def + by metis + then show "is_semialg_function m (\a\carrier (Q\<^sub>p\<^bsup>m\<^esup>). f (g a))" + using 0 1 semialg_function_on_carrier'[of m "f \ g" "\ a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). f (g a)" ] + by blast +qed + +lemma tl_comp_in_SA: + assumes "f \ carrier (SA n)" + shows "(\ a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + shows "SA_poly_to_SA_fun n (f \\<^bsub>UP (SA n)\<^esub> g) a = SA_poly_to_SA_fun n f a \\<^bsub>Q\<^sub>p\<^esub> SA_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 \\<^bsub>UP (SA n)\<^esub> g) = SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> SA_poly_to_SA_fun n g" +proof fix x + show " SA_poly_to_SA_fun n (f \\<^bsub>UP (SA n)\<^esub> g) x = (SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> SA_poly_to_SA_fun n g) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)") + 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 \\<^bsub>UP (SA n)\<^esub> g) x = undefined" + unfolding SA_poly_to_SA_fun_def + by (meson False restrict_apply) + have F1: "(SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + shows "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = (f (tl a))\(hd a)[^]\<^bsub>Q\<^sub>p\<^esub>k " +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 Q\<^sub>p_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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "t \ carrier Q\<^sub>p" + shows "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) (t#x) = (f x)\t[^]\<^bsub>Q\<^sub>p\<^esub>k " +proof- + have 0: "hd (t#x) = t" + by simp + have 1: "tl (t#x) = x" + by auto + have 2: "(t#x) \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>n\<^esup>). a!0)" + using assms index_is_semialg_function restrict_is_semialg by blast + have 1: "restrict (\a\carrier (Q\<^sub>p\<^bsup>n\<^esup>). a ! 0) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict lead_coeff (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" + proof fix x + show "restrict (\a\carrier (Q\<^sub>p\<^bsup>n\<^esup>). a ! 0) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x = restrict lead_coeff (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) x" + unfolding restrict_def + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)") + 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 (Q\<^sub>p\<^bsup>n\<^esup>). 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>))" + 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[^]\<^bsub>SA (Suc n)\<^esub> 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) \\<^bsub>SA (Suc n)\<^esub> h[^]\<^bsub>SA (Suc n)\<^esub> 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>) \ SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = ((f \ tl) \\<^bsub>SA (Suc n)\<^esub> h[^]\<^bsub>SA (Suc n)\<^esub> k) a" + proof- fix a assume "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + then show "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) a = ((f \ tl) \\<^bsub>SA (Suc n)\<^esub> h[^]\<^bsub>SA (Suc n)\<^esub> 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 [^]\<^bsub>SA (Suc n)\<^esub> 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) \\<^bsub>SA (Suc n)\<^esub> h[^]\<^bsub>SA (Suc n)\<^esub> k)" + proof fix x + show "SA_poly_to_SA_fun n (up_ring.monom (UP (SA n)) f k) x = ((f \ tl) \\<^bsub>SA (Suc n)\<^esub> h [^]\<^bsub>SA (Suc n)\<^esub> k) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)") + 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 \\<^bsub>UP (SA n)\<^esub> g) = SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> SA_poly_to_SA_fun n g" +proof(rule function_ring_car_eqI[of _ "Suc n"]) + show "SA_poly_to_SA_fun n (f \\<^bsub>UP (SA n)\<^esub> g) \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)) Q\<^sub>p)" + proof- + have "f \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> g"] SA_car_in_Qp_funs_car[of "Suc n"] + by blast + qed + show "SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> SA_poly_to_SA_fun n g \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)) Q\<^sub>p)" + 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>) \ SA_poly_to_SA_fun n (f \\<^bsub>UP (SA n)\<^esub> g) a = (SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> SA_poly_to_SA_fun n g) a" + proof- fix a assume A: "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + then obtain t b where tb_def: "a = t#b" + using cartesian_power_car_memE[of a Q\<^sub>p "Suc n"] by (meson length_Suc_conv) + have t_closed: "t \ carrier Q\<^sub>p" + using tb_def A cartesian_power_head[of a Q\<^sub>p n] by (metis list.sel(1)) + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using tb_def A cartesian_power_tail[of a Q\<^sub>p n] by (metis list.sel(3)) + have 0: "f \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> g) a = (SA_poly_to_Qp_poly n b (f \\<^bsub>UP (SA n)\<^esub> g))\t" + using SA_poly_to_SA_fun_formula[of "f \\<^bsub>UP (SA n)\<^esub> g" n b t] t_closed b_closed tb_def 0 assms(1) + by blast + have 2: "(SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> 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 \\<^bsub>SA (Suc n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>UP (SA n)\<^esub> g) a = (SA_poly_to_SA_fun n f \\<^bsub>SA (Suc n)\<^esub> 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 (\\<^bsub>UP (SA n)\<^esub>) = \\<^bsub>SA (Suc n)\<^esub>" +proof(rule function_ring_car_eqI[of _ "Suc n"]) + have "\\<^bsub>UP (SA n)\<^esub> \ carrier (UP (SA n))" + using UP_SA_n_is_cring cring.cring_simprules(6) by blast + thus " SA_poly_to_SA_fun n \\<^bsub>UP (SA n)\<^esub> \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)) Q\<^sub>p)" + using SA_poly_to_SA_fun_is_SA[of "\\<^bsub>UP (SA n)\<^esub>"] SA_car_in_Qp_funs_car[of "Suc n"] SA_is_UP_cring[of n] + by blast + show "\\<^bsub>SA (Suc n)\<^esub> \ carrier (function_ring (carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)) Q\<^sub>p)" + unfolding SA_one + using function_one_closed by blast + show "\a. a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>) \ SA_poly_to_SA_fun n \\<^bsub>UP (SA n)\<^esub> a = \\<^bsub>SA (Suc n)\<^esub> a" + proof- fix a assume A: " a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + then obtain t b where tb_def: "a = t#b" + using cartesian_power_car_memE[of a Q\<^sub>p "Suc n"] by (meson length_Suc_conv) + have t_closed: "t \ carrier Q\<^sub>p" + using tb_def A cartesian_power_head[of a Q\<^sub>p n] by (metis list.sel(1)) + have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + using tb_def A cartesian_power_tail[of a Q\<^sub>p n] by (metis list.sel(3)) + have 0: "SA_poly_to_SA_fun n \\<^bsub>UP (SA n)\<^esub> a = (SA_poly_to_Qp_poly n b \\<^bsub>UP (SA n)\<^esub>)\t" + using SA_poly_to_SA_fun_formula \\\<^bsub>UP (SA n)\<^esub> \ carrier (UP (SA n))\ b_closed t_closed tb_def + by blast + have 1: "SA_poly_to_Qp_poly n b \\<^bsub>UP (SA n)\<^esub> = \\<^bsub>UP Q\<^sub>p\<^esub>" + using SA_poly_to_Qp_poly_is_hom[of b] ring_hom_one[of _ "UP (SA n)" "UP Q\<^sub>p"] b_closed + by blast + thus "SA_poly_to_SA_fun n \\<^bsub>UP (SA n)\<^esub> a = \\<^bsub>SA (Suc n)\<^esub> 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 (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "t \ carrier Q\<^sub>p" + 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 Q\<^sub>p (c x) f k) \(t \ c x)[^]\<^bsub>Q\<^sub>p\<^esub> k" +proof- + have 0: "hd (t#x) = t" + by simp + have 1: "tl (t#x) = x" + by auto + have 2: "(t#x) \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "(\\<^bsub>SA m\<^esub>i\I. F i) x = (\i\I. F i x)" +proof- + have "F \ I \ carrier (SA m) \ (\\<^bsub>SA m\<^esub>i\I. F i) x = (\i\I. 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 = (\i\A. 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 = (\i\A. F i x)" + using IH by blast + show "finsum (SA m) F (insert a A) x = (\i\insert 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 Q\<^sub>p" proof fix i assume "i \ A" thus "F i x \ carrier Q\<^sub>p" + using 0 assms(3) SA_car_memE(3)[of "F i" m] by blast qed + have F1: "F a x \ carrier Q\<^sub>p" + 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 \\<^bsub>SA m\<^esub> finsum (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 (\i\I. F i) = (\\<^bsub>S\<^esub>i\I. h (F i))" +proof- + have "F \ I \ carrier R \ h (\i\I. F i) = (\\<^bsub>S\<^esub>i\I. 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 {} = \\<^bsub>S\<^esub>\ 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) = (\\<^bsub>S\<^esub>i\A. 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) = (\\<^bsub>S\<^esub>i\A. 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: "(\\<^bsub>S\<^esub>i\A. 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)) = (\\<^bsub>S\<^esub>i\insert 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: "(\\<^bsub>S\<^esub>i\insert a A. h (F i)) = h (F a) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub>i\A. 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 = (\\<^bsub>UP (SA m)\<^esub>i\I. F i)" + assumes "x \ carrier (Q\<^sub>p\<^bsup>Suc m\<^esup>)" + shows "SA_poly_to_SA_fun m f x = (\i\I. 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 = (\\<^bsub>SA (Suc m)\<^esub>i\I. SA_poly_to_SA_fun m (F i))" + unfolding assms + apply(rule finsum_ring_hom) + apply (simp add: R.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 (Q\<^sub>p\<^bsup>Suc m\<^esup>)" + 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 = (\\<^bsub>UP (SA m)\<^esub>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 Q\<^sub>p" + using assms cartesian_power_head by blast + have tl_closed: "tl x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + using assms cartesian_power_tail by blast + obtain t a where ta_def: " x = t#a" + using assms cartesian_power_car_memE[of x Q\<^sub>p "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 Q\<^sub>p (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 Q\<^sub>p (c (tl x)) (SA_poly_to_Qp_poly m (tl x) f) i) \((hd x) \ c (tl x))[^]\<^bsub>Q\<^sub>p\<^esub> i)" + using 1 2 unfolding 4 by blast + have 6: "taylor_expansion Q\<^sub>p (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 Q\<^sub>p (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\<^bsub>m\<^esub> Q\<^sub>p)" + assumes "UP_ring.lcf (SA m) g \ Units (SA m)" + assumes "\x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). (SA_poly_to_SA_fun m g) (\ x#x) = \" + shows "\ = \\<^bsub>SA m\<^esub>(g 0)\\<^bsub>SA m\<^esub> (inv\<^bsub>SA m\<^esub> (g 1))" +proof(rule ext) + fix x show " \ x = (\\<^bsub>SA m\<^esub> g 0 \\<^bsub>SA m\<^esub> inv\<^bsub>SA m\<^esub> g 1) x" + proof(cases "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)") + 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 Q\<^sub>p (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 Q\<^sub>p)" + using assms True SA_poly_to_Qp_poly_closed by blast + show "deg Q\<^sub>p (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 Q\<^sub>p _ "\ 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 Q\<^sub>p" + 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 Q\<^sub>p" + using True assms Qp.function_ring_car_memE by blast + have T5: "g 1 x \ nonzero Q\<^sub>p" + 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: "(inv\<^bsub>SA m\<^esub> g 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: "(\\<^bsub>SA m\<^esub> g 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: "((\\<^bsub>SA m\<^esub> g 0) \\<^bsub>SA m\<^esub> inv\<^bsub>SA m\<^esub> g 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 = (\\<^bsub>SA m\<^esub> g 0 \\<^bsub>SA m\<^esub> inv\<^bsub>SA m\<^esub> g 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\<^bsub>m\<^esub> Q\<^sub>p)" + assumes "UP_ring.lcf (SA m) g \ Units (SA m)" + assumes "\x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). (SA_poly_to_SA_fun m g) (\ x#x) = \" + shows "\ \ carrier (SA m)" +proof- + have 0: "\ = \\<^bsub>SA m\<^esub>(g 0)\\<^bsub>SA m\<^esub> (inv\<^bsub>SA m\<^esub> (g 1))" + using assms SA_deg_one_eval by blast + have 1: "(inv\<^bsub>SA m\<^esub> (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: "\\<^bsub>SA m\<^esub>(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 Q\<^sub>p" + assumes "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "t#x \ carrier (Q\<^sub>p\<^bsup>Suc m\<^esup>)" + using assms cartesian_power_cons[of x Q\<^sub>p m t] Suc_eq_plus1 by presburger + +lemma Qp_pow_ConsE: + assumes "x \ carrier (Q\<^sub>p\<^bsup>Suc m\<^esup>)" + shows "tl x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + "hd x \ carrier Q\<^sub>p" + using assms cartesian_power_tail apply blast + using assms cartesian_power_head by blast + +lemma(in ring) add_monoid_one: +"\\<^bsub>add_monoid R\<^esub> = \" + by (metis add.generate_empty add.group_l_invI add.l_inv_ex empty_iff group.generate_empty insert_iff) + +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 Q\<^sub>p" + shows "SA_poly_to_SA_fun m f (t#x) = (\i\inds. (a i x)\(t \ c x)[^]i)" +proof- + have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>Suc m\<^esup>)" + using x_closed assms(8) cartesian_power_cons[of x Q\<^sub>p 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 = \\<^bsub>SA m\<^esub>" + 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 Q\<^sub>p" + 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 Q\<^sub>p" + using x_closed function_ring_car_closed SA_car_memE(2) by blast + have 32: "c x \ carrier Q\<^sub>p" + using assms x_closed SA_car_memE(3) by blast + hence 33: "(t \ c x)[^] i \ carrier Q\<^sub>p" + using assms by blast + show "taylor_expansion (SA m) c f i x \ (t \ c x) [^] i \ carrier Q\<^sub>p" + using 30 31 32 33 by blast + qed + have 4: "SA_poly_to_SA_fun m f (t#x) = (\i\inds. 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 Q\<^sub>p" + proof fix i assume A: "i \ inds" + show "a i x \ (t \ c x) [^] i \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "c \ carrier (SA m)" + shows "taylor_expansion (SA m) c f i x = + taylor_expansion Q\<^sub>p (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 Q\<^sub>p (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 Q\<^sub>p (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 (Q\<^sub>p\<^bsup>n\<^esup>))" + +lemma Qp_ev_hom_ev: + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + 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 (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_ev_hom n f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" + using Qp_ev_hom_ev assms by (metis Pi_I eval_at_point_closed) + +lemma Qp_ev_hom_is_semialg_function: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + 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 (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_ev_hom n f \ carrier (Fun\<^bsub>n\<^esub> Q\<^sub>p)" + 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 (Q\<^sub>p[\\<^bsub>n\<^esub>])" + 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 (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_ev_hom n (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_ev_hom n f) \\<^bsub>SA n\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)" + have " Qp_ev_hom n (f \\<^bsub>Q\<^sub>p [\\<^bsub>n\<^esub>]\<^esub> 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 \\<^bsub>Q\<^sub>p [\\<^bsub>n\<^esub>]\<^esub> g) a = (Qp_ev_hom n f \\<^bsub>SA n\<^esub> Qp_ev_hom n g) a" + using A SA_add by blast +qed + +lemma Qp_ev_hom_mult: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" + shows "Qp_ev_hom n (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_ev_hom n f) \\<^bsub>SA n\<^esub> (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 (Q\<^sub>p\<^bsup>n\<^esup>)" + have " Qp_ev_hom n (f \\<^bsub>Q\<^sub>p [\\<^bsub>n\<^esub>]\<^esub> 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 \\<^bsub>Q\<^sub>p [\\<^bsub>n\<^esub>]\<^esub> g) a = (Qp_ev_hom n f \\<^bsub>SA n\<^esub> Qp_ev_hom n g) a" + using A SA_mult by blast +qed + +lemma Qp_ev_hom_one: + shows "Qp_ev_hom n \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> = \\<^bsub>SA n\<^esub>" + 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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) (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 Q\<^sub>p" + shows "Qp_ev_hom n (Qp.indexed_const c) = \\<^bsub>n\<^esub> c" + 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 ("\\<^bsub>_, _\<^esub>") + +lemma Qp_ev_hom_pvar: + assumes "i < n" + shows "Qp_ev_hom n (pvar Q\<^sub>p i) = \\<^bsub>n, i\<^esub>" + 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 (Q\<^sub>p\<^bsup>m\<^esup>). eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) (SA n) (Qp_ev_hom n) \ ring_hom (UP (Q\<^sub>p[\\<^bsub>n\<^esub>])) (UP (SA n))" + using UP_cring.poly_lift_hom_is_hom[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) (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 (Q\<^sub>p[\\<^bsub>Suc n\<^esub>]) (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 (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + shows "coord_ring_to_UP_SA n (f \\<^bsub>Q\<^sub>p[\\<^bsub>Suc n\<^esub>]\<^esub>g) = coord_ring_to_UP_SA n f \\<^bsub>UP (SA n)\<^esub> 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 (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + assumes "g \ carrier (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + shows "coord_ring_to_UP_SA n (f \\<^bsub>Q\<^sub>p[\\<^bsub>Suc n\<^esub>]\<^esub>g) = coord_ring_to_UP_SA n f \\<^bsub>UP (SA n)\<^esub> 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 \\<^bsub>Q\<^sub>p[\\<^bsub>Suc n\<^esub>]\<^esub> = \\<^bsub>UP (SA n)\<^esub>" + using coord_ring_to_UP_SA_is_hom ring_hom_one + by blast + +lemma coord_ring_to_UP_SA_closed: + assumes "f \ carrier (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + 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 Q\<^sub>p" + shows "coord_ring_to_UP_SA n (Qp.indexed_const c) = to_polynomial (SA n) (\\<^bsub>n\<^esub> c)" +proof- + have 0: "pre_to_univ_poly (Suc n) 0 (Qp.indexed_const c) = ring.indexed_const (Q\<^sub>p[\\<^bsub>n\<^esub>]) (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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) (Qp.indexed_const c) " + unfolding to_univ_poly_def + using UP_cring.IP_to_UP_indexed_const[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "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) = \\<^bsub>n\<^esub> c" + using Qp_ev_hom_constant[of c n] assms by blast + have 3: "poly_lift_hom (Q\<^sub>p [\\<^bsub>n\<^esub>]) (SA n) (Qp_ev_hom n) (to_polynomial (Q\<^sub>p [\\<^bsub>n\<^esub>]) (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 "Q\<^sub>p[\\<^bsub>n\<^esub>]" "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 (Q\<^sub>p [\\<^bsub>n\<^esub>]) (SA n) (Qp_ev_hom n) (to_univ_poly (Suc n) 0 (Qp.indexed_const c) ) = to_polynomial (SA n) (\\<^bsub>n\<^esub> c)" + using "1" "2" by presburger + show ?thesis + using assms 4 Qp.indexed_const_closed[of c "{..p {..p 0) = up_ring.monom (UP (SA n)) \\<^bsub>SA n\<^esub> 1" +proof- + have 0: "pre_to_univ_poly (Suc n) 0 (pvar Q\<^sub>p 0) = pvar (Q\<^sub>p[\\<^bsub>n\<^esub>]) 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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) 0) = up_ring.monom (UP (Q\<^sub>p[\\<^bsub>n\<^esub>])) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> 1" + using cring.IP_to_UP_var[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "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 Q\<^sub>p 0) = up_ring.monom (UP (Q\<^sub>p[\\<^bsub>n\<^esub>])) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> 1" + unfolding to_univ_poly_def using 0 1 comp_apply + by metis + have 3: "poly_lift_hom (Q\<^sub>p[\\<^bsub>n\<^esub>]) (SA n) (Qp_ev_hom n) (up_ring.monom (UP (Q\<^sub>p[\\<^bsub>n\<^esub>])) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> 1) + = up_ring.monom (UP (SA n)) \\<^bsub>SA n\<^esub> 1" + using UP_cring.poly_lift_hom_monom[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "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 Q\<^sub>p i) = to_polynomial (SA n) (\\<^bsub>n, i-1\<^esub>)" +proof- + have 0: "pre_to_univ_poly (Suc n) 0 (pvar Q\<^sub>p i) = ring.indexed_const (Q\<^sub>p[\\<^bsub>n\<^esub>]) (pvar Q\<^sub>p (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 (Q\<^sub>p[\\<^bsub>n\<^esub>]) (pvar Q\<^sub>p (i-1))) = to_polynomial (Q\<^sub>p[\\<^bsub>n\<^esub>]) (pvar Q\<^sub>p (i-1))" + using UP_cring.IP_to_UP_indexed_const[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "pvar Q\<^sub>p (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 Q\<^sub>p i) = to_polynomial (Q\<^sub>p[\\<^bsub>n\<^esub>]) (pvar Q\<^sub>p (i-1))" + unfolding to_univ_poly_def using 0 1 comp_apply + by metis + have 3: "poly_lift_hom (Q\<^sub>p[\\<^bsub>n\<^esub>]) (SA n) (Qp_ev_hom n) (to_polynomial (Q\<^sub>p[\\<^bsub>n\<^esub>]) (pvar Q\<^sub>p (i-1))) + = to_polynomial (SA n) (\\<^bsub>n, (i-1)\<^esub>)" + using UP_cring.poly_lift_hom_extends_hom[of "Q\<^sub>p[\\<^bsub>n\<^esub>]" "SA n" "Qp_ev_hom n" "pvar Q\<^sub>p (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 (Q\<^sub>p[\\<^bsub>Suc n\<^esub>])" + assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" + assumes "t \ carrier Q\<^sub>p" + 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 (Q\<^sub>p\<^bsup>Suc n\<^esup>)" + using assms cartesian_power_cons by (metis Suc_eq_plus1) + show "f \ carrier (Q\<^sub>p [\\<^bsub>Suc n\<^esub>])" + by (simp add: assms) + show "\c. c \ carrier Q\<^sub>p \ eval_at_point Q\<^sub>p (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 Q\<^sub>p" + have 0: "eval_at_point Q\<^sub>p (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 Q\<^sub>p 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 Q\<^sub>p (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 (Q\<^sub>p [\\<^bsub>Suc n\<^esub>]) \ + q \ carrier (Q\<^sub>p [\\<^bsub>Suc n\<^esub>]) \ + eval_at_point Q\<^sub>p (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \ t \ + eval_at_point Q\<^sub>p (t # a) q = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n q) \ t \ + eval_at_point Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> q) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> q)) \ t" + proof- fix p q assume A: "p \ carrier (Q\<^sub>p [\\<^bsub>Suc n\<^esub>])" "q \ carrier (Q\<^sub>p [\\<^bsub>Suc n\<^esub>])" + "eval_at_point Q\<^sub>p (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \ t" + "eval_at_point Q\<^sub>p (t # a) q = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n q) \ t" + have 0: "eval_at_point Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> q) = eval_at_point Q\<^sub>p (t # a) p \ eval_at_point Q\<^sub>p (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 \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> q))= + SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> q) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> 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 (Q\<^sub>p [\\<^bsub>Suc n\<^esub>]) \ + i < Suc n \ + eval_at_point Q\<^sub>p (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \ t \ + eval_at_point Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) \ t" + proof- fix p i assume A: "p \ carrier (Q\<^sub>p [\\<^bsub>Suc n\<^esub>])" "i < Suc n" + "eval_at_point Q\<^sub>p (t # a) p = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \ t" + show " eval_at_point Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) \ t" + proof- + have 0: "eval_at_point Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = eval_at_point Q\<^sub>p (t # a) p \ eval_at_point Q\<^sub>p (t # a) (pvar Q\<^sub>p 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 \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = coord_ring_to_UP_SA n p \\<^bsub>UP (SA n)\<^esub> coord_ring_to_UP_SA n (pvar Q\<^sub>p 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 \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) = + SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \\<^bsub>UP Q\<^sub>p\<^esub>SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (pvar Q\<^sub>p 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 Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) \ t" + proof(cases "i = 0") + case True + have T0: "SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) = + SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) \ 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 Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = eval_at_point Q\<^sub>p (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 "\\<^bsub>n, i - 1\<^esub> a = 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 \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i)) = + SA_poly_to_Qp_poly n a (coord_ring_to_UP_SA n p) \\<^bsub>UP Q\<^sub>p\<^esub> to_polynomial Q\<^sub>p (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 "\\<^bsub>n, i - 1\<^esub>"] + 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 Q\<^sub>p (t # a) (p \\<^bsub>Q\<^sub>p [\\<^bsub>Suc n\<^esub>]\<^esub> pvar Q\<^sub>p i) = eval_at_point Q\<^sub>p (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 = \\<^bsub>SA m\<^esub>" + 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 = \\<^bsub>SA m\<^esub>" + using 0 assms UPSA.deg_leE by blast + have 3: "g n = \\<^bsub>SA m\<^esub>" + using 1 assms UPSA.deg_leE by blast + show "fun_glue m S (f n) (g n) = \\<^bsub>SA m\<^esub>" + unfolding SA_zero function_ring_def ring_record_simps function_zero_def 2 3 + proof fix x + show " fun_glue m S (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) x = (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)") + 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 = \\<^bsub>SA m\<^esub>" + 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 = \\<^bsub>SA m\<^esub>" + using 0 assms UPSA.deg_leE by blast + have 3: "g n = \\<^bsub>SA m\<^esub>" + using 1 assms UPSA.deg_leE by blast + show "fun_glue m S (f n) (g n) = \\<^bsub>SA m\<^esub>" + unfolding SA_zero function_ring_def ring_record_simps function_zero_def 2 3 + proof fix x + show " fun_glue m S (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) x = (\x\carrier (Q\<^sub>p\<^bsup>m\<^esup>). \) x" + apply(cases "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)") + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 \ \\<^sub>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 \ \\<^sub>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 \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>)" + shows "\ \<^bsub>UP (SA m)\<^esub> \ 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 \\<^bsub>UP (SA m)\<^esub> = \\<^bsub>UP Q\<^sub>p\<^esub>" + 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 \\<^bsub>UP (SA m)\<^esub> i \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "f \ integral_on m B" + assumes "g \ integral_on m B" + shows "f \\<^bsub>UP (SA m)\<^esub> g \ integral_on m B" +proof(rule integral_on_memI) + show "f \\<^bsub>UP (SA m)\<^esub> 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 \\<^bsub>UP (SA m)\<^esub> g) i \ \\<^sub>p" + proof- fix x i assume A: "x \ B" + have 0: "SA_poly_to_Qp_poly m x (f \\<^bsub>UP (SA m)\<^esub> g) = SA_poly_to_Qp_poly m x f + \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 \\<^bsub>UP (SA m)\<^esub> 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 \\<^bsub>UP (SA m)\<^esub> g) i \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "f \ integral_on m B" + assumes "g \ integral_on m B" + shows "f \\<^bsub>UP (SA m)\<^esub> g \ integral_on m B" +proof(rule integral_on_memI) + show "f \\<^bsub>UP (SA m)\<^esub> 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 \\<^bsub>UP (SA m)\<^esub> g) i \ \\<^sub>p" + proof- fix x i assume A: "x \ B" + have 0: "SA_poly_to_Qp_poly m x (f \\<^bsub>UP (SA m)\<^esub> g) = SA_poly_to_Qp_poly m x f + \\<^bsub>UP Q\<^sub>p\<^esub> SA_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 (Q\<^sub>p \ carrier := \\<^sub>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 Q\<^sub>p). \n. h n \ \\<^sub>p}" + unfolding S_def using UPQ.UP_ring_subring_car[of \\<^sub>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 \\<^bsub>UP (SA m)\<^esub> g) \ carrier S" + unfolding 0 + using cring.cring_simprules(5)[of S]3 4 1 UPQ.UP_ring_subring_mult[of \\<^sub>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 \\<^bsub>UP (SA m)\<^esub> g) i \ \\<^sub>p" + using 5 unfolding 2 mem_Collect_eq by blast + qed +qed + +lemma integral_on_a_minus: + assumes "B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + assumes "f \ integral_on m B" + shows "\\<^bsub>UP (SA m)\<^esub> 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 (\\<^bsub>UP (SA m)\<^esub> f) = \\<^bsub>UP Q\<^sub>p\<^esub> SA_poly_to_Qp_poly m x f" + apply(rule UP_cring.ring_hom_uminus[of "UP Q\<^sub>p" "UP (SA m)" "SA_poly_to_Qp_poly m x" f ] ) + unfolding UP_cring_def + apply (simp add: UPQ.M_cring) + using UPSA.P.is_ring apply blast + 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 Q\<^sub>p) \ (\\<^bsub>UP Q\<^sub>p\<^esub> f) i = \ (f i)" + using UPQ.cfs_a_inv by blast + have 1: "SA_poly_to_Qp_poly m x (\\<^bsub>UP (SA m)\<^esub> 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 (\\<^bsub>UP (SA m)\<^esub> f) i \ \\<^sub>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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 "\\<^bsub>UP (SA m)\<^esub> \ integral_on m B" + using one_integral_on assms by blast + show " \h. h \ integral_on m B \ \\<^bsub>UP (SA m)\<^esub> 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 \\<^bsub>UP (SA m)\<^esub> 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 \\<^bsub>UP (SA m)\<^esub> h2 \ integral_on m B" + using integral_on_plus assms by blast +qed + +lemma val_ring_add_pow: + assumes "a \ carrier Q\<^sub>p" + 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 Q\<^sub>p)" + assumes "\ i. f i \ \\<^sub>p" + shows "\x. x \ \\<^sub>p \ f \ x \ \\<^sub>p" +proof- fix x assume A: "x \ \\<^sub>p" + obtain S where S_def: "S = (Q\<^sub>p \ carrier := \\<^sub>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 Q\<^sub>p 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 \\<^sub>p] assms val_ring_subring by blast + have 3: "to_function S f x \ \\<^sub>p" + using UPQ.UP_subring_eval_closed[of \\<^sub>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 \ \\<^sub>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 \ \\<^sub>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 (\\<^bsub>UP (SA m)\<^esub>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>). 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 "\\<^bsub>UP (SA m)\<^esub>" 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 "\\<^bsub>UP (SA m)\<^esub>" 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ g i x \ \\<^sub>p" + proof- fix i x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + show "g i x \ \\<^sub>p" + unfolding g_def apply(cases "x \ B") + using SA_poly_glue_cfs1[of f m "\\<^bsub>UP (SA m)\<^esub>" 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 "\\<^bsub>UP (SA m)\<^esub>" 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ x \ B \ g i x = \\<^bsub>UP (SA m)\<^esub> i x" + unfolding g_def apply(rule SA_poly_glue_cfs2[of f m "\\<^bsub>UP (SA m)\<^esub>" B ]) + by(rule assms, blast, rule assms, blast, blast) + have 4: "\i x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ x \ B \ g i x \ \\<^sub>p " + using 3 cfs_one[of m] 2 by blast + have 5: "\i x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ g i x \ \\<^sub>p" + using 1 4 assms 2 by blast + have 6: "is_semialgebraic m {x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). 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 \ \\<^sub>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 (\\<^bsub>UP (SA m)\<^esub>)" + 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ g i x \ \\<^sub>p" + proof- fix i x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" + show "g i x \ \\<^sub>p" + unfolding g_def apply(cases "x \ B") + using SA_poly_glue_cfs1[of f m "\\<^bsub>UP (SA m)\<^esub>" 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 "\\<^bsub>UP (SA m)\<^esub>" 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ x \ B \ g i x = \\<^bsub>UP (SA m)\<^esub> i x" + unfolding g_def apply(rule SA_poly_glue_cfs2[of f m "\\<^bsub>UP (SA m)\<^esub>" B ]) + by(rule assms, blast, rule assms, blast, blast) + have 4: "\i x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ x \ B \ g i x \ \\<^sub>p " + using 3 cfs_one[of m] 2 by blast + have 5: "\i x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ g i x \ \\<^sub>p" + using 1 4 assms 2 by blast + have 6: "\ i x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>) \ 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 (Q\<^sub>p\<^bsup>m\<^esup>)" + 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 \ (\C\poly_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 \ (\C\poly_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 \\<^bsub>P\<^esub> ltrm 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 + diff --git a/thys/Padic_Field/ROOT b/thys/Padic_Field/ROOT new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/ROOT @@ -0,0 +1,20 @@ +chapter AFP + +session Padic_Field (AFP) = "Padic_Ints" + + options [timeout = 3600] +sessions + "Localization_Ring" +theories + Fraction_Field + Cring_Multivariable_Poly + Indices + Ring_Powers + Padic_Fields + Padic_Field_Polynomials + Padic_Field_Topology + Generated_Boolean_Algebra + Padic_Field_Powers + Padic_Semialgebraic_Function_Ring + document_files + "root.tex" + "root.bib" diff --git a/thys/Padic_Field/Ring_Powers.thy b/thys/Padic_Field/Ring_Powers.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/Ring_Powers.thy @@ -0,0 +1,7011 @@ +theory Ring_Powers + imports "HOL-Algebra.Chinese_Remainder" "HOL-Combinatorics.List_Permutation" + Padic_Ints.Function_Ring "HOL-Algebra.Generated_Rings" Cring_Multivariable_Poly Indices +begin + +type_synonym arity = nat +type_synonym 'a tuple = "'a list" + +(**************************************************************************************************) +(**************************************************************************************************) +section\Cartesian Powers of a Ring\ +(**************************************************************************************************) +(**************************************************************************************************) + + (**********************************************************************) + (**********************************************************************) + subsection\Constructing the Cartesian Power of a Ring\ + (**********************************************************************) + (**********************************************************************) +text\Powers of a ring\ + +text\\texttt{R\_list n R} produces the list $[R, ... , R]$ of length n\ + +fun R_list :: "nat \ ('a, 'b) ring_scheme \ (('a, 'b) ring_scheme ) list" where +"R_list n R = map (\_. R) (index_list n)" + +text\Cartesian powers of a ring\ + +definition cartesian_power :: "('a, 'b) ring_scheme \ nat \ ('a list) ring" ("_\<^bsup>_\<^esup>" 80) where +"R\<^bsup>n\<^esup> \ RDirProd_list (R_list n R)" + +lemma R_list_length: +"length (R_list n R) = n" + apply(induction n) by auto + +lemma R_list_nth: +"i < n \ R_list n R ! i = R" + by (simp add: index_list_length) + +lemma cartesian_power_car_memI: + assumes "length as = n" + assumes "set as \ carrier R" + shows "as \ carrier (R\<^bsup>n\<^esup>)" + unfolding cartesian_power_def + apply(rule RDirProd_list_carrier_memI) + using R_list_length assms(1) apply auto[1] + by (metis R_list_length R_list_nth assms(1) assms(2) nth_mem subsetD) + +lemma cartesian_power_car_memI': + assumes "length as = n" + assumes "\i. i < n \ as ! i \ carrier R" + shows "as \ carrier (R\<^bsup>n\<^esup>)" + unfolding cartesian_power_def + apply(rule RDirProd_list_carrier_memI) + using R_list_length assms(1) apply auto[1] + by (metis R_list_length R_list_nth assms(2)) + +lemma cartesian_power_car_memE: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "length as = n" + using RDirProd_list_carrier_mem(1) + by (metis R_list_length assms cartesian_power_def) + +lemma cartesian_power_car_memE': + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "i < n" + shows " as ! i \ carrier R" + using assms RDirProd_list_carrier_mem(2) + by (metis (no_types, lifting) R_list_length R_list_nth cartesian_power_def) + +lemma cartesian_power_car_memE'': + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "set as \ carrier R" + using cartesian_power_car_memE' + by (metis assms cartesian_power_car_memE in_set_conv_nth subsetI) + +lemma cartesian_power_car_memI'': + assumes "length as = n + k" + assumes "take n as \ carrier (R\<^bsup>n\<^esup>)" + assumes "drop n as \ carrier (R\<^bsup>k\<^esup>)" + shows "as \ carrier (R\<^bsup>n+k\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (simp add: assms(1)) +proof- fix i assume A: "i < n + k" + show " as ! i \ carrier R" + apply(cases "i < n") + apply (metis assms(2) cartesian_power_car_memE' nth_take) + by (metis A add_diff_inverse_nat add_less_imp_less_left + append_take_drop_id assms(2) assms(3) cartesian_power_car_memE + cartesian_power_car_memE' nth_append_length_plus) +qed + +lemma cartesian_power_cons: + assumes " as \ carrier (R\<^bsup>n\<^esup>)" + assumes "a \ carrier R" + shows "a#as \ carrier (R\<^bsup>n+1\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (metis One_nat_def assms(1) cartesian_power_car_memE list.size(4)) + by (metis assms(1) assms(2) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth set_ConsD subsetI) + +lemma cartesian_power_append: + assumes " as \ carrier (R\<^bsup>n\<^esup>)" + assumes "a \ carrier R" + shows "as@[a] \ carrier (R\<^bsup>n+1\<^esup>)" + apply(rule cartesian_power_car_memI'') + apply (metis add.commute assms(1) cartesian_power_car_memE length_append_singleton plus_1_eq_Suc) + apply (metis append_eq_append_conv_if assms(1) butlast_snoc cartesian_power_car_memE length_append_singleton lessI take_butlast) + by (metis add.commute add.right_neutral append_eq_conv_conj assms(1) assms(2) bot_least + cartesian_power_car_memE cartesian_power_car_memI cartesian_power_cons + list.set(1) list.size(3)) + +lemma cartesian_power_head: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "hd as \ carrier R" + by (metis assms cartesian_power_car_memE cartesian_power_car_memE'' list.set_sel(1) list.size(3) old.nat.distinct(1) subsetD) + +lemma cartesian_power_tail: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "tl as \ carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (metis add_diff_cancel_left' assms cartesian_power_car_memE length_tl plus_1_eq_Suc) + by (metis assms cartesian_power_car_memE cartesian_power_car_memE'' list.set_sel(2) list.size(3) nat.simps(3) subsetD subsetI) + +lemma insert_at_index_closed: + assumes "length as = n" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "a \ carrier R" + assumes "k \ n" + shows "(insert_at_index as a k) \ carrier (R\<^bsup>Suc n\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (metis Groups.add_ac(2) assms(1) insert_at_index_length plus_1_eq_Suc) + by (smt R_list_length Suc_le_eq assms(1) assms(2) assms(3) assms(4) + cartesian_power_car_memE' insert_at_index_eq insert_at_index_eq' + insert_at_index_eq'' less_Suc_eq less_Suc_eq_0_disj not_less_eq_eq) + +lemma insert_at_index_pow_not_car: + assumes "k \n" + assumes "length x = n" + assumes "(insert_at_index x a k) \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "x \ carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (simp add: assms(2)) + by (metis Suc_mono assms(1) assms(2) assms(3) + cartesian_power_car_memE' insert_at_index_eq' + insert_at_index_eq'' leI less_SucI) + +lemma insert_at_index_pow_not_car': + assumes "k \n" + assumes "length x = n" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "(insert_at_index x a n) \ carrier (R\<^bsup>Suc n\<^esup>)" + by (metis assms(2) assms(3) insert_at_index_pow_not_car lessI less_Suc_eq_le) + +lemma take_closed: + assumes "k \n" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "take k x \ carrier (R\<^bsup>k\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (metis assms(1) assms(2) cartesian_power_car_memE length_take min.absorb_iff2) + by (meson assms(2) cartesian_power_car_memE'' set_take_subset subset_trans) + +lemma drop_closed: + assumes "k < n" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "drop k x \ carrier (R\<^bsup>n - k\<^esup>)" + apply(rule cartesian_power_car_memI[of "drop k x" "n - k"] ) + using assms(2) cartesian_power_car_memE length_drop apply blast + by (metis add_diff_inverse_nat assms(1) assms(2) cartesian_power_car_memE + cartesian_power_car_memE' in_set_conv_nth length_drop less_imp_le_nat + nat_add_left_cancel_less nth_drop order.asym subsetI) + +lemma last_closed: + assumes "n > 0" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "last x \ carrier R" + using assms + by (metis Suc_diff_1 cartesian_power_car_memE cartesian_power_car_memE' + last_conv_nth lessI list.size(3) neq0_conv) + +lemma cartesian_power_concat: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "b \ carrier (R\<^bsup>k\<^esup>)" + shows "a@b \ carrier (R\<^bsup>n+k\<^esup>)" + "b@a \ carrier (R\<^bsup>n+k\<^esup>)" + apply (metis (no_types, lifting) append_eq_conv_conj assms(1) assms(2) + cartesian_power_car_memE cartesian_power_car_memI'' length_append) + by (metis (no_types, lifting) add.commute append_eq_conv_conj assms(1) assms(2) + cartesian_power_car_memE cartesian_power_car_memI'' length_append) + +lemma cartesian_power_decomp: + assumes "a \ carrier (R\<^bsup>n+k\<^esup>)" + obtains a0 a1 where "a0 \ carrier (R\<^bsup>n\<^esup>) \ a1 \ carrier (R\<^bsup>k\<^esup>) \ a0@a1 = a" + using assms + by (metis (no_types, lifting) add_diff_cancel_left' append.assoc append_eq_append_conv + append_take_drop_id cartesian_power_car_memE drop_closed le_add1 + le_neq_implies_less length_append take_closed) + +lemma list_segment_pow: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "j \ n" + assumes "i \ j" + shows "list_segment i j as \ carrier (R\<^bsup>j - i\<^esup>)" + apply(rule cartesian_power_car_memI) + using list_segment_length assms cartesian_power_car_memE + apply blast + using assms + by (metis cartesian_power_car_memE cartesian_power_car_memE'' + dual_order.trans list_segment_subset_list_set) + +lemma nth_list_segment: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "j \n" + assumes "i \ j" + assumes "k < j - i" + shows "(list_segment i j as) ! k = as ! (i + k)" + unfolding list_segment_def + using assms nth_map_upt[of k j i "((!) as)" ] + by blast + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Mapping the Carrier of a Ring to its 1-Dimensional Cartesian Power.\ +(**************************************************************************************************) +(**************************************************************************************************) + +context cring +begin + +lemma R1_carI: + assumes "length as = 1" + assumes "as!0 \ carrier R" + shows "as \ carrier (R\<^bsup>1\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms + apply blast + using assms + by (metis in_set_conv_nth less_one subsetI) + +abbreviation(input) to_R1 where +"to_R1 a \ [a]" + +abbreviation(input) to_R :: "'a list \ 'a" where +"to_R as \ as!0" + +lemma to_R1_to_R: + assumes "a \ carrier (R\<^bsup>1\<^esup>)" + shows "to_R1 (to_R a) = a" +proof- + have "length a = 1" + using assms cartesian_power_car_memE by blast + then obtain b where "a = [b]" + by (metis One_nat_def length_0_conv length_Suc_conv) + then show ?thesis + using assms + by (metis nth_Cons_0) +qed + +lemma to_R_to_R1: + shows "to_R (to_R1 a) = a" + by (meson nth_Cons_0) + +lemma to_R1_closed: + assumes "a \ carrier R" + shows "to_R1 a \ carrier (R\<^bsup>1\<^esup>)" +proof(rule R1_carI) + show "length [a] = 1" + by simp + show "[a] ! 0 \ carrier R" + using assms to_R_to_R1 by presburger +qed + +lemma to_R_pow_closed: + assumes "a \ carrier (R\<^bsup>1\<^esup>)" + shows "to_R a \ carrier R" + using assms cartesian_power_car_memE' by blast + +lemma to_R1_intersection: + assumes "A \ carrier R" + assumes "B \ carrier R" + shows "to_R1 ` (A \ B) = to_R1` A \ to_R1 ` B" +proof + show "(\a. [a]) ` (A \ B) \ (\a. [a]) ` A \ (\a. [a]) ` B" + by blast + show "(\a. [a]) ` A \ (\a. [a]) ` B \ (\a. [a]) ` (A \ B)" + using assms + by blast +qed + +lemma to_R1_finite: + assumes "finite A" + shows "finite (to_R1` A)" + "card A = card (to_R1` A)" + using assms + apply blast + apply(rule finite.induct[of A]) + apply (simp add: assms(1)) + apply simp + by (smt card_insert_if finite_imageI image_iff image_insert list.inject) + +lemma to_R1_carrier: +"to_R1` (carrier R)= carrier (R\<^bsup>1\<^esup>)" +proof + show "(\a. [a]) ` carrier R \ carrier (R\<^bsup>1\<^esup>)" + proof fix x + assume "x \ (\a. [a]) ` carrier R" + then show "x \ carrier (R\<^bsup>1\<^esup>)" + using cartesian_power_car_memI[of x 1 R] + by (metis (no_types, lifting) image_iff to_R1_closed) + qed + show "carrier (R\<^bsup>1\<^esup>) \ (\a. [a]) ` carrier R" + proof fix x + assume "x \ carrier (R\<^bsup>1\<^esup>)" + then obtain a where a_def: "a \ carrier R \ x = [a]" + using cartesian_power_car_memE'[of x R 1] cartesian_power_car_memE[of x R 1] + by (metis less_numeral_extra(1) to_R1_to_R) + then show "x \ (\a. [a]) ` carrier R" + by blast + qed +qed + +lemma to_R1_diff: +"to_R1` (A - B) = to_R1` A - to_R1` B" +proof + show "(\a. [a]) ` (A - B) \ (\a. [a]) ` A - (\a. [a]) ` B" + by blast + show "(\a. [a]) ` A - (\a. [a]) ` B \ (\a. [a]) ` (A - B)" + by blast +qed + +lemma to_R1_complement: + shows "to_R1` (carrier R - A) = carrier (R\<^bsup>1\<^esup>) - to_R1` A" + by (metis to_R1_carrier to_R1_diff) + +lemma to_R1_subset: + assumes "A \ B" + shows "to_R1` A \ to_R1` B" + using assms + by blast + +lemma to_R1_car_subset: + assumes "A \ carrier R" + shows "to_R1` A \ carrier (R\<^bsup>1\<^esup>)" + using assms to_R1_carrier + by blast +end + + (**********************************************************************) + (**********************************************************************) + subsection\Simple Cartesian Products\ + (**********************************************************************) + (**********************************************************************) +definition cartesian_product :: "('a list) set \ ('a list) set \ ('a list) set" where +"cartesian_product A B \ {xs. \as \ A. \bs \ B. xs = as@bs}" + +lemma cartesian_product_closed: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + shows "cartesian_product A B \ carrier (R\<^bsup>n + m\<^esup>)" +proof + fix x + assume A: "x \ cartesian_product A B " + then obtain as bs where as_bs_def: "x = as@bs \ as \ A \ bs \ B" + unfolding cartesian_product_def by blast + show "x \ carrier (R\<^bsup>n + m\<^esup>) " + apply(rule cartesian_power_car_memI') + apply (metis as_bs_def assms cartesian_power_car_memE length_append subsetD) + using A unfolding cartesian_product_def + by (metis (no_types, lifting) add_diff_inverse_nat as_bs_def assms(1) + assms(2) cartesian_power_car_memE cartesian_power_car_memE' + nat_add_left_cancel_less nth_append subsetD) +qed + +lemma cartesian_product_closed': + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "b \ carrier (R\<^bsup>m\<^esup>)" + shows "(a@b) \ carrier (R\<^bsup>n + m\<^esup>)" +proof- + have "a@b \ cartesian_product {a} {b}" + using cartesian_product_def by blast + then show ?thesis + using cartesian_product_closed[of "{a}" R n "{b}" m] + assms + by blast +qed + +lemma cartesian_product_carrier: +"cartesian_product (carrier (R\<^bsup>n\<^esup>)) (carrier (R\<^bsup>m\<^esup>)) = carrier (R\<^bsup>n + m\<^esup>)" +proof + show "cartesian_product (carrier (R\<^bsup>n\<^esup>)) (carrier (R\<^bsup>m\<^esup>)) \ carrier (R\<^bsup>n + m\<^esup>)" + using cartesian_product_closed[of "(carrier (R\<^bsup>n\<^esup>))" R n "(carrier (R\<^bsup>m\<^esup>)) " m] + by blast + show "carrier (R\<^bsup>n + m\<^esup>) \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) (carrier (R\<^bsup>m\<^esup>))" + proof + fix x + assume A: "x \ carrier (R\<^bsup>n + m\<^esup>)" + have 0: "take n x \ carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (metis A cartesian_power_car_memE le_add1 length_take min.absorb2) + by (metis A add.commute cartesian_power_car_memE' + nth_take trans_less_add2) + have 1: "drop n x \ carrier (R\<^bsup>m\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (metis A add_diff_cancel_left' cartesian_power_car_memE length_drop) + by (metis A cartesian_power_car_memE cartesian_power_car_memE' le_add1 nat_add_left_cancel_less nth_drop) + show "x \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) (carrier (R\<^bsup>m\<^esup>))" + using 0 1 + by (smt A cartesian_power_decomp cartesian_product_def mem_Collect_eq) +qed + + +text\Higher function rings\ +qed + +lemma cartesian_product_memI: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "take n a \ A" + assumes "drop n a \ B" + shows "a \ cartesian_product A B" +proof- + have "a = (take n a) @ (drop n a)" + by (metis append_take_drop_id) + then show ?thesis + using assms(3) assms(4) cartesian_product_def by blast +qed + +lemma cartesian_product_memI': + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "a \ A" + assumes "b \ B" + shows "a@b \ cartesian_product A B" + using assms unfolding cartesian_product_def + by blast + +lemma cartesian_product_memE: +assumes "a \ cartesian_product A B" +assumes "A \ carrier (R\<^bsup>n\<^esup>)" +shows "take n a \ A" + "drop n a \ B" + using assms unfolding cartesian_product_def + apply (smt append_eq_conv_conj cartesian_power_car_memE in_mono mem_Collect_eq) + using assms unfolding cartesian_product_def + by (smt append_eq_conv_conj cartesian_power_car_memE in_mono mem_Collect_eq) + +lemma cartesian_product_intersection: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "C \ carrier (R\<^bsup>n\<^esup>)" + assumes "D \ carrier (R\<^bsup>m\<^esup>)" + shows "cartesian_product A B \ cartesian_product C D = cartesian_product (A \ C) (B \ D)" +proof + show "cartesian_product A B \ cartesian_product C D \ cartesian_product (A \ C) (B \ D)" + proof fix x + assume "x \ cartesian_product A B \ cartesian_product C D" + then show "x \ cartesian_product (A \ C) (B \ D)" + using assms cartesian_product_memE[of x C D] cartesian_product_memE[of x A B] + cartesian_product_memI[of "A \ C" R n "B \ D" m x] + by (smt Int_iff inf.coboundedI1) + qed + show "cartesian_product (A \ C) (B \ D) \ cartesian_product A B \ cartesian_product C D" + proof fix x + assume "x \ cartesian_product (A \ C) (B \ D)" + then show "x \ cartesian_product A B \ cartesian_product C D" + using assms cartesian_product_memI[of C R n D m] cartesian_product_memI[of A R n B m] + cartesian_product_memE[of x "A \ B" "C \ D" R n ] + by (metis (no_types, lifting) Int_iff cartesian_product_memE(1) cartesian_product_memE(2) inf_le1 subset_trans) + qed +qed + +lemma cartesian_product_subsetI: + assumes "C \ A" + assumes "D \ B" + shows "cartesian_product C D \ cartesian_product A B" + using assms unfolding cartesian_product_def + by blast + +lemma cartesian_product_binary_union_right: + assumes "C \ carrier (R\<^bsup>n\<^esup>)" + assumes "D \ carrier (R\<^bsup>n\<^esup>)" + shows "cartesian_product A (C \ D) = (cartesian_product A C) \ (cartesian_product A D)" +proof + show "cartesian_product A (C \ D) \ cartesian_product A C \ cartesian_product A D" + unfolding cartesian_product_def by blast + show "cartesian_product A C \ cartesian_product A D \ cartesian_product A (C \ D)" + unfolding cartesian_product_def by blast +qed + +lemma cartesian_product_binary_union_left: + assumes "C \ carrier (R\<^bsup>n\<^esup>)" + assumes "D \ carrier (R\<^bsup>n\<^esup>)" + shows "cartesian_product (C \ D) A = (cartesian_product C A) \ (cartesian_product D A)" +proof + show "cartesian_product (C \ D) A \ cartesian_product C A \ cartesian_product D A" + unfolding cartesian_product_def by blast + show "cartesian_product C A \ cartesian_product D A \ cartesian_product (C \ D) A" + unfolding cartesian_product_def by blast +qed + +lemma cartesian_product_binary_intersection_right: + assumes "C \ carrier (R\<^bsup>n\<^esup>)" + assumes "D \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R\<^bsup>m\<^esup>)" + shows "cartesian_product A (C \ D) = (cartesian_product A C) \ (cartesian_product A D)" +proof + show "cartesian_product A (C \ D) \ cartesian_product A C \ cartesian_product A D" + unfolding cartesian_product_def by blast + show "cartesian_product A C \ cartesian_product A D \ cartesian_product A (C \ D)" + proof fix x assume A: "x \ cartesian_product A C \ cartesian_product A D" + show "x \ cartesian_product A (C \ D)" apply(rule cartesian_product_memI[of A R m _ n ]) + apply (simp add: assms(3)) + apply (simp add: assms(1) inf.coboundedI1) + apply (meson A IntD1 assms(3) cartesian_product_memE(1)) + by (meson A Int_iff assms(3) cartesian_product_memE(2)) + qed +qed + +lemma cartesian_product_binary_intersection_left: + assumes "C \ carrier (R\<^bsup>n\<^esup>)" + assumes "D \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R\<^bsup>m\<^esup>)" + shows "cartesian_product (C \ D) A = (cartesian_product C A) \ (cartesian_product D A)" +proof + show "cartesian_product (C \ D) A \ cartesian_product C A \ cartesian_product D A" + unfolding cartesian_product_def by blast + show "cartesian_product C A \ cartesian_product D A \ cartesian_product (C \ D) A" + proof fix x assume A: "x \ cartesian_product C A \ cartesian_product D A" + show "x \ cartesian_product (C \ D) A" apply(rule cartesian_product_memI[of _ R n _ m ]) + apply (simp add: assms(2) inf.coboundedI2) + apply (simp add: assms(3)) + apply (meson A Int_iff assms(1) assms(2) cartesian_product_memE(1)) + by (meson A IntD1 assms(1) cartesian_product_memE(2)) + qed +qed + +lemma cartesian_product_car_complement_right: + assumes "A \ carrier (R\<^bsup>m\<^esup>)" + shows "carrier (R\<^bsup>n + m\<^esup>) - cartesian_product (carrier (R\<^bsup>n\<^esup>)) A = + cartesian_product (carrier (R\<^bsup>n\<^esup>)) ((carrier (R\<^bsup>m\<^esup>)) - A)" +proof + show "carrier (R\<^bsup>n + m\<^esup>) - cartesian_product (carrier (R\<^bsup>n\<^esup>)) A \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) ((carrier (R\<^bsup>m\<^esup>)) - A)" + proof fix x assume A: "x \ (carrier (R\<^bsup>n + m\<^esup>) - cartesian_product (carrier (R\<^bsup>n\<^esup>)) A)" + show "x \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) ((carrier (R\<^bsup>m\<^esup>)) - A)" + apply(rule cartesian_product_memI[of _ R n _ m]) + apply simp + apply simp + apply (meson A DiffE le_add1 take_closed) + apply(rule ccontr) + proof- + assume A': "drop n x \ (carrier (R\<^bsup>m\<^esup>) - A)" + have "drop n x \ A" + proof- + have "x \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) (carrier (R\<^bsup>m\<^esup>))" + using A + by (metis (mono_tags, lifting) DiffD1 cartesian_product_carrier) + then show ?thesis + using A' cartesian_product_memE[of x "(carrier (R\<^bsup>n\<^esup>))" "(carrier (R\<^bsup>m\<^esup>))" R n] + by blast + qed + then show False + using A cartesian_product_memI[of "(carrier (R\<^bsup>n\<^esup>))" R n A m x] + by (meson DiffD1 DiffD2 assms le_add1 order_refl take_closed) + qed + qed + show "cartesian_product (carrier (R\<^bsup>n\<^esup>)) ((carrier (R\<^bsup>m\<^esup>)) - A) \ carrier (R\<^bsup>n + m\<^esup>) - cartesian_product (carrier (R\<^bsup>n\<^esup>)) A" + proof fix x assume A: "x \ cartesian_product (carrier (R\<^bsup>n\<^esup>)) ((carrier (R\<^bsup>m\<^esup>)) - A)" + show "x \ carrier (R\<^bsup>n + m\<^esup>) - cartesian_product (carrier (R\<^bsup>n\<^esup>)) A" + apply(rule ccontr) + using A cartesian_product_memE[of x "carrier (R\<^bsup>n\<^esup>)" A R n] + using A cartesian_product_memE[of x "(carrier (R\<^bsup>n\<^esup>))" "(carrier (R\<^bsup>m\<^esup>)) - A" R n] + by (metis (no_types, lifting) DiffD1 DiffD2 DiffI + append_take_drop_id cartesian_product_closed' order_refl) + qed +qed + +lemma cartesian_product_car_complement_left: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + shows "carrier (R\<^bsup>n + m\<^esup>) - cartesian_product A (carrier (R\<^bsup>m\<^esup>)) = + cartesian_product ((carrier (R\<^bsup>n\<^esup>)) - A) (carrier (R\<^bsup>m\<^esup>)) " +proof + show "carrier (R\<^bsup>n + m\<^esup>) - cartesian_product A (carrier (R\<^bsup>m\<^esup>)) \ + cartesian_product ((carrier (R\<^bsup>n\<^esup>)) - A) (carrier (R\<^bsup>m\<^esup>)) " + proof fix x assume A: " x \ carrier (R\<^bsup>n + m\<^esup>) - cartesian_product A (carrier (R\<^bsup>m\<^esup>))" + show "x \ cartesian_product ((carrier (R\<^bsup>n\<^esup>)) - A) (carrier (R\<^bsup>m\<^esup>)) " + proof(rule cartesian_product_memI[of _ R n _ m]) + show "carrier (R\<^bsup>n\<^esup>) - A \ carrier (R\<^bsup>n\<^esup>)" + by simp + show "carrier (R\<^bsup>m\<^esup>) \ carrier (R\<^bsup>m\<^esup>)" + by simp + show "take n x \ carrier (R\<^bsup>n\<^esup>) - A" + by (metis (no_types, lifting) A DiffD1 DiffD2 DiffI assms + cartesian_product_carrier cartesian_product_memE(2) cartesian_product_memI + le_add1 order_refl take_closed) + show "drop n x \ carrier (R\<^bsup>m\<^esup>)" + by (metis A DiffD1 cartesian_product_carrier cartesian_product_memE(2) order_refl) + qed + qed + show "cartesian_product ((carrier (R\<^bsup>n\<^esup>)) - A) (carrier (R\<^bsup>m\<^esup>)) \ + carrier (R\<^bsup>n + m\<^esup>) - cartesian_product A (carrier (R\<^bsup>m\<^esup>)) " + proof fix x assume A: " x \ cartesian_product ((carrier (R\<^bsup>n\<^esup>)) - A) (carrier (R\<^bsup>m\<^esup>))" + show "x \ carrier (R\<^bsup>n + m\<^esup>) - cartesian_product A (carrier (R\<^bsup>m\<^esup>))" + apply(rule ccontr) + using A cartesian_product_memE[of x "((carrier (R\<^bsup>n\<^esup>)) - A)" "(carrier (R\<^bsup>m\<^esup>))"] + cartesian_product_memE[of x A "(carrier (R\<^bsup>m\<^esup>))"] + by (smt DiffD1 DiffD2 DiffI Diff_subset append_take_drop_id assms cartesian_product_closed') + qed +qed + +lemma cartesian_product_complement_right: + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + shows "cartesian_product A (carrier (R\<^bsup>m\<^esup>)) - (cartesian_product A B) = + cartesian_product A ((carrier (R\<^bsup>m\<^esup>)) - B)" +proof + show "cartesian_product A (carrier (R\<^bsup>m\<^esup>)) - cartesian_product A B \ cartesian_product A ((carrier (R\<^bsup>m\<^esup>)) - B)" + unfolding cartesian_product_def by blast + show "cartesian_product A ((carrier (R\<^bsup>m\<^esup>)) - B) \ cartesian_product A ((carrier (R\<^bsup>m\<^esup>))) - cartesian_product A B" + proof fix x assume A: "x \ cartesian_product A ((carrier (R\<^bsup>m\<^esup>)) - B)" + have 0: "x \ cartesian_product A (carrier (R\<^bsup>m\<^esup>))" + using A unfolding cartesian_product_def by blast + show "x \ cartesian_product A (carrier (R\<^bsup>m\<^esup>)) - cartesian_product A B " + apply(rule ccontr) + using assms 0 A cartesian_product_memE[of x A "((carrier (R\<^bsup>m\<^esup>)) - B)" R n] + cartesian_product_memE[of x A B R n] + by blast + qed +qed + +lemma cartesian_product_complement_left: + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + shows "cartesian_product (carrier (R\<^bsup>m\<^esup>)) A - (cartesian_product B A) = + cartesian_product ((carrier (R\<^bsup>m\<^esup>)) - B) A " +proof + show "cartesian_product (carrier (R\<^bsup>m\<^esup>)) A - cartesian_product B A \ cartesian_product ((carrier (R\<^bsup>m\<^esup>)) - B) A" + unfolding cartesian_product_def by blast + show "cartesian_product ((carrier (R\<^bsup>m\<^esup>)) - B) A \ cartesian_product (carrier (R\<^bsup>m\<^esup>)) A - cartesian_product B A" + proof fix x assume A: "x \ cartesian_product ((carrier (R\<^bsup>m\<^esup>)) - B) A" + have 0: "x \ cartesian_product (carrier (R\<^bsup>m\<^esup>)) A" + using A unfolding cartesian_product_def by blast + have 1: "take m x \ (carrier (R\<^bsup>m\<^esup>)) - B" + using A cartesian_product_memE[of x "((carrier (R\<^bsup>m\<^esup>)) - B)" A R m] + by blast + have 2: "drop m x \ A" + using cartesian_product_memE[of x "((carrier (R\<^bsup>m\<^esup>)) - B)" A R m] + by (metis A Diff_subset) + show "x \ cartesian_product (carrier (R\<^bsup>m\<^esup>)) A - cartesian_product B A" + apply(rule ccontr) + using A 0 1 2 cartesian_product_memE[of x B A R m] assms + by blast + qed +qed + +lemma cartesian_product_empty_right: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B = {[]}" + shows "cartesian_product A B = A" +proof + show "cartesian_product A B \ A" + using assms unfolding cartesian_product_def + by (smt append_Nil2 mem_Collect_eq singletonD subsetI) + show "A \ cartesian_product A B" + using assms unfolding cartesian_product_def + by blast +qed + +lemma cartesian_product_empty_left: + assumes "B \ carrier (R\<^bsup>n\<^esup>)" + assumes "A = {[]}" + shows "cartesian_product A B = B" +proof + show "cartesian_product A B \ B" + using assms unfolding cartesian_product_def + by (smt append.simps(1) mem_Collect_eq singletonD subsetI) + show "B \ cartesian_product A B" + using assms unfolding cartesian_product_def + by blast +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Cartesian Products at Arbitrary Indices\ + (**********************************************************************) + (**********************************************************************) + +definition(in ring) ring_pow_proj :: "nat \ (nat set) \ ('a list) \ ('a list) " ("\\<^bsub>_, _\<^esub>") where +"ring_pow_proj n S \ restrict (project_at_indices S) (carrier (R\<^bsup>n\<^esup>))" + +text\The projection at an arbitrary index set\ + +lemma project_at_indices_closed: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "S \ indices_of a" + shows "\\<^bsub>S\<^esub> a \ carrier (R\<^bsup>card S\<^esup>)" + apply(rule cartesian_power_car_memI') + using assms proj_at_index_list_length apply blast + using assms project_at_indices_nth[of S] + by (smt cartesian_power_car_memE cartesian_power_car_memE' indices_of_def lessThan_iff nth_elem_closed subsetD) + +lemma(in ring) ring_pow_proj_is_map: + assumes "S \ {..\<^bsub>n,S\<^esub> \ struct_maps (R\<^bsup>n\<^esup>) (R\<^bsup>card S\<^esup>)" +proof(rule struct_maps_memI) + show "\x. x \ carrier (R\<^bsup>n\<^esup>) \ \\<^bsub>n,S\<^esub> x \ carrier (R\<^bsup>card S\<^esup>)" + using project_at_indices_closed unfolding ring_pow_proj_def + by (metis assms cartesian_power_car_memE indices_of_def restrict_apply') + show " \x. x \ carrier (R\<^bsup>n\<^esup>) \ \\<^bsub>n, S\<^esub> x = undefined" + by (metis restrict_apply ring_pow_proj_def) +qed + +lemma(in ring) project_at_indices_ring_pow_proj: + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "\\<^bsub>S\<^esub> x = \\<^bsub>n,S\<^esub> x" + unfolding ring_pow_proj_def + by (metis assms restrict_apply') + +text\ + Cartesian products where the first factor \A\ occurs at the entries of some arbitrary index set. + Note that this product isn't completely arbitrary because the entries of the factor of \A\ + still occurs in ascending order.\ + +definition twisted_cartesian_product ("Prod\<^bsub>_, _\<^esub>") where +"twisted_cartesian_product S S' A B = {a . length a = card S + card S' \ \\<^bsub>S\<^esub> a \ A \ \\<^bsub>S'\<^esub> a \ B}" + +lemma twisted_cartesian_product_mem_length: + assumes "card S = n" + assumes "card S' = m" + assumes "a \ Prod\<^bsub>S,S'\<^esub> A B" + shows "length a = n + m" + using assms unfolding twisted_cartesian_product_def + by blast + +lemma twisted_cartesian_product_closed: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "card S = n" + assumes "card S' = m" + assumes "S \ S' = {.. carrier (R\<^bsup>n + m\<^esup>)" +proof(rule subsetI) + fix x assume A: "x \ twisted_cartesian_product S S' A B" + show "x \ carrier (R\<^bsup>n + m\<^esup>)" + proof(rule cartesian_power_car_memI') + show "length x = n + m" + using twisted_cartesian_product_mem_length \x \ twisted_cartesian_product S S' A B\ assms(1) assms(2) assms(3) assms(4) assms(5) by blast + fix i assume A': "i < n + m" + have 0: "indices_of x = {..length x = n + m\ indices_of_def) + show "x ! i \ carrier R" + proof(cases "i \ S") + case True + have "x!i = \\<^bsub>S\<^esub>x ! (set_rank S i)" + using A' 0 assms + by (metis True Un_upper1 project_at_indices_nth') + then show ?thesis + using project_at_indices_closed[of x R "n + m" S] A A' + cartesian_power_car_memE'[of "\\<^bsub>S\<^esub> x" R "card S"] + by (metis (no_types, lifting) True UnI2 Un_upper1 assms(1) assms(3) assms(5) + finite_lessThan finite_subset mem_Collect_eq set_rank_range sup.absorb_iff1 + twisted_cartesian_product_def) + next + case False + have "x!i = \\<^bsub>S'\<^esub>x ! (set_rank S' i)" + using A' 0 assms + by (metis False UnE lessThan_iff project_at_indices_nth' sup.absorb_iff1 sup.right_idem) + then show ?thesis + using project_at_indices_closed[of x R "n + m" S'] A A' + cartesian_power_car_memE'[of "\\<^bsub>S'\<^esub> x" R "card S'"] + by (metis (no_types, lifting) False UnE UnI2 Un_upper2 assms(2) assms(4) assms(5) + finite_lessThan finite_subset lessThan_iff mem_Collect_eq set_rank_range sup.absorb_iff1 + twisted_cartesian_product_def) + qed + qed +qed + +lemma twisted_cartesian_product_memE: + assumes "a \ twisted_cartesian_product S S' A B" + shows "\\<^bsub>S\<^esub> a \ A" "\\<^bsub>S'\<^esub> a \ B" + using assms(1) unfolding twisted_cartesian_product_def apply blast + using assms(1) unfolding twisted_cartesian_product_def by blast + +lemma twisted_cartesian_product_memI: + assumes "\\<^bsub>S\<^esub> a \ A" + assumes "\\<^bsub>S'\<^esub> a \ B" + assumes "length a = card S + card S'" + shows "a \ twisted_cartesian_product S S' A B" + by (metis (mono_tags, lifting) assms(1) assms(2) assms(3) mem_Collect_eq twisted_cartesian_product_def) + +lemma twisted_cartesian_product_empty_left_factor: + assumes "A = {}" + shows "twisted_cartesian_product S S' A B = {}" + by (metis assms emptyE equals0I twisted_cartesian_product_memE(1)) + +lemma twisted_cartesian_product_empty_right_factor: + assumes "B = {}" + shows "twisted_cartesian_product S S' A B = {}" + by (metis assms emptyE equals0I twisted_cartesian_product_memE(2)) + +lemma twisted_cartesian_project_left: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "A \ {}" + assumes "B \ {}" + assumes "card S = n" + assumes "card S' = m" + assumes "S \ S' = {..\<^bsub>S\<^esub> ` (Prod\<^bsub>S,S'\<^esub> A B) = A" +proof + have f0: "S \ S' = {}" + proof- + have "card (S \ S') = card S + card S'" + by (simp add: assms(5) assms(6) assms(7)) + thus ?thesis + by (metis Nat.add_diff_assoc2 add.right_neutral add_diff_cancel_left' assms(6) + assms(7) card_0_eq card_Un_Int finite_Int finite_Un finite_lessThan le_add1) +qed + show "\\<^bsub>S\<^esub> ` (Prod\<^bsub>S,S'\<^esub> A B) \ A" + unfolding twisted_cartesian_product_def + by blast + show "A \ \\<^bsub>S\<^esub> ` (Prod\<^bsub>S,S'\<^esub> A B)" + proof fix x assume A: "x \ A" + obtain y where y_def: "y \ B" + using assms(4) by blast + obtain a where a_def: + "a = map (\i. if i \ S then (x ! set_rank S i) else (y ! set_rank S' i)) [0.. indices_of a" + by (metis (no_types, lifting) Un_upper1 a_def assms(7) diff_zero indices_of_def length_map length_upt) + have 1: "S' \ indices_of a" + by (metis (no_types, lifting) Un_upper2 a_def assms(7) diff_zero indices_of_def length_map length_upt) + have 2: "\\<^bsub>S\<^esub> a = x" + proof- + have 20: "length (\\<^bsub>S\<^esub> a) = n" + by (metis (no_types, lifting) Un_upper1 a_def assms(5) assms(7) diff_zero indices_of_def length_map length_upt proj_at_index_list_length) + have "\i. i < n \ \\<^bsub>S\<^esub> a ! i = x ! i" + proof- fix i assume A: "i < n" show "\\<^bsub>S\<^esub> a ! i = x ! i" + using 0 assms a_def project_at_indices_nth'[of S a "nth_elem S i"] set_rank_nth_elem_inv[of S i] + nth_map[of i "[0..\<^bsub>S'\<^esub> a = y" + proof- + have 20: "length (\\<^bsub>S'\<^esub> a) = m" + using "1" assms(6) proj_at_index_list_length by blast + have "\i. i < m \ \\<^bsub>S'\<^esub> a ! i = y ! i" + proof- fix i assume A: "i < m" + have "nth_elem S' i \ S" + using nth_elem_closed[of i S'] f0 A assms(6) by blast + thus "\\<^bsub>S'\<^esub> a ! i = y ! i" + using 0 assms a_def project_at_indices_nth'[of S' a "nth_elem S' i"] set_rank_nth_elem_inv[of S' i] + nth_map[of i "[0.. (Prod\<^bsub>S,S'\<^esub> A B)" + apply(rule twisted_cartesian_product_memI) + apply (simp add: "2" A) + apply (simp add: "3" y_def) + by (metis (no_types, lifting) a_def assms(5) assms(6) diff_zero length_map length_upt) + thus "x \ \\<^bsub>S\<^esub> ` (Prod\<^bsub>S,S'\<^esub> A B)" + using "2" by blast + qed +qed + +lemma twisted_cartesian_product_swap: + shows "(Prod\<^bsub>S,S'\<^esub> A B) = (Prod\<^bsub>S',S\<^esub> B A)" + unfolding twisted_cartesian_product_def + by (metis add.commute) + +lemma twisted_cartesian_project_right: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>m\<^esup>)" + assumes "A \ {}" + assumes "B \ {}" + assumes "card S = n" + assumes "card S' = m" + assumes "S \ S' = {..\<^bsub>S'\<^esub> ` (Prod\<^bsub>S,S'\<^esub> A B) = B" + using assms twisted_cartesian_project_left[of B R m A n S' S] twisted_cartesian_product_swap + by (metis add.commute sup_commute) + +text \ + Cartesian products which send points $a = (a_1, \dots, a_{m})$ and $b = (b_1, \dots, b_{n})$ to + the point $(a_1, \dots, a_i, b_1, \dots, b_{n},a_{i+1}, \dots, a_m)$ +\ +definition splitting_permutation :: "nat \ nat \ nat \ + nat \ nat" where +"splitting_permutation l1 l2 i j = (if j < i then j else + (if i \ j \ j < l1 then (l2 + j) else + (if j < l1 + l2 then j - l1 + i else j)))" + +lemma splitting_permutation_case_1_unique: + assumes "i \ l1" + assumes "y < i" + assumes "splitting_permutation l1 l2 i j = y" + shows "j = y" + unfolding splitting_permutation_def + using assms(2) assms(3) splitting_permutation_def by auto + +lemma splitting_permutation_case_1_exists: + assumes "i \ l1" + assumes "y < i" + shows "splitting_permutation l1 l2 i y = y" + unfolding splitting_permutation_def + by (simp add: assms(2)) + +lemma splitting_permutation_case_2_unique: + assumes "i \ l1" + assumes "i \ y \ y < l2 + i" + assumes "splitting_permutation l1 l2 i j = y" + shows "j = y + l1 - i" + unfolding splitting_permutation_def + using assms(1) assms(2) assms(3) le_add_diff_inverse2 not_less_iff_gr_or_eq + splitting_permutation_def trans_less_add2 by auto + +lemma splitting_permutation_case_2_exists: + assumes "i \ l1" + assumes "i \ y \ y < l2 + i" + shows "splitting_permutation l1 l2 i (y + l1 - i) = y" + unfolding splitting_permutation_def + using assms(1) assms(2) less_diff_conv2 by auto + +lemma splitting_permutation_case_3_unique: + assumes "i \ l1" + assumes "l2 + i \ y \ y < l1 + l2" + assumes "splitting_permutation l1 l2 i j = y" + shows "j = y - l2" + unfolding splitting_permutation_def + by (smt Nat.le_diff_conv2 add_diff_cancel_left' add_diff_cancel_right' add_leD2 + assms(2) assms(3) le_add1 le_diff_iff not_le splitting_permutation_def) + +lemma splitting_permutation_case_3_exists: + assumes "i \ l1" + assumes "l2 + i \ y \ y < l1 + l2" + shows "splitting_permutation l1 l2 i (y - l2) = y" + unfolding splitting_permutation_def + by (metis Nat.le_diff_conv2 add.commute add_leD1 assms(2) leD le_add_diff_inverse less_diff_conv2) + +lemma splitting_permutation_case_4_unique: + assumes "i \ l1" + assumes "l1 + l2 \ y" + assumes "splitting_permutation l1 l2 i j = y" + shows "j = y" + using assms(1) assms(2) assms(3) le_add_diff_inverse2 less_le_trans + splitting_permutation_def by auto + +lemma splitting_permutation_case_4_exists: + assumes "i \ l1" + assumes "l1 + l2 \y" + shows "splitting_permutation l1 l2 i y = y" + unfolding splitting_permutation_def + using assms(2) by auto + +lemma splitting_permutation_permutes: + assumes "i \ l1" + shows "(splitting_permutation l1 l2 i) permutes {..< l1 + l2}" +proof- + have 0: "(\x. x \ {.. splitting_permutation l1 l2 i x = x)" + proof fix x show "x \ {.. splitting_permutation l1 l2 i x = x" + proof assume A: "x \ {..y. \!x. splitting_permutation l1 l2 i x = y)" + proof fix y + show "\!x. splitting_permutation l1 l2 i x = y" + proof(cases "y < i") + case True + then show ?thesis + using splitting_permutation_case_1_exists splitting_permutation_case_1_unique assms + by (metis splitting_permutation_def) + next + case F0: False + show ?thesis + proof(cases "i \ y \ y < l2 + i") + case True + then show ?thesis + using F0 splitting_permutation_case_2_exists splitting_permutation_case_2_unique assms + by metis + next + case F1: False + show ?thesis + proof(cases "l2 + i \ y \ y < l1 + l2") + case True + then show ?thesis + using F0 F1 splitting_permutation_case_3_exists splitting_permutation_case_3_unique assms + by metis + next + case F2: False + show ?thesis + using F0 F1 F2 splitting_permutation_case_4_exists splitting_permutation_case_4_unique assms + by (metis leI not_less) + qed + qed + qed + qed + show ?thesis + using 0 1 permutes_def + by blast +qed + +lemma splitting_permutation_action: + assumes "i \l1" + assumes "length a1 = l1" + assumes "length a2 = l2" + shows "permute_list (splitting_permutation l1 l2 i) ((take i a1) @ a2 @ (drop i a1)) = + a1@a2" +proof- + obtain x where x_def: "x = permute_list (splitting_permutation l1 l2 i) ((take i a1) @ a2 @ (drop i a1))" + by blast + obtain y where y_def: "y = a1 @ a2" + by blast + have 0: "length x = length y" + using x_def y_def assms splitting_permutation_permutes[of i l1 l2] + by (smt add.commute add.left_commute le_add_diff_inverse length_append + length_drop length_permute_list length_take min.absorb2) + have 1: "\i. i < l1 + l2 \ x ! i = y ! i" + proof- fix j assume A: "j < l1 + l2" + show "x ! j = y ! j" + apply(cases "j < i") + apply (smt "0" A append_take_drop_id assms(1) assms(2) assms(3) length_append length_permute_list length_take less_le_trans min.absorb2 nth_append permute_list_nth splitting_permutation_case_1_exists splitting_permutation_permutes x_def y_def) + apply(cases "i \ j \ j < l1") + apply (smt "0" A add.left_commute append_take_drop_id assms(1) assms(2) assms(3) le_add_diff_inverse length_append length_permute_list length_take min.absorb2 nth_append nth_append_length_plus permute_list_nth splitting_permutation_def splitting_permutation_permutes x_def y_def) + using x_def y_def assms + by (smt "0" A add.commute add_diff_cancel_left' add_diff_inverse_nat length_append length_permute_list length_take less_diff_conv min.absorb2 not_le nth_append permute_list_nth splitting_permutation_case_1_unique splitting_permutation_def splitting_permutation_permutes) + qed + have 2: "length x = l1 + l2" + by (simp add: x_def assms(2) assms(3)) + have 3: "x = y" + using 0 1 2 + by (metis nth_equalityI) + then show ?thesis + using x_def y_def + by blast +qed + +definition scp_permutation where +"scp_permutation l1 l2 i = fun_inv (splitting_permutation l1 l2 i)" + +lemma scp_permutation_action: + assumes "i \l1" + assumes "length a1 = l1" + assumes "length a2 = l2" + shows "permute_list (scp_permutation l1 l2 i) (a1@a2) = ((take i a1) @ a2 @ (drop i a1))" +proof- + have "(scp_permutation l1 l2 i) \ (splitting_permutation l1 l2 i) = id" + by (metis assms(1) fun_inv_def permutes_inv_o(2) scp_permutation_def splitting_permutation_permutes) + then have "permute_list ((scp_permutation l1 l2 i) \ (splitting_permutation l1 l2 i) ) ((take i a1) @ a2 @ (drop i a1)) = + ((take i a1) @ a2 @ (drop i a1))" + by (metis permute_list_id) + then show ?thesis using splitting_permutation_action permute_list_compose + by (smt \scp_permutation l1 l2 i \ splitting_permutation l1 l2 i = id\ assms(1) + assms(2) assms(3) fun_inv_def length_append length_permute_list permutes_inv permutes_inv_o(1) scp_permutation_def splitting_permutation_permutes) +qed + +lemma scp_permutes: + assumes "i \l1" + shows "(scp_permutation l1 l2 i) permutes {.. A" + assumes "b \ B" + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "length a1 = i" + shows "a1@b@a2 \ split_cartesian_product l1 l2 i A B" +proof- + have P: "a1@a2@b \ cartesian_product A B" + by (metis append.assoc assms(1) assms(2) assms(3) assms(4) cartesian_product_memI') + have 0: "i \ l1" + using assms + by (metis cartesian_power_car_memE le_add1 length_append subset_iff) + have 1: "length (a1@a2) = l1" + using assms(1) assms(3) cartesian_power_car_memE + by blast + have 2: "length b = l2" + using assms(2) assms(4) cartesian_power_car_memE + by blast + have 3: "take i (a1 @ a2) = a1" + by (simp add: assms(5)) + have 4: "drop i (a1 @ a2) = a2" + by (simp add: assms(5)) + have "permute_list (scp_permutation l1 l2 i) ((a1 @ a2) @ b) = take i (a1 @ a2) @ b @ drop i (a1 @ a2)" + using 0 1 2 scp_permutation_action[of i l1 "a1@a2" b l2] + by blast + then have "permute_list (scp_permutation l1 l2 i) ((a1@a2)@b) = a1@b@a2 " + by(simp only: 3 4) + then have "permute_list (scp_permutation l1 l2 i) (a1@a2@b) = a1@b@a2 " + by simp + then show ?thesis + using P unfolding split_cartesian_product_def + by (metis (mono_tags, lifting) image_eqI) +qed + +lemma split_cartesian_product_memI': + assumes "a \ A" + assumes "b \ B" + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "i \ l1" + shows "(take i a)@b@(drop i a) \ split_cartesian_product l1 l2 i A B" + using assms split_cartesian_product_memI[of "take i a" "drop i a" A b B R l1 l2 i] + by (metis append_take_drop_id cartesian_power_car_memE length_take min.absorb2 subset_iff) + +lemma split_cartesian_product_memE: + assumes "a \ split_cartesian_product l1 l2 i A B" + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "i \ l1" + shows "(take i a)@(drop (i + l2) a) \ A" + "(drop i (take (i + l2) a)) \ B" +proof- + obtain b where b_def: "b \ cartesian_product A B \ a = permute_list (scp_permutation l1 l2 i) b" + using assms split_cartesian_product_def + by (metis (mono_tags, lifting) image_iff) + then have 0: "(take l1 b) \ A \ (drop l1 b) \ B" + using assms(2) cartesian_product_memE(1)[of b A B R l1] cartesian_product_memE(2)[of b A B R l1] + by metis + have 1: "a = (take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))" + using "0" append_take_drop_id assms(2) assms(3) assms(4) b_def + cartesian_power_car_memE scp_permutation_action subsetD + by smt + have 2: "(take i a) = (take i (take l1 b))" + using 0 1 + by (metis (no_types, lifting) append_eq_append_conv append_take_drop_id + assms(4) b_def length_permute_list length_take min.absorb1 take_take) + have "drop (i + l2) a = drop i (take l1 b)" + proof- + have "drop (i + l2) ( (take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))) = drop i (take l1 b)" + using assms + by (metis "0" "1" "2" add.commute append_eq_conv_conj append_take_drop_id + cartesian_power_car_memE drop_drop subsetD) + then show ?thesis + using 1 + by blast + qed + then show "take i a @ drop (i + l2) a \ A" + by (metis "0" "2" append_take_drop_id) + have 3: "length b = l1 + l2 " + by (metis "0" append_take_drop_id assms(2) assms(3) cartesian_power_car_memE length_append subsetD) + then have "(drop i (take (i + l2) ((take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))))) = (drop l1 b)" + proof- + have 0: "take (i + l2) ((take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))) = + take (i + l2) ((take i b)@(drop l1 b)@(drop i (take l1 b)))" + using assms(4) + by (metis min.absorb1 take_take) + have 1: "length ((take i b)@(drop l1 b)) = i + l2" + using 3 assms + by (metis (no_types, opaque_lifting) add_diff_cancel_left' b_def length_append length_drop + length_permute_list length_take min.absorb2 trans_le_add1) + have 2: "take (i + l2) (((take i b)@(drop l1 b))@(drop i (take l1 b))) = (take i b)@(drop l1 b)" + using 1 + by (metis append_eq_conv_conj) + have 3: "take (i + l2) ((take i b)@(drop l1 b)@(drop i (take l1 b))) = (take i b)@(drop l1 b)" + using 2 + by (metis append.assoc) + have 4: "take (i + l2) ((take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))) = (take i b)@(drop l1 b)" + using "0" "3" + by presburger + then have 5: "(drop i (take (i + l2) ((take i (take l1 b))@(drop l1 b)@(drop i (take l1 b))))) = + drop i ((take i b)@(drop l1 b))" + by presburger + have "length (take i b) = i" + by (metis "1" append_take_drop_id assms(4) le_add1 length_take min.absorb2 min.bounded_iff nat_le_linear take_all) + then show ?thesis using 5 + by (metis append_eq_conv_conj) + qed + then have "drop i (take (i + l2) a) = drop l1 b" + using 1 by blast + then show "(drop i (take (i + l2) a)) \ B" + using 0 + by presburger +qed + +lemma split_cartesian_product_mem_length: + assumes "a \ split_cartesian_product l1 l2 i A B" + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "i \ l1" + shows "length a = l1 + l2" + using assms unfolding split_cartesian_product_def + using cartesian_product_closed[of A R l1 B l2] scp_permutes[of i l1 l2] + by (smt cartesian_power_car_memE imageE in_mono length_permute_list scp_permutation_def) + +lemma split_cartesian_product_memE': + assumes "a1@b@a2 \ split_cartesian_product l1 l2 i A B" + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "i \ l1" + assumes "length a1 = i" + assumes "length b = l2" + assumes "length as = (l1 - i)" + shows "a1@a2 \ A" + "b \ B" + using assms split_cartesian_product_memE(1)[of "a1@b@a2" l1 l2 i A B R] + apply (metis append.assoc append_eq_conv_conj length_append) + using assms split_cartesian_product_memE(2)[of "a1@b@a2" l1 l2 i A B R] + by (metis add_diff_cancel_left' append_eq_conv_conj drop_take) + +lemma split_cartesian_product_closed: + assumes "A \ carrier (R\<^bsup>l1\<^esup>)" + assumes "B \ carrier (R\<^bsup>l2\<^esup>)" + assumes "i \ l1" + shows "split_cartesian_product l1 l2 i A B \ carrier (R\<^bsup>l1 + l2\<^esup>)" +proof fix x + assume A: "x \ split_cartesian_product l1 l2 i A B" + show "x \ carrier (R\<^bsup>l1 + l2\<^esup>)" + apply(rule cartesian_power_car_memI) + apply (meson \x \ split_cartesian_product l1 l2 i A B\ assms(1) + assms(2) assms(3) split_cartesian_product_mem_length) + using assms A unfolding split_cartesian_product_def + using cartesian_product_closed[of A R l1 B l2] + by (smt A cartesian_power_car_memE'' image_iff length_permute_list + scp_permutes set_permute_list split_cartesian_product_mem_length subsetD) +qed + +text\General function for permuting the elements of a simple cartesian product:\ + +definition intersperse :: "(nat \ nat) \ 'a tuple \ 'a tuple \ 'a tuple" where +"intersperse \ as bs = permute_list \ (as@bs) " + +lemma intersperseE: + assumes "\ permutes ({.. as bs) = n" + by (metis assms(2) intersperse_def length_append length_permute_list) + +lemma intersperseE': + assumes "\ permutes ({.. i < k" + shows "(intersperse \ as bs)! i = as ! \ i" +proof- + have "permute_list \ (as @ bs) ! i = (as @ bs) ! \ i" + using assms permute_list_nth[of \ "(as@bs)" i] + unfolding intersperse_def + by (metis length_append lessThan_iff permutes_not_in trans_less_add1) + then show ?thesis using assms + by (metis intersperse_def nth_append) +qed + +lemma intersperseE'': + assumes "\ permutes ({.. i \ k" + shows "(intersperse \ as bs)! i = bs ! ((\ i) - k)" +proof- + have 0: "permute_list \ (as @ bs) ! i = (as @ bs) ! \ i" + using assms permute_list_nth[of \ "(as@bs)" i] + unfolding intersperse_def + proof - + have "(as @ bs) ! \ i = (as @ bs) ! \ ([0..i < n\) + then show ?thesis + by (metis (no_types) \i < n\ \length as + length bs = n\ diff_zero length_append + length_upt nth_map permute_list_def) + qed + have 1: "\ i < n" + using assms + by (meson lessThan_iff permutes_in_image) + have 2: "(\ i) - k < length bs" + using "1" assms(2) assms(3) assms(5) by linarith + have "(as @ bs) ! (\ i) = bs ! (\ i - length as)" + using assms 1 2 nth_append[of as bs "(\ i)"] + by (meson not_le) + then have 3: "(as @ bs) ! (\ i) = bs ! (\ i - k)" + using assms + by blast + have 4: "permute_list \ (as @ bs) ! i = (as @ bs) ! (\ i)" + using "0" by blast + show ?thesis using 4 3 unfolding intersperse_def + by auto +qed + +text\Some more lemmas about the project\_at\_indices function.\ + +lemma project_at_indices_consecutive_ind_length: + assumes "(i::nat) < j" + assumes "j \ n" + assumes "length a = n" + shows "length (project_at_indices {i.. n" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "length (project_at_indices {i.. n" + assumes "length a = n" + assumes "k < j - i" + shows "(project_at_indices {i.. n" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "k < j - i" + shows "(project_at_indices {i.. carrier (R\<^bsup>n\<^esup>)" + shows "tl a = project_at_indices {1::nat..i. i < n - 1 \ (tl a) ! i = (project_at_indices {1::nat..Function Rings on Cartesian Powers\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\Complement operator\ + +definition ring_pow_comp :: "('a, 'b) ring_scheme \ arity \ 'a tuple set \ 'a tuple set" where +"ring_pow_comp R n S \ carrier (R\<^bsup>n\<^esup>) - S" + +lemma ring_pow_comp_closed: +"ring_pow_comp R n S \ carrier (R\<^bsup>n\<^esup>)" + by (simp add: ring_pow_comp_def) + +lemma ring_pow_comp_disjoint: +"ring_pow_comp R n S \ S = {}" + by (simp add: ring_pow_comp_def inf_sup_aci(1)) + +lemma ring_pow_comp_union: + assumes "S \ carrier (R\<^bsup>n\<^esup>)" + shows "(ring_pow_comp R n S) \ S = carrier (R\<^bsup>n\<^esup>)" + by (metis ring_pow_comp_def Un_Diff_cancel2 assms sup.absorb_iff1) + +lemma ring_pow_comp_carrier: +"ring_pow_comp R n (carrier (R\<^bsup>n\<^esup>)) = {}" + by (simp add: ring_pow_comp_def) + +lemma ring_pow_comp_empty: +"ring_pow_comp R n {} = (carrier (R\<^bsup>n\<^esup>)) " + by (simp add: ring_pow_comp_def) + +lemma ring_pow_comp_demorgans: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>n\<^esup>)" + shows "ring_pow_comp R n (A \ B) = (ring_pow_comp R n A) \ (ring_pow_comp R n B)" + by (simp add: ring_pow_comp_def Diff_Un ) + +lemma ring_pow_comp_demorgans': + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (R\<^bsup>n\<^esup>)" + shows "ring_pow_comp R n (A \ B) = (ring_pow_comp R n A) \ (ring_pow_comp R n B)" + by (simp add: ring_pow_comp_def Diff_Int) + +lemma ring_pow_comp_inv: + assumes "A \ carrier (R\<^bsup>n\<^esup>)" + shows "ring_pow_comp R n (ring_pow_comp R n A) = A" + by (simp add: ring_pow_comp_def assms double_diff) + +text\The function ring defined on the powers of a ring:\ +abbreviation(input) ring_pow_function_ring ("Fun\<^bsub>_\<^esub> _") where +"ring_pow_function_ring n R \ function_ring (carrier (R\<^bsup>n\<^esup>)) R" + +text \ + Partial function application. Given a function $f(x_1, \dots, x_{n+1})$, an index $i$ and a + point $a \in \text{carrier R}$ returns the function + $(x_1,..,x_n) \mapsto f(x_1, \dots, x_{i-1}, a, x_i, \dots, x_n)$ \ + +lemma ring_pow_function_ring_car_memE: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + shows "f \ extensional (carrier (R\<^bsup>n\<^esup>))" + "f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" + using ring_functions.function_ring_car_memE[of R f "carrier (R\<^bsup>n\<^esup>)"] assms + unfolding ring_functions_def + using function_ring_def partial_object.select_convs(1) apply (metis PiE_iff) + using Int_iff assms PiE_iff function_ring_def partial_object.select_convs(1) + by (simp add: PiE_iff function_ring_def) + +definition partial_eval :: "('a, 'b) ring_scheme \ arity \ nat \ ('a list \ 'a) \ 'a \ ('a list \ 'a)" where +"partial_eval R m n f c = restrict (\ as. f (insert_at_index as c n)) (carrier (R\<^bsup>m\<^esup>))" + +context ring +begin + +lemma function_ring_car_mem_closed: + assumes "f \ carrier (function_ring S R)" + assumes "s \ S" + shows "f s \ carrier R" + using assms unfolding function_ring_def ring_record_simps by blast + +lemma function_ring_car_mem_closed': + assumes "f \ carrier (Fun\<^bsub>Suc k\<^esub> R)" + assumes "s \ carrier (R\<^bsup>Suc k\<^esup>)" + shows "f s \ carrier R" + using assms unfolding function_ring_def ring_record_simps by blast + +lemma(in ring) partial_eval_domain: + assumes "f \ carrier (Fun\<^bsub>Suc k\<^esub> R)" + assumes "a \ carrier R" + assumes "n \k" + shows "(partial_eval R k n f a) \ carrier (Fun\<^bsub>k\<^esub> R)" + apply(rule ring_functions.function_ring_car_memI) +proof- + show "\x. x \ carrier (R\<^bsup>k\<^esup>) \ (partial_eval R k n f a) x \ (carrier R)" + proof- + fix x + assume A: "x \ carrier (R\<^bsup>k\<^esup>)" + show "(partial_eval R k n f a) x \ (carrier R)" + proof(cases "n = k") + case True + then have "(partial_eval R k n f a) x = f (insert_at_index x a n)" + by (metis (no_types, lifting) A restrict_apply' partial_eval_def) + then show "(partial_eval R k n f a) x \ carrier R" + using insert_at_index_closed[of x k R a n] assms ring_functions.function_ring_car_memE[of R] + unfolding ring_functions_def + by (smt A cartesian_power_car_memE funcset_carrier ring_pow_function_ring_car_memE(2)) + + next + case False + then have F0: "(partial_eval R k n f a) x = f (insert_at_index x a n)" + unfolding partial_eval_def + using assms + by (meson A restrict_apply') + have F1: "(insert_at_index x a n) \ carrier (R\<^bsup>Suc k\<^esup>)" + using A assms insert_at_index_closed[of x k R a n] cartesian_power_car_memE + by blast + show "(partial_eval R k n f a) x \ carrier R" + unfolding F0 apply(rule function_ring_car_mem_closed[of f "carrier (R\<^bsup>Suc k\<^esup>)"]) + apply (simp add: assms(1)) + by(rule F1) + qed + qed + show "\x. x \ carrier (R\<^bsup>k\<^esup>) \ (partial_eval R k n f a) x = undefined" + proof- + fix x + assume "x \ carrier (R\<^bsup>k\<^esup>)" + show "(partial_eval R k n f a) x = undefined" + unfolding partial_eval_def + by (meson \x \ carrier (R\<^bsup>k\<^esup>)\ restrict_apply) + qed + show "ring_functions R" + unfolding ring_functions_def + by (simp add: ring_axioms) +qed + +text\Pullbacks preserve ring power functions\ + +lemma fun_struct_maps: +"struct_maps (R\<^bsup>n\<^esup>) R = carrier (Fun\<^bsub>n\<^esub> R)" +proof + show "struct_maps (R\<^bsup>n\<^esup>) R \ carrier Fun\<^bsub>n\<^esub> R" + by (smt function_ring_car_memI struct_maps_memE(1) struct_maps_memE(2) subsetI) + show "carrier (Fun\<^bsub>n\<^esub> R) \ struct_maps (R\<^bsup>n\<^esup>) R" + using struct_maps_memI ring_functions.function_ring_car_memE + by (smt function_ring_car_mem_closed ring_axioms ring_functions.function_ring_not_car ring_functions.intro subsetI) +qed + +lemma pullback_fun_closed: + assumes "f \ struct_maps (R\<^bsup>n\<^esup>) (R\<^bsup>m\<^esup>)" + assumes "g \ carrier (Fun\<^bsub>m\<^esub> R)" + shows "pullback (R\<^bsup>n\<^esup>) f g \ carrier (Fun\<^bsub>n\<^esub> R)" + using assms(1) assms(2) fun_struct_maps pullback_closed by blast + +end + + +text\Includes $R^{|S|}$ into $R^n$ by pulling back along the projection $R^n \mapsto R^{|S|}$ at indices $S$ \ + +context ring +begin + +definition(in ring) ring_pow_inc :: " (nat set) \ arity \ ('a tuple \ 'a) => ('a tuple \ 'a) " where +"ring_pow_inc S n f = pullback (R\<^bsup>n\<^esup>) (\\<^bsub>n,S\<^esub>) f" + +lemma ring_pow_inc_is_fun: + assumes "S \ {.. carrier (Fun\<^bsub>card S\<^esub> R)" + shows "ring_pow_inc S n f \ carrier (Fun\<^bsub>n\<^esub> R)" + by (metis assms(1) assms(2) ring_pow_proj_is_map pullback_fun_closed ring_pow_inc_def) + +text\The "standard" inclusion of powers of function rings into one another\ + +abbreviation(input) std_proj:: "nat \ nat \ ('a list) \ ('a list)" where +"std_proj n m \ ring_pow_proj n ({.. n" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "i < m" + shows "std_proj n m as ! i = as ! i" +proof- + have "{.. indices_of as" + using assms cartesian_power_car_memE unfolding indices_of_def + by blast + thus ?thesis + unfolding ring_pow_proj_def + using assms nth_elem_upto[of i m] + project_at_indices_nth[of "{.. nat \ ('a list \ 'a) => ('a list \ 'a)" where +"std_inc n m f \ ring_pow_inc ({.. n" + shows "std_proj n m \ struct_maps (R\<^bsup>n\<^esup>) (R\<^bsup>m\<^esup>)" + by (metis assms card_lessThan lessThan_subset_iff ring_pow_proj_is_map) + +end +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Coordinate Functions\ +(**************************************************************************************************) +(**************************************************************************************************) + + +definition var :: "('a, 'b) ring_scheme \ nat \ nat \ ('a list \ 'a)" where +"var R n i = restrict (\x. x!i) (carrier (R\<^bsup>n\<^esup>))" + +context ring +begin + +lemma var_in_car: + assumes "i < n" + shows "var R n i \ carrier (Fun\<^bsub>n\<^esub> R)" + apply(rule function_ring_car_memI) + unfolding var_def + apply (metis assms cartesian_power_car_memE' restrict_apply') + by (meson restrict_apply) + + +lemma varE[simp]: + assumes "i < n" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "var R n i x = x ! i" + unfolding var_def + using assms(2) + by (meson restrict_apply') + +lemma std_inc_of_var: + assumes "i < n" + assumes "n \m" + shows "std_inc m n (var R n i) = (var R m i)" + apply(rule ext) +proof- + fix x + show "std_inc m n (var R n i) x = var R m i x" + apply(cases "x \ carrier (R\<^bsup>m\<^esup> )") + proof- + show "x \ carrier (R\<^bsup>m\<^esup>) \ std_inc m n (var R n i) x = var R m i x" + proof- + assume A: "x \ carrier (R\<^bsup>m\<^esup>)" + have "(restrict (project_at_indices ({..m\<^esup>))) x = ((project_at_indices ({.. carrier (R\<^bsup>m\<^esup>) \ std_inc m n (var R n i) x = var R m i x" + by (metis (mono_tags, lifting) assms(1) assms(2) card_lessThan lessThan_subset_iff less_SucI ring_axioms nat_induct_at_least ring.fun_struct_maps ring_pow_inc_is_fun struct_maps_memE(2) var_in_car) + qed +qed + +abbreviation variable ("\\<^bsub>_\<^esub>") where +"variable n i \ var R n i" + +end + +definition var_set :: "('a, 'b) ring_scheme \ nat \ ('a list \ 'a) set" where +"var_set R n = var R n ` {.. var_set R n" + obtains i where "f = var R n i \ i \ {.. {.. var_set R n" + using assms(1) assms(2) var_set_def + by blast + +context ring +begin + +lemma var_set_in_fun_ring_car: + shows "var_set R n \ carrier Fun\<^bsub>n\<^esub> R" +proof + fix x + assume "x \ var_set R n" + then obtain i where i_def: "i \ {.. x = var R n i" + unfolding var_set_def + by blast + have "i < n"using i_def + using atLeastLessThan_iff by blast + then show "x \ carrier Fun\<^bsub>n\<^esub> R" + using i_def var_in_car by blast +qed + + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Graphs of functions\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition fun_graph:: "('a, 'b) ring_scheme \ nat \ ('a list \ 'a) \ 'a list set" where +"fun_graph R n f = {as. (\x \ carrier (R\<^bsup>n\<^esup>). as = x @ [f x])}" + +context ring +begin + +lemma function_ring_car_memE: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "f a \ carrier R" + using ring_functions.function_ring_car_memE(1)[of R f] + unfolding ring_functions_def + by (meson assms(1) assms(2) ring_axioms function_ring_car_mem_closed ring_functions_def) + +lemma graph_range: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + shows "fun_graph R n f \ carrier (R\<^bsup>Suc n\<^esup> )" +proof + fix x + assume x_def: "x \ fun_graph R n f" + obtain a where a_def: "a \ carrier (R\<^bsup>n\<^esup>) \ x = a@[f a]" + using x_def fun_graph_def + by (smt mem_Collect_eq) + have f_closed: "f a \ carrier R" + using assms function_ring_car_memE a_def + by blast + show "x \ carrier (R\<^bsup>Suc n\<^esup> )" + proof(rule cartesian_power_car_memI) + show "length x = Suc n" + using x_def a_def cartesian_power_car_memE[of a R n] + by (metis length_append_singleton) + have "set x = insert (f a) (set a)" + using a_def + by (metis Un_insert_right append_Nil2 list.simps(15) set_append) + thus "set x \ carrier R" + using a_def + by (metis cartesian_power_car_memE'' f_closed insert_subset) + qed +qed + +lemma fun_graph_memE: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + assumes "p \ fun_graph R n f" + shows "(take n p) \ carrier (R\<^bsup>n\<^esup>)" + using assms unfolding fun_graph_def + by (metis (no_types, lifting) assms(2) graph_range le_add2 plus_1_eq_Suc subsetD take_closed) + +lemma fun_graph_memE': + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + assumes "p \ fun_graph R n f" + shows "f (take n p) = p!n" + using assms + unfolding fun_graph_def + by (smt Cons_nth_drop_Suc append_take_drop_id assms(2) butlast_snoc cartesian_power_car_memE + drop_all graph_range last_snoc le_Suc_eq lessI mem_Collect_eq subsetD) + +text\ + apply a function f to the tuple consisting of the first n indices, leaving the remaining indices + unchanged +\ + +definition partial_image :: "arity \ ('c list \ 'c) \ 'c list \ 'c list" where +"partial_image n f as = (f (take n as)) # (drop n as) " + +lemma partial_image_range: + assumes "f \ carrier (Fun\<^bsub>n\<^esub> R)" + assumes "m \ n" + assumes "as \ carrier (R\<^bsup>m\<^esup>)" + shows "partial_image n f as \ carrier (R\<^bsup>m - n + 1\<^esup>)" +proof(cases "m = n") + case True + then have "f (take n as) = f as" + by (metis assms(2) assms(3) cartesian_power_car_memE take_all) + then have 0: "f (take n as) \ carrier R" + using True assms(1) assms(3) function_ring_car_memE by presburger + have 1: "(drop n as) = []" + using True assms(3) cartesian_power_car_memE drop_all by blast + then show ?thesis + unfolding partial_image_def + using 0 1 + by (metis (no_types, lifting) One_nat_def assms(3) cartesian_power_car_memE + cartesian_power_car_memI empty_iff insert_iff length_drop list.set(1) + list.set(2) list.size(4) subsetI) +next + case False + then have 0: "(drop n as) \ carrier (R\<^bsup>m - n\<^esup>)" + using assms drop_closed[of n m as R] le_neq_implies_less + by blast + have 1: "f (take n as) \ carrier R" + using assms(1) assms(2) assms(3) function_ring_car_memE take_closed by blast + show ?thesis + apply(rule cartesian_power_car_memI) + apply (metis "0" One_nat_def cartesian_power_car_memE list.size(4) partial_image_def) + by (smt "1" assms(3) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth + partial_image_def set_ConsD set_drop_subset subsetD subsetI) +qed + +end + + +(**************************************************************************************************) +(**************************************************************************************************) +section\Coordinate Rings on Cartesian Powers\ +(**************************************************************************************************) +(**************************************************************************************************) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Basic Facts and Definitions\ +(**************************************************************************************************) +(**************************************************************************************************) +locale cring_coord_rings = UP_cring + + assumes one_neq_zero: "\ \ \" + +text\coordinate polynomial ring in n variables over a commutative ring\ + +definition coord_ring :: "('a, 'b) ring_scheme \ nat \ ('a, ('a, nat) mvar_poly) module" + ("_ [\\<^bsub>_\<^esub>]" 80) where "R[\\<^bsub>n\<^esub>] \ Pring R {..< n::nat}" + +sublocale cring_coord_rings < cring_functions R "carrier (R\<^bsup>n\<^esup>)" "Fun\<^bsub>n\<^esub> R" + unfolding cring_functions_def ring_functions_def + apply (simp add: R.ring_axioms R_cring) + by simp + +sublocale cring_coord_rings < MP?: cring "R[\\<^bsub>n\<^esub>]" + by (simp add: R.Pring_is_cring R_cring coord_ring_def) + +sublocale cring_coord_rings < F?: cring "Fun\<^bsub>n\<^esub> R" + by (simp add: function_ring_is_cring) + +context cring_coord_rings +begin + +lemma coord_cring_cring: +"cring (R[\\<^bsub>n\<^esub>])" unfolding coord_ring_def + by (simp add: R.Pring_is_cring R_cring) + +text\coordinate constant functions\ + +abbreviation(input) coord_const :: "'a \ ('a, nat) mvar_poly" where +"coord_const k \ ring.indexed_const R k" + +lemma coord_const_ring_hom: +"ring_hom_ring R (R[\\<^bsub>n\<^esub>]) coord_const" + unfolding coord_ring_def + apply(rule ring_hom_ringI) + apply (simp add: R.ring_axioms) + apply (simp add: R.Pring_is_ring) + apply (simp add: R.indexed_const_closed) + apply (simp add: R.indexed_const_mult) + apply (simp add: R.Pring_add R.indexed_padd_const) + by (simp add: R.Pring_one) + +text\coordinate functions\ + +lemma pvar_closed: + assumes "i < n" + shows "pvar R i \ carrier (R[\\<^bsub>n\<^esub>])" + unfolding var_to_IP_def +proof- + have "set_mset {#i#} \ {.. carrier (R[\\<^bsub>n\<^esub>])" + by (simp add: R.ring_axioms coord_ring_def R.Pring_car ring.mset_to_IP_closed) +qed + +text\relationship between multiplciation by a variable and index multiplcation\ + +lemma pvar_indexed_pmult: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(p \ i) = p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i" +proof- + have "p \ Pring_set R {..<(n::nat)} " + using R.Pring_car assms + by (metis coord_ring_def) + then have "p \ Pring_set R (UNIV::nat set)" + using R.Pring_set_restrict + by blast + then show ?thesis + using assms R.poly_index_mult[of p UNIV i] unfolding var_to_IP_def + by (metis R.Pring_mult UNIV_I coord_ring_def) +qed + +lemma coord_ring_cfs_closed: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "p m \ carrier R" + using assms unfolding coord_ring_def + using R.Pring_carrier_coeff' by blast + +lemma coord_ring_plus: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) m = p m \ Q m" + using assms unfolding coord_ring_def + by (metis R.Pring_add R.indexed_padd_def) + +lemma coord_ring_uminus: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> p) m = \ (p m)" + using assms unfolding coord_ring_def + using MP.add.inv_closed MP.minus_minus coord_ring_cfs_closed coord_ring_def + coord_ring_plus is_abelian_group R.is_cring + R.ring_axioms + by (metis P_ring_uminus_def R.Pring_a_inv assms) + +lemma coord_ring_minus: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) m = p m \ Q m" + using assms R.Pring_add[of _ p Q] coord_ring_cfs_closed + unfolding indexed_padd_def coord_ring_def a_minus_def + by (metis (no_types, lifting) MP.add.inv_closed coord_ring_def coord_ring_plus + cring_coord_rings.coord_ring_uminus cring_coord_rings_axioms) + +lemma coord_ring_one: +"\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m = (if m = {#} then \ else \)" + by (metis R.Pring_one coord_ring_def R.indexed_const_def) + +lemma coord_ring_zero: +"\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m = \" + by (metis MP.minus_zero MP.r_zero MP.zero_closed R_cring coord_ring_cfs_closed coord_ring_plus coord_ring_uminus cring.cring_simprules(17)) + +text\Evaluation of a polynomial at a point\ + +end + +abbreviation(input) point_to_eval_map where +"point_to_eval_map R as \ (\i. (if i< length as then as ! i else \\<^bsub>R\<^esub>))" + +definition eval_at_point :: "('a, 'b) ring_scheme \ 'a list \ ('a, nat) mvar_poly \ 'a" where +"eval_at_point R as p \ total_eval R (\i. (if i< length as then as ! i else \\<^bsub>R\<^esub>)) p" + + +lemma(in cring_coord_rings) eval_at_point_factored: +"eval_at_point R as p = total_eval R (point_to_eval_map R as) p" + using eval_at_point_def by blast + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Total Evaluation of a Polynomial\ +(**************************************************************************************************) +(**************************************************************************************************) + +abbreviation(input) eval_at_poly where +"eval_at_poly R p as \ eval_at_point R as p" + + +text\evaluation of coordinate polynomials\ + +context cring_coord_rings +begin + +lemma eval_at_point_closed: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "eval_at_point R a p \ carrier R" +proof- + have 0: "R.indexed_pset ({.. carrier (R[\\<^bsub>n\<^esub>])" + unfolding coord_ring_def + by (simp add: R.Pring_car R.Pring_carrier_subset) + have 1 : "poly_eval R UNIV (\i. if i < length a then a ! i else \) p \ R.indexed_pset ({..i. if i < length a then a ! i else \) p \carrier (R[\\<^bsub>n\<^esub>])" + using 0 by blast + show ?thesis + unfolding eval_at_point_def total_eval_def eval_in_ring_def + using 1 R.Pring_car R.Pring_cfs_closed cartesian_power_car_memE cartesian_power_car_memE' R.closed_funI coord_ring_def R.zero_closed + by blast +qed + +lemma eval_pvar: + assumes "i < (n::nat)" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "eval_at_point R a (pvar R i) = a!i" + unfolding eval_at_point_def +proof- + have "pvar R i = \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> \ i" + unfolding var_to_IP_def + by (metis R.Pring_one coord_ring_def R.monom_add_mset R.one_mset_to_IP) + then show "total_eval R (\i. if i < length a then a ! i else \) (pvar R i) = a ! i" + using assms R.total_eval_var[of "(\i. (if i< length a then a ! i else \\<^bsub>R\<^esub>))" i ] + by (smt cartesian_power_car_memE cartesian_power_car_memE' R.closed_funI var_to_IP_def R.zero_closed) +qed + +lemma eval_at_point_const: + assumes "k \ carrier R" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "eval_at_point R a (R.indexed_const k) = k" + unfolding eval_at_point_def + using assms(1) R.total_eval_const by blast + +lemma eval_at_point_add: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (coord_ring R n)" + shows "eval_at_point R a (A \\<^bsub>coord_ring R n\<^esub> B) = + eval_at_point R a A \\<^bsub>R\<^esub> eval_at_point R a B" + unfolding eval_at_point_def + using R.total_eval_add[of A "{.. carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier ((R[\\<^bsub>n\<^esub>]))" + shows "eval_at_point R a (A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B) = + eval_at_point R a A \\<^bsub>R\<^esub> eval_at_point R a B" + unfolding eval_at_point_def + using R.total_eval_mult[of A "{.. carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "i < n" + shows "eval_at_point R a (A \ i) = + eval_at_point R a A \\<^bsub>R\<^esub> (a!i)" +proof- + have "(A \ i) = A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (pvar R i)" + using assms(2) pvar_indexed_pmult by blast + then show ?thesis + using assms eval_at_point_mult eval_pvar pvar_closed + by presburger +qed + +lemma eval_at_point_ring_hom: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "ring_hom_ring (coord_ring R I) R (eval_at_point R a)" + unfolding eval_at_point_def + using R.total_eval_ring_hom + by (smt assms cartesian_power_car_memE cartesian_power_car_memE' R.closed_funI coord_ring_def R.zero_closed) + +lemma eval_at_point_scalar_mult: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "k \ carrier R" + shows "eval_at_point R a (poly_scalar_mult R k A) = k \\<^bsub>R\<^esub> (eval_at_point R a A)" + using assms unfolding eval_at_point_def total_eval_def eval_in_ring_def + using R.poly_eval_scalar_mult[of k "(\i. if i < length a then a ! i else \)" A "{.. carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "k \ carrier R" + shows "eval_at_point R a (k \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> A) = k \\<^bsub>R\<^esub> (eval_at_point R a A)" + by (metis R.Pring_smult assms(1) assms(2) assms(3) coord_ring_def eval_at_point_scalar_mult) + +lemma eval_at_point_subtract: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (coord_ring R n)" + shows "eval_at_point R a (A \\<^bsub>coord_ring R n\<^esub> B) = + eval_at_point R a A \\<^bsub>R\<^esub> eval_at_point R a B" + using assms eval_at_point_add[of a n A "\\<^bsub>coord_ring R n\<^esub> B"] + abelian_group.a_inv_closed[of "R[\\<^bsub>n\<^esub>]" B] + unfolding a_minus_def + abelian_group.a_inv_closed abelian_group.minus_minus abelian_group.r_neg1 abelian_groupE(1) abelian_groupE(4) coord_cring_cring cring_def eval_at_point_add eval_at_point_closed is_abelian_group ring_def + by (smt MP.add.inv_closed MP.l_neg MP.r_zero MP.zero_closed R.add.inv_closed R.add.m_assoc R.l_neg R.r_zero R.zero_closed eval_at_point_add eval_at_point_closed) + +lemma eval_at_point_a_inv: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "B \ carrier (coord_ring R n)" + shows "eval_at_point R a (\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B) = \\<^bsub>R\<^esub> eval_at_point R a B" + using assms eval_at_point_subtract[of a n "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>" B] + by (smt MP.add.inv_eq_1_iff MP.l_zero MP.minus_add MP.zero_closed R.is_abelian_group R.r_neg R.r_neg2 a_minus_def abelian_group.a_inv_closed abelian_groupE(4) eval_at_point_add eval_at_point_closed) + +lemma eval_at_point_nat_pow: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + shows "eval_at_point R a (A[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>(k::nat)) = (eval_at_point R a A)[^]k" + apply(induction k) + apply (metis Group.nat_pow_0 R.Pring_one assms(1) coord_ring_def eval_at_point_const R.one_closed) +proof- fix k::nat assume IH: "eval_at_poly R (A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> k) a = eval_at_poly R A a [^] k" + have "A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Suc k = A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> k \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> A" + using MP.nat_pow_Suc by blast + then have "eval_at_poly R (A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Suc k) a = + eval_at_poly R (A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> k) a \ eval_at_poly R A a" + using monoid.nat_pow_closed[of "(R[\\<^bsub>n\<^esub>])" A k] eval_at_point_mult[of a n "A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> k" A] assms + by (metis R.Pring_is_monoid coord_ring_def) + then show " eval_at_poly R (A [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Suc k) a = eval_at_poly R A a [^] Suc k" + using IH R.nat_pow_Suc + by auto +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Partial Evaluation of a Polynomial\ +(**************************************************************************************************) +(**************************************************************************************************) + + +definition coord_partial_eval :: + "('a, 'b) ring_scheme \ nat set \ 'a list \ ('a, nat) mvar_poly \ ('a, nat) mvar_poly" where +"coord_partial_eval R S as = poly_eval R S (point_to_eval_map R as)" + +context cring_coord_rings +begin + +lemma point_to_eval_map_closed: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "closed_fun R (point_to_eval_map R as)" + using assms + by (smt cartesian_power_car_memE cartesian_power_car_memE' R.closed_funI R.zero_closed) + +lemma coord_partial_eval_hom: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "coord_partial_eval R S as \ ring_hom (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n\<^esub>])" + unfolding coord_partial_eval_def + using point_to_eval_map_closed[of as n] assms + R.poly_eval_ring_hom[of "{.. carrier (R\<^bsup>n\<^esup>)" + shows "coord_partial_eval R S as \ ring_hom (R[\\<^bsub>n\<^esub>]) (Pring R ({.. {.. I" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "coord_partial_eval R S as p \ carrier (Pring R I)" + unfolding coord_partial_eval_def + using R.poly_eval_closed[of "point_to_eval_map R as" p "{.. carrier (R\<^bsup>n\<^esup>)" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "coord_partial_eval R S as (p \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> Q) = + (coord_partial_eval R S as p) \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> (coord_partial_eval R S as Q)" + using assms R.poly_eval_add[of p "{.. carrier (R\<^bsup>n\<^esup>)" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "coord_partial_eval R S as (p \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> Q) = + (coord_partial_eval R S as p) \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> (coord_partial_eval R S as Q)" + using assms R.poly_eval_mult[of p "{.. \ \" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "i \ S \ {.. S" using assms + by blast + have "i < length as" + by (metis IntD2 assms(2) assms(3) cartesian_power_car_memE lessThan_iff) + then have "(point_to_eval_map R as i) = as!i" + by presburger + then show ?thesis + unfolding coord_partial_eval_def var_to_IP_def + using 0 assms point_to_eval_map_closed[of as n] + R.poly_eval_index[of "point_to_eval_map R as" S i ] + by presburger +qed + +lemma coord_partial_eval_pvar': + assumes "\ \ \" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "i \ S" + shows "coord_partial_eval R S as (pvar R i) = (pvar R i)" + unfolding coord_partial_eval_def + using R.poly_eval_index[of "point_to_eval_map R as" S i ] + by (smt assms(1) assms(2) assms(3) cartesian_power_car_memE cartesian_power_car_memE' R.closed_funI var_to_IP_def R.zero_closed) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\An induction rule for coordinate rings\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma coord_ring_induct: + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "\a. a \ carrier R \ p (coord_const a)" + assumes "\i Q. Q \ carrier (R[\\<^bsub>n\<^esub>]) \ p Q \ i < n \ p (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>pvar R i)" + assumes "\Q0 Q1. Q0 \ carrier (R[\\<^bsub>n\<^esub>]) \ Q1 \ carrier (R[\\<^bsub>n\<^esub>]) \ p Q0 \ p Q1 \ p (Q0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q1)" + shows "p A" + apply(rule R.indexed_pset.induct[of A "{.. Pring_set R {.. carrier (R[\\<^bsub>n\<^esub>])" + using R.Pring_car + by (simp add: \\I. carrier (Pring R I) = Pring_set R I\ \a \ Pring_set R {.. coord_ring_def) + assume 1: "p a" + assume "i \ {..< n}" + then have 2: "i < n" + using assms + by blast + have "p (a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>pvar R i)" + using "0" "1" "2" assms(3) by blast + then show "p (a \ i)" + using "0" pvar_indexed_pmult + by presburger +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Algebraic Sets in Cartesian Powers\ +(**************************************************************************************************) +(**************************************************************************************************) + + (**********************************************************************) + (**********************************************************************) + subsubsection\The Zero Set of a Single Polynomial\ + (**********************************************************************) + (**********************************************************************) +definition zero_set :: "('a, 'b) ring_scheme \ nat \ ('a, nat) mvar_poly \ 'a list set" + ("V\") where +"zero_set R n p = {a \ carrier (R\<^bsup>n\<^esup>). eval_at_point R a p =\\<^bsub>R\<^esub>}" + +context cring_coord_rings +begin + +lemma zero_setI: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "eval_at_point R a p =\\<^bsub>R\<^esub>" + shows "a \ zero_set R n p" + using assms + by (metis (mono_tags, lifting) mem_Collect_eq zero_set_def) + +lemma zero_setE: + assumes "a \ zero_set R n p" + shows "a \ carrier (R\<^bsup>n\<^esup>)" + "eval_at_point R a p =\\<^bsub>R\<^esub>" + using assms zero_set_def + apply blast + by (metis (mono_tags, lifting) assms mem_Collect_eq zero_set_def) + +lemma zero_set_closed: + "zero_set R n p \ carrier (R\<^bsup>n\<^esup>)" +unfolding zero_set_def + by blast + +end + + (**********************************************************************) + (**********************************************************************) + subsubsection\The Zero Set of a Collection of Polynomials\ + (**********************************************************************) + (**********************************************************************) +definition affine_alg_set :: "('a, 'b) ring_scheme \ nat \ ('a, nat) mvar_poly set \ 'a list set" + where "affine_alg_set R n as = {a \ carrier (R\<^bsup>n\<^esup>). \ b \ as. a \ (zero_set R n b)}" + +context cring_coord_rings +begin + +lemma affine_alg_set_empty: +"affine_alg_set R n {} = carrier (R\<^bsup>n\<^esup>)" + unfolding affine_alg_set_def by blast + +lemma affine_alg_set_subset_zero_set: + assumes "b \ as" + shows " affine_alg_set R n as \ (zero_set R n b)" + using assms affine_alg_set_def + by blast + +lemma(in cring_coord_rings) affine_alg_set_memE: + assumes "b \ as" + assumes "a \ affine_alg_set R n as" + shows "eval_at_poly R b a = \" + using affine_alg_set_subset_zero_set zero_set_def assms(1) assms(2) + by blast + +lemma affine_alg_set_subset: + assumes "as \ bs" + shows " affine_alg_set R n bs \ affine_alg_set R n as " + using assms affine_alg_set_def + by blast + +lemma affine_alg_set_empty_set: + assumes "as = {}" + shows " affine_alg_set R n as = carrier (R\<^bsup>n\<^esup>)" + unfolding affine_alg_set_def + using assms by blast + +lemma affine_alg_set_closed: + shows "affine_alg_set R n as \ carrier (R\<^bsup>n\<^esup>)" + unfolding affine_alg_set_def + by blast + +lemma affine_alg_set_singleton: +"affine_alg_set R n {a} = zero_set R n a" + unfolding affine_alg_set_def using zero_set_closed + by blast + +lemma affine_alg_set_insert: +"affine_alg_set R n (insert a A) = zero_set R n a \ (affine_alg_set R n A)" +proof + show "affine_alg_set R n (insert a A) \ V\<^bsub>R\<^esub> n a \ affine_alg_set R n A" + using affine_alg_set_subset + by (metis Int_greatest affine_alg_set_subset_zero_set insertI1 subset_insertI) + show "V\<^bsub>R\<^esub> n a \ affine_alg_set R n A \ affine_alg_set R n (insert a A)" + unfolding affine_alg_set_def + by blast +qed + +lemma affine_alg_set_intersect: +"affine_alg_set R n (A \ B) = (affine_alg_set R n A) \ (affine_alg_set R n B)" + unfolding affine_alg_set_def by blast + +lemma affine_alg_set_memI: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "\p. p \ B \ eval_at_point R a p = \" + shows "a \ (affine_alg_set R n B)" + unfolding affine_alg_set_def zero_set_def + using assms + by blast + +lemma affine_alg_set_not_memE: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "a \ (affine_alg_set R n B)" + shows "\b \ B. eval_at_poly R b a \ \" + using assms affine_alg_set_memI by blast + + + (**********************************************************************) + (**********************************************************************) + subsubsection\Finite Unions and Intersections of Algebraic Sets are Algebraic\ + (**********************************************************************) + (**********************************************************************) +text\The product set of two sets in an arbitrary ring. That is, the set $\{ xy \mid x \in A \land y \in B \}$ for two sets $A$, $B$.\ +definition(in ring) prod_set :: "'a set \ 'a set \ 'a set" where +"prod_set as bs = (\x. fst x \ snd x) ` (as \ bs)" + +lemma(in ring) prod_setI: + assumes "c \ prod_set as bs" + shows "\a \ as. \b \ bs. c = a \ b" +proof- + obtain p where p_def: "p \ (as \ bs) \ c = fst p \ snd p" + using assms prod_set_def[of as bs] + by (metis (no_types, lifting) image_iff) + then show ?thesis + using mem_Times_iff by blast +qed + +lemma(in ring) prod_set_closed: + assumes "as \ carrier R" + assumes "bs \ carrier R" + shows "prod_set as bs \ carrier R" +proof + fix x + assume " x \ prod_set as bs" + then obtain a b where "a \ as \ b \ bs \ x = a \ b" + by (meson ring_axioms ring.prod_setI) + then have "a \ carrier R \ b \ carrier R \ x = a \ b" + using assms + by blast + then show "x \ carrier R" + by blast +qed + +text\The set of products of elements from two finite sets is again finite.\ +lemma(in ring) prod_set_finite: + assumes "finite as" + assumes "finite bs" + shows "finite (prod_set as bs)" "card (prod_set as bs) \ card as * card bs" +proof- + have "finite (as \ bs)" + using assms + by blast + then show "finite (prod_set as bs)" + using prod_set_def + by (metis (no_types, lifting) finite_imageI) + have "card (prod_set as bs) \ card (as \ bs)" + using assms + unfolding prod_set_def + using \finite (as \ bs)\ card_image_le by blast + then show "card (prod_set as bs) \ card as * card bs" + by (simp add: card_cartesian_product) +qed + +definition poly_prod_set where +"poly_prod_set n as bs = ring.prod_set (R[\\<^bsub>n\<^esub>]) as bs" + +lemma poly_prod_setE: + assumes "c \ poly_prod_set n as bs" + shows "\a \ as. \b \ bs. c = a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> b" + using ring.prod_setI[of "R[\\<^bsub>n\<^esub>]"] R.Pring_is_ring assms poly_prod_set_def coord_cring_cring cring.axioms(1) + by blast + +lemma poly_prod_setI: + assumes "a \ as" + assumes "b \ bs" + shows "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> b \ poly_prod_set n as bs" +proof- + have 0: "(a,b) \ (as \ bs)" + using assms by blast + have 1: "(\x. fst x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> snd x) (a, b) = a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> b" + by (metis fst_conv snd_conv) + have 2: "(\x. fst x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> snd x) (a, b) \ ((\x. fst x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> snd x) ` (as \ bs))" + using 0 by blast + have 3: "ring (R[\\<^bsub>n\<^esub>])" + by (simp add: R.Pring_is_ring coord_ring_def) + then show ?thesis + unfolding poly_prod_set_def using 0 1 2 3 ring.prod_set_def[of "R[\\<^bsub>n\<^esub>]" as bs] + by presburger +qed + +lemma poly_prod_set_closed: + assumes "as \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "bs \ carrier (R[\\<^bsub>n\<^esub>])" + shows "poly_prod_set n as bs \ carrier (R[\\<^bsub>n\<^esub>])" + using ring.prod_set_closed[of "R[\\<^bsub>n\<^esub>]"] R.Pring_is_ring assms(1) assms(2) poly_prod_set_def + by (simp add: coord_cring_cring cring.axioms(1)) + +lemma poly_prod_set_finite: + assumes "finite as" + assumes "finite bs" + shows "finite (poly_prod_set n as bs)" "card (poly_prod_set n as bs) \ card as * card bs" + + using ring.prod_set_finite[of "R[\\<^bsub>n\<^esub>]"] + apply (simp add: R.Pring_is_ring assms(1) assms(2) poly_prod_set_def) + using ring.prod_set_finite[of "R[\\<^bsub>n\<^esub>]"] + apply (simp add: assms(1) assms(2) coord_cring_cring cring.axioms(1)) + by (simp add: assms(1) assms(2) coord_cring_cring cring.axioms(1) poly_prod_set_def ring.prod_set_finite(2)) + +end + +locale domain_coord_rings = cring_coord_rings + domain + +lemma(in domain_coord_rings) poly_prod_set_algebraic_set: + assumes "as \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "bs \ carrier (R[\\<^bsub>n\<^esub>])" + shows "affine_alg_set R n as \ affine_alg_set R n bs = affine_alg_set R n (poly_prod_set n as bs)" +proof + show "affine_alg_set R n as \ affine_alg_set R n bs \ affine_alg_set R n (poly_prod_set n as bs)" + proof fix x + assume A: "x \ affine_alg_set R n as \ affine_alg_set R n bs" + show "x \ affine_alg_set R n (poly_prod_set n as bs)" + proof(rule affine_alg_set_memI) + show "x \ carrier (R\<^bsup>n\<^esup>)" + using A affine_alg_set_closed + by blast + show "\p. p \ poly_prod_set n as bs \ eval_at_poly R p x = \" + proof- fix p + assume B: "p \ poly_prod_set n as bs" + show "eval_at_poly R p x = \" + proof- + obtain p0 p1 where C: "p0 \ as \ p1 \ bs \ p = p0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> p1" + using B poly_prod_setE by blast + then have D: "eval_at_poly R p x = (eval_at_poly R p0 x) \ (eval_at_poly R p1 x)" + using \x \ carrier (R\<^bsup>n\<^esup>)\ assms(1) assms(2) eval_at_point_mult + by blast + show ?thesis proof(cases "x \ affine_alg_set R n as") + case True + then have "(eval_at_poly R p0 x) = \" + using C affine_alg_set_memE by blast + then show ?thesis + by (smt C D \x \ carrier (R\<^bsup>n\<^esup>)\ assms(2) eval_at_point_closed R.semiring_axioms semiring.l_null subsetD) + next + case False + then have "x \ affine_alg_set R n bs" + using A + by blast + then have "(eval_at_poly R p1 x) = \" + using C affine_alg_set_memE by blast + then show ?thesis + using C A False + by (smt D \x \ carrier (R\<^bsup>n\<^esup>)\ assms(1) eval_at_point_closed R.r_null subsetD) + qed + qed + qed + qed + qed + show "affine_alg_set R n (poly_prod_set n as bs) \ affine_alg_set R n as \ affine_alg_set R n bs" + proof fix x + assume A: "x \ affine_alg_set R n (poly_prod_set n as bs)" + have x_car: "x \ carrier (R\<^bsup>n\<^esup>)" + using A affine_alg_set_closed + by blast + show "x \ affine_alg_set R n as \ affine_alg_set R n bs" + proof(cases "x \ affine_alg_set R n as") + case True + then show ?thesis by blast + next + case False + have "x \ affine_alg_set R n bs" + proof(rule affine_alg_set_memI) + show "x \ carrier (R\<^bsup>n\<^esup>)" + using A affine_alg_set_closed by blast + show "\p. p \ bs \ eval_at_poly R p x = \" + proof- + fix p assume p_def: "p \ bs" + obtain a where a_def: "a \ as \ eval_at_poly R a x \ \" + using False affine_alg_set_not_memE \x \ carrier (R\<^bsup>n\<^esup>)\ + by blast + then have "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> p \ (poly_prod_set n as bs)" + using poly_prod_setI[of a as p bs] p_def + by blast + then have "eval_at_poly R (a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> p) x = \" + using A affine_alg_set_memE + by blast + + then have "eval_at_poly R a x \ eval_at_poly R p x = \" + using eval_at_point_mult[of x n a p] + by (metis (no_types, lifting) \x \ carrier (R\<^bsup>n\<^esup>)\ a_def assms(1) assms(2) p_def subsetD) + then show "eval_at_poly R p x = \" + using a_def p_def + by (meson assms(1) assms(2) eval_at_point_closed integral_iff subsetD x_car) + qed + qed + then show ?thesis + by blast + qed + qed +qed + +definition is_algebraic :: "('a, 'b) ring_scheme \ nat \ 'a list set \ bool" where +"is_algebraic R n S = (\ps. finite ps \ ps \ carrier (R[\\<^bsub>n\<^esub>]) \ S = affine_alg_set R n ps)" + +context cring_coord_rings +begin + +lemma is_algebraicE: + assumes "is_algebraic R n S" + obtains ps where "finite ps" "ps \ carrier (R[\\<^bsub>n\<^esub>])" "S = affine_alg_set R n ps" + using assms + by (meson is_algebraic_def) + +lemma is_algebraicI: + assumes "finite ps" + assumes "ps \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "S = affine_alg_set R n ps" + shows "is_algebraic R n S" + using is_algebraic_def assms + by blast + +lemma is_algebraicI': + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "S = zero_set R n p" + shows "is_algebraic R n S" + by (metis affine_alg_set_singleton assms(1) assms(2) empty_subsetI finite.emptyI finite.intros(2) insert_subset is_algebraic_def) + +end + +definition alg_sets :: "arity \ ('a, 'b) ring_scheme \ ('a list set) set" where +"alg_sets n R = {S. is_algebraic R n S}" + +context cring_coord_rings +begin + +lemma intersection_is_alg: + assumes "is_algebraic R n A" + assumes "is_algebraic R n B" + shows "is_algebraic R n (A \ B)" +proof- + obtain as where as_def: "finite as \ as \ carrier (R[\\<^bsub>n\<^esub>]) \ A = affine_alg_set R n as" + by (meson assms(1) is_algebraicE) + obtain bs where bs_def: "finite bs \ bs \ carrier (R[\\<^bsub>n\<^esub>]) \ B = affine_alg_set R n bs" + by (meson assms(2) is_algebraicE) + show ?thesis apply(rule is_algebraicI[of "as \ bs"]) + using as_def bs_def apply blast + using as_def bs_def apply blast + by (simp add: affine_alg_set_intersect as_def bs_def) +qed + +lemma(in domain_coord_rings) union_is_alg: + assumes "is_algebraic R n A" + assumes "is_algebraic R n B" + shows "is_algebraic R n (A \ B)" +proof- + obtain as where as_def: "finite as \ as \ carrier (R[\\<^bsub>n\<^esub>]) \ A = affine_alg_set R n as" + by (meson assms(1) is_algebraicE) + obtain bs where bs_def: "finite bs \ bs \ carrier (R[\\<^bsub>n\<^esub>]) \ B = affine_alg_set R n bs" + by (meson assms(2) is_algebraicE) + show ?thesis apply(rule is_algebraicI[of "poly_prod_set n as bs"]) + using as_def bs_def + apply (simp add: poly_prod_set_finite(1)) + using as_def bs_def poly_prod_set_closed apply blast + using as_def bs_def poly_prod_set_algebraic_set + by simp +qed + +lemma zero_set_zero: +"zero_set R n \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> = carrier (R\<^bsup>n\<^esup>)" + by (metis R.add.r_cancel_one cring.cring_simprules(2) cring.cring_simprules(8) + coord_cring_cring cring_coord_rings.eval_at_point_add cring_coord_rings.eval_at_point_closed + cring_coord_rings_axioms subsetI subset_antisym zero_setI zero_set_closed) + +lemma affine_alg_set_set: +"affine_alg_set R n {\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>} = carrier (R\<^bsup>n\<^esup>)" +using affine_alg_set_singleton zero_set_zero +by blast + +lemma car_is_alg: +"is_algebraic R n (carrier (R\<^bsup>n\<^esup>))" + apply(rule is_algebraicI[of "{\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>}"]) + apply blast + using R.Pring_zero_closed + apply blast + using affine_alg_set_set by blast + +lemma zero_set_nonzero_constant: + assumes "a \ \" + assumes "a \ carrier R" + shows "zero_set R n (coord_const a) = {}" +proof(rule ccontr) + assume "V n (coord_const a) \ {}" + then obtain x where "x \ V n (coord_const a)" + by blast + then show False + by (metis assms(1) assms(2) cring_coord_rings.eval_at_point_const cring_coord_rings.zero_setE(1) cring_coord_rings.zero_setE(2) cring_coord_rings_axioms) +qed + +lemma zero_set_one: + assumes "a \ \" + assumes "a \ carrier R" + shows "zero_set R n \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> = {}" + using zero_set_nonzero_constant + by (metis R.Pring_one coord_ring_def one_neq_zero R.one_closed) + +lemma empty_set_as_affine_alg_set: +"affine_alg_set R n {\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>} = {}" + using affine_alg_set_singleton local.one_neq_zero zero_set_one by blast + +lemma empty_is_alg: +"is_algebraic R n {}" + apply(rule is_algebraicI'[of "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>"]) + apply blast + using local.one_neq_zero zero_set_one by blast + + + (**********************************************************************) + (**********************************************************************) + subsubsection\Finite Sets Are Algebraic\ + (**********************************************************************) + (**********************************************************************) +text\the function mapping a point in $R^n$ to the unique linear polynomial vanishing exclusively at that point\ + +definition pvar_trans :: "nat \ nat \ 'a \ ('a, nat) mvar_poly" where +"pvar_trans n i a = (pvar R i) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const a" + +lemma pvar_trans_closed: + assumes "a \ carrier R" + assumes "i < n" + shows "pvar_trans n i a \ carrier (R[\\<^bsub>n\<^esub>])" + unfolding pvar_trans_def using assms + by (metis MP.minus_closed coord_ring_def R.indexed_const_closed local.pvar_closed) + +lemma pvar_trans_eval: + assumes "a \ carrier R" + assumes "b \ carrier (R\<^bsup>n\<^esup>)" + assumes "i < n" + shows "eval_at_point R b (pvar_trans n i a) = (b!i) \ a" +proof- + have "eval_at_point R b (pvar_trans n i a) = + (eval_at_point R b (pvar R i)) \ (eval_at_point R b (\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (coord_const a)))" + unfolding pvar_trans_def a_minus_def using assms + by (metis MP.add.inv_closed coord_ring_def eval_at_point_add R.indexed_const_closed local.pvar_closed) + + then show ?thesis + by (metis a_minus_def assms(1) assms(2) assms(3) coord_ring_def eval_at_point_a_inv eval_at_point_const eval_pvar R.indexed_const_closed) +qed + +definition point_to_polys :: "'a list \ ('a, nat) mvar_poly list" where +"point_to_polys as = map (\ x. pvar_trans (length as) (snd x) (fst x)) (zip as (index_list (length as)))" + +lemma point_to_polys_length: +"length (point_to_polys as) = length as" + unfolding point_to_polys_def + by (smt index_list_length length_map list.map_ident map_eq_imp_length_eq zip_eq_conv) + +lemma point_to_polysE: + assumes "i < length as" + shows "(point_to_polys as) ! i = (pvar_trans (length as) i (as ! i))" +proof- + have " (zip as (index_list (length as)))!i = ((as!i), i)" + by (metis assms index_list_indices index_list_length nth_zip) + then have 0: "(point_to_polys as) ! i = (\ x. pvar_trans (length as) (snd x) (fst x)) ((as!i), i)" + unfolding point_to_polys_def + using assms nth_map[of i "(zip as (index_list (length as)))" "(\x. pvar_trans (length as) (snd x) (fst x))" ] + by (metis index_list_length length_map map_fst_zip) + then show ?thesis + by (metis fst_conv snd_conv) +qed + +lemma point_to_polysE': + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "i < n" + shows "eval_at_point R as ((point_to_polys as) ! i) = \" + by (metis assms(1) assms(2) cartesian_power_car_memE cartesian_power_car_memE' point_to_polysE pvar_trans_eval R.r_right_minus_eq) + +lemma point_to_polysE'': + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "b \ set (point_to_polys as)" + shows "eval_at_point R as b = \" + using point_to_polysE' + by (metis assms(1) assms(2) cartesian_power_car_memE in_set_conv_nth point_to_polys_length) + +lemma point_to_polys_zero_set: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "b \ set (point_to_polys as)" + shows "as \ zero_set R n b" + using assms(1) assms(2) point_to_polysE'' zero_setI by blast + +lemma point_to_polys_closed: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "set (point_to_polys as) \ carrier (R[\\<^bsub>n\<^esub>])" + using assms point_to_polysE pvar_trans_closed + by (smt cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth point_to_polys_length subsetI) + +lemma point_to_polys_affine_alg_set: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "affine_alg_set R n (set (point_to_polys as)) = {as}" +proof + show "affine_alg_set R n (set (point_to_polys as)) \ {as}" + proof + fix x + assume A0: "x \ affine_alg_set R n (set (point_to_polys as))" + then have 0: "length x = n" using affine_alg_set_closed[of n " (set (point_to_polys as))"] + using cartesian_power_car_memE by blast + have "\i. i < n \ x!i = as!i" + proof- + fix i + assume A1: "i < n" + show " x!i = as!i" + using A0 + by (smt A1 affine_alg_set_closed affine_alg_set_memE assms cartesian_power_car_memE + cartesian_power_car_memE' nth_mem point_to_polysE point_to_polys_length + pvar_trans_eval R.r_right_minus_eq subsetD) + qed + then have "x = as" + by (metis "0" assms cartesian_power_car_memE nth_equalityI) + then show "x \ {as}" + by blast + qed + show "{as} \ affine_alg_set R n (set (point_to_polys as))" + proof- + have "as \ affine_alg_set R n (set (point_to_polys as))" + using affine_alg_set_not_memE assms point_to_polysE'' + by blast + then show ?thesis + by blast + qed +qed + +lemma singleton_is_algebraic: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "is_algebraic R n {as}" + apply(rule is_algebraicI[of "(set (point_to_polys as))"]) + apply blast + using point_to_polys_affine_alg_set + using assms point_to_polys_closed apply blast + by (simp add: assms point_to_polys_affine_alg_set) + +lemma(in domain_coord_rings) finite_sets_are_algebraic: + assumes "finite F" + shows "F \ carrier (R\<^bsup>n\<^esup>) \ is_algebraic R n F" + apply(rule finite.induct) + apply (simp add: assms) + using empty_is_alg apply blast + using singleton_is_algebraic + by (metis union_is_alg insert_is_Un insert_subset) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Polynomial Maps\ +(**************************************************************************************************) +(**************************************************************************************************) + + (**********************************************************************) + (**********************************************************************) + subsection\The Action of Index Permutations on Polynomials\ + (**********************************************************************) + (**********************************************************************) + +definition permute_poly_args :: + "nat \ (nat \ nat) \ ('a, nat) mvar_poly \ ('a, nat) mvar_poly" where +"permute_poly_args (n::nat) \ p = indexed_poly_induced_morphism {..\<^bsub>n\<^esub>]) coord_const (\i. pvar R (\ i)) p" + +lemma permute_poly_args_characterization: + assumes "\ permutes {..< n}" + shows "(ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n\<^esub>]) (permute_poly_args (n::nat) \))" + "(\i \ {..) (pvar R i) = pvar R (\ i))" + "(\a \ carrier R. permute_poly_args (n::nat) \ (coord_const a) = (coord_const a))" +proof- + have 0: "cring (R[\\<^bsub>n\<^esub>])" + by (simp add: MP.is_cring) + have 1: "(\i. pvar R (\ i)) \ {.. carrier (R[\\<^bsub>n\<^esub>])" + proof + fix x + assume A: "x \ {.. x \ {.. x < n" + using assms + by auto + then show "pvar R (\ x) \ carrier (R[\\<^bsub>n\<^esub>]) " + using pvar_closed[of "\ x" n] + by blast + qed + have 2: " ring_hom_ring R (R[\\<^bsub>n\<^esub>]) coord_const" + using R.indexed_const_ring_hom unfolding coord_ring_def + by blast + have 3: " indexed_poly_induced_morphism {..\<^bsub>n\<^esub>]) coord_const (\i. pvar R (\ i)) = + indexed_poly_induced_morphism {..\<^bsub>n\<^esub>]) coord_const (\i. pvar R (\ i))" + by blast + show "ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n\<^esub>]) (permute_poly_args n \)" + using 0 1 2 3 + R.Pring_universal_prop[of "(R[\\<^bsub>n\<^esub>])" " (\i. pvar R (\ i))" "{.." ] + unfolding permute_poly_args_def + by (metis coord_ring_def) + show "\i\{.. (pvar R i) = pvar R (\ i)" + using 0 1 2 3 + R.Pring_universal_prop(2)[of "(R[\\<^bsub>n\<^esub>])" " (\i. pvar R (\ i))" "{.." ] + unfolding permute_poly_args_def var_to_IP_def + by blast + show "\a\carrier R. permute_poly_args n \ (coord_const a) = coord_const a" + using 0 1 2 3 + R.Pring_universal_prop[of "(R[\\<^bsub>n\<^esub>])" " (\i. pvar R (\ i))" "{.." ] + unfolding permute_poly_args_def var_to_IP_def + by blast +qed + +lemma permute_poly_args_closed: + assumes "\ permutes {.. carrier (R[\\<^bsub>n\<^esub>])" + shows "permute_poly_args n \ p \ carrier (R[\\<^bsub>n\<^esub>])" +proof- + have "(ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n\<^esub>]) (permute_poly_args (n::nat) \))" + using assms permute_poly_args_characterization(1) + by blast + then have "(permute_poly_args (n::nat) \) \ carrier (R[\\<^bsub>n\<^esub>]) \ carrier (R[\\<^bsub>n\<^esub>])" + unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def + by blast + then show ?thesis + using assms + by blast +qed + + +lemma permute_poly_args_constant: + assumes "a \ carrier R" + assumes "\ permutes {.. (coord_const a) = (coord_const a)" + using assms permute_poly_args_characterization(3) + by blast + +lemma permute_poly_args_add: + assumes "\ permutes {.. carrier (R[\\<^bsub>n\<^esub>])" + assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "permute_poly_args n \ (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q) = (permute_poly_args n \ p) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (permute_poly_args n \ q)" + using permute_poly_args_characterization(1)[of \] assms + unfolding ring_hom_ring_def ring_hom_ring_axioms_def + by (metis (no_types, lifting) ring_hom_add) + +lemma permute_poly_args_mult: + assumes "\ permutes {.. carrier (R[\\<^bsub>n\<^esub>])" + assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "permute_poly_args n \ (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q) = (permute_poly_args n \ p) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (permute_poly_args n \ q)" + using permute_poly_args_characterization(1)[of \] assms + unfolding ring_hom_ring_def ring_hom_ring_axioms_def + using ring_hom_mult + by (metis (mono_tags, lifting)) + +lemma permute_poly_args_indexed_pmult: + assumes "\ permutes {.. carrier (R[\\<^bsub>n\<^esub>])" + assumes "i \ {.. (p \ i)) = (permute_poly_args n \ p) \ (\ i)" +proof + fix x + show "permute_poly_args n \ (p \ i) x = (permute_poly_args n \ p \ \ i) x" + proof- + have 0: "(p \ i) = (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i)" + using assms pvar_indexed_pmult + by blast + have 1: "(permute_poly_args n \ p) \ (\ i) = (permute_poly_args n \ p) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R (\ i)" + using assms permute_poly_args_closed pvar_indexed_pmult by blast + have 2: "permute_poly_args n \ (p \ i) x = permute_poly_args n \ (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) x" + using \p \ i = p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i\ by presburger + then show ?thesis using 1 R.Pring_var_closed assms(1) assms(2) assms(3) assms + permute_poly_args_mult R.is_cring permute_poly_args_characterization(2) R.zero_closed + by (metis coord_ring_def) + qed +qed + +lemma permute_list_closed: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "\ permutes {.. a) \ carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms cartesian_power_car_memE length_permute_list apply blast +proof- + have 0: "set a \ carrier R" + using assms(1) cartesian_power_car_memE'' by blast + have "\ permutes {.. a) = set a" + using assms set_permute_list[of \ a] \\ permutes {.. + by blast + then show "set (permute_list \ a) \ carrier R" + by (simp add: "1" "0") +qed + +lemma permute_list_set: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "\ permutes {.. a) = set a" +proof- + have "\ permutes {.. a) = set a" + using assms set_permute_list[of \ a] + by blast +qed + +end + +definition perm_map :: "('a, 'b) ring_scheme \ nat \ (nat \ nat) \ 'a list \ 'a list" where +"perm_map R n \ = restrict (permute_list \) (carrier (R\<^bsup>n\<^esup>))" + +context cring_coord_rings +begin + +lemma perm_map_is_struct_map: + assumes "\ permutes {.. \ struct_maps (R\<^bsup>n\<^esup>) (R\<^bsup>n\<^esup>)" + apply(rule struct_maps_memI) + unfolding perm_map_def restrict_def using assms permute_list_closed[of _ n \] + apply metis + by metis + +lemma permute_poly_args_eval: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "\ permutes {.. carrier (R[\\<^bsub>n\<^esub>])" + shows "eval_at_point R a (permute_poly_args n \ p) = eval_at_point R (permute_list \ a) p" + apply(rule R.indexed_pset.induct[of p "{..p Q. p \ Pring_set R {.. + eval_at_point R a (permute_poly_args n \ p) = eval_at_point R (permute_list \ a) p \ + Q \ Pring_set R {.. + eval_at_point R a (permute_poly_args n \ Q) = eval_at_point R (permute_list \ a) Q \ + eval_at_point R a (permute_poly_args n \ (p \ Q)) = eval_at_point R (permute_list \ a) (p \ Q)" + proof- + fix p Q assume A0: "p \ Pring_set R {.. p) = eval_at_point R (permute_list \ a) p " + assume A2: "Q \ Pring_set R {.. Q) = eval_at_point R (permute_list \ a) Q" + have A0': "p \ carrier (R[\\<^bsub>n\<^esub>]) " + using A0 R.Pring_car unfolding coord_ring_def by blast + have A2': "Q \ carrier (R[\\<^bsub>n\<^esub>]) " + using A2 R.Pring_car unfolding coord_ring_def by blast + have A4: "(permute_poly_args n \ (p \ Q)) = (permute_poly_args n \ p) \ (permute_poly_args n \ Q)" + proof- + have "(permute_poly_args n \ (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q)) = (permute_poly_args n \ p) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (permute_poly_args n \ Q)" + using A0' A2' assms permute_poly_args_add by blast + then show ?thesis + unfolding coord_ring_def + by (metis R.Pring_add) + qed + show A5: "eval_at_point R a (permute_poly_args n \ (p \ Q)) = eval_at_point R (permute_list \ a) (p \ Q)" + using eval_at_point_add[of a n "permute_poly_args n \ p" "permute_poly_args n \ Q" ] + permute_poly_args_add[of \ n p Q] A0' A1 A2' A3 A4 permute_poly_args_closed assms + by (metis R.Pring_add cartesian_power_car_memE cartesian_power_car_memE'' + cartesian_power_car_memI coord_ring_def eval_at_point_add length_permute_list permute_list_set) + qed + show "\p i. p \ Pring_set R {.. + eval_at_point R a (permute_poly_args n \ p) = eval_at_point R (permute_list \ a) p \ + i \ {.. eval_at_point R a (permute_poly_args n \ (p \ i)) = eval_at_point R (permute_list \ a) (p \ i)" + proof- + fix p i + assume A0: "p \ Pring_set R {.. p) = eval_at_point R (permute_list \ a) p " + assume A2: "i \ {.. (p \ i)) = eval_at_point R a (permute_poly_args n \ p \ \ i)" + using permute_poly_args_indexed_pmult[of \ n p i ] A0 A1 A2 assms + by (metis R.Pring_car coord_ring_def) + then have LHS' : "eval_at_point R a (permute_poly_args n \ (p \ i)) = + eval_at_point R a (permute_poly_args n \ p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R (\ i))" + using A0 R.Pring_car assms(1) assms permute_poly_args_closed pvar_indexed_pmult + by (metis coord_ring_def) + have "eval_at_point R a (permute_poly_args n \ p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R (\ i)) = + eval_at_point R a (permute_poly_args n \ p) \ eval_at_point R a (pvar R (\ i))" + proof- + have 1: "permute_poly_args n \ p \ carrier (R[\\<^bsub>n\<^esub>])" + using A0 R.Pring_car assms(1) assms permute_poly_args_closed + by (metis coord_ring_def) + have "pvar R (\ i) \ carrier (R[\\<^bsub>n\<^esub>])" + proof- + have "\ i \ {.. i) < n" + by blast + then show ?thesis + using pvar_closed[of "\ i" n] + by blast + qed + then have LHS'' : "eval_at_point R a (permute_poly_args n \ (p \ i)) = + (eval_at_point R a (permute_poly_args n \ p)) \\<^bsub>R\<^esub> eval_at_point R a (pvar R (\ i))" + using LHS' "1" eval_at_point_mult assms + by presburger + then show ?thesis + using LHS' by presburger + qed + then have LHS'': "eval_at_point R a (permute_poly_args n \ (p \ i)) = + eval_at_point R a (permute_poly_args n \ p) \ eval_at_point R a (pvar R (\ i))" + using LHS' by presburger + have 0: "eval_at_point R a (pvar R (\ i)) = a! (\ i)" + proof- + have "\ i \ {.. i < n" + by blast + have 1: "permute_list \ a \ carrier (R\<^bsup>n\<^esup>)" + using assms(1) assms(2) assms(3) permute_list_closed by blast + show ?thesis + using 0 1 eval_pvar[of "\ i" n a] assms + by blast + qed + have 1: "(permute_list \ a)! i = a! \ i" + proof- + have "length a = n" + using assms cartesian_power_car_memE + by blast + then have "{.. permutes {..{.. + by blast + show ?thesis using 0 1 permute_list_nth[of \ a i] + by blast + qed + have LHS''': "eval_at_point R a (permute_poly_args n \ (p \ i)) = + eval_at_point R (permute_list \ a) p \ a! (\ i)" + using 0 LHS'' A1 + by presburger + have RHS: "eval_at_point R (permute_list \ a) (p \ i) = + (eval_at_point R (permute_list \ a) p) \\<^bsub>R\<^esub> (eval_at_point R (permute_list \ a) (pvar R i))" + proof- + have "(p \ i) = p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (pvar R i)" + using A0 R.Pring_car pvar_indexed_pmult unfolding coord_ring_def + by blast + then show ?thesis + using eval_at_point_mult[of "(permute_list \ a)" n p "(pvar R i)" ] + A0 A2 R.Pring_car R.Pring_var_closed assms(1) assms(2) assms(3) permute_list_closed + by (metis coord_ring_def) + qed + then have RHS': "eval_at_point R (permute_list \ a) (p \ i) = + (eval_at_point R (permute_list \ a) p) \\<^bsub>R\<^esub> (permute_list \ a)! i" + proof- + have 0: "i < n" + using A2 assms + by blast + have 1: "permute_list \ a \ carrier (R\<^bsup>n\<^esup>)" + using assms permute_list_closed + by blast + show ?thesis + using 0 1 eval_pvar[of i n "(permute_list \ a)" ] RHS + by presburger + qed + then show "eval_at_point R a (permute_poly_args n \ (p \ i)) = eval_at_point R (permute_list \ a) (p \ i)" + using LHS''' A1 1 + by presburger + qed +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Inverse Images of Sets by Tuples of Polynomials\ + (**********************************************************************) + (**********************************************************************) + +definition is_poly_tuple :: "nat \ ('a, nat) mvar_poly list \ bool" where +"is_poly_tuple (n::nat) fs = (set (fs) \ carrier (R[\\<^bsub>n\<^esub>]))" + +lemma is_poly_tupleE: + assumes "is_poly_tuple n fs" + assumes "j < length fs" + shows "fs ! j \ carrier (R[\\<^bsub>n\<^esub>])" + using assms is_poly_tuple_def nth_mem + by blast + +lemma is_poly_tuple_Cons: + assumes "is_poly_tuple n fs" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + shows "is_poly_tuple n (f#fs)" + using assms unfolding is_poly_tuple_def + by (metis (no_types, lifting) set_ConsD subset_iff) + +lemma is_poly_tuple_append: + assumes "is_poly_tuple n fs" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + shows "is_poly_tuple n (fs@[f])" + using assms set_append unfolding is_poly_tuple_def + by (metis (no_types, lifting) Un_subset_iff append_Nil2 set_ConsD subset_code(1)) + +definition poly_tuple_eval :: "('a, nat) mvar_poly list \ 'a list \ 'a list" where +"poly_tuple_eval fs as = map (\ f. eval_at_poly R f as) fs " + +lemma poly_tuple_evalE: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "j < m" + shows "(poly_tuple_eval fs as)!j \ carrier R" +proof- + have 0: "(poly_tuple_eval fs as)!j = (eval_at_poly R (fs!j) as)" + using poly_tuple_eval_def + by (metis assms(2) assms(4) nth_map) + have 1: "(fs!j) \ carrier (R[\\<^bsub>n\<^esub>])" + using assms is_poly_tupleE + by blast + show ?thesis + using assms 0 1 eval_at_point_closed + by presburger +qed + +lemma poly_tuple_evalE': + shows "length (poly_tuple_eval fs as) = length fs" + unfolding poly_tuple_eval_def + using length_map by blast + +lemma poly_tuple_evalE'': + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "j < m" + shows "(poly_tuple_eval fs as)!j = (eval_at_poly R (fs!j) as)" + using assms + unfolding poly_tuple_eval_def + using nth_map + by blast + +lemma poly_tuple_eval_closed: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "(poly_tuple_eval fs as) \ carrier (R\<^bsup>m\<^esup>)" +proof(rule cartesian_power_car_memI) + show "length (poly_tuple_eval fs as) = m" + using assms + by (simp add: assms poly_tuple_evalE') + show "set (poly_tuple_eval fs as) \ carrier R" + proof fix x + assume "x \ set (poly_tuple_eval fs as)" + then obtain j where j_def: "j< m \ x = (poly_tuple_eval fs as)!j" + using assms + by (metis \length (poly_tuple_eval fs as) = m\ in_set_conv_nth) + then show "x \ carrier R" + using assms(1) assms(2) assms(3) poly_tuple_evalE assms by blast + qed +qed + +lemma poly_tuple_eval_Cons: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(poly_tuple_eval (f#fs) as) = (eval_at_point R as f)#(poly_tuple_eval fs as)" + using assms poly_tuple_eval_def + by (metis list.simps(9)) + +definition poly_tuple_pullback :: + "nat \ 'a list set \ ('a, nat) mvar_poly list \ 'a list set" where +"poly_tuple_pullback n S fs = ((poly_tuple_eval fs) -` S) \ (carrier (R\<^bsup>n\<^esup>)) " + +lemma poly_pullbackE: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "S \ carrier (R\<^bsup>m\<^esup>)" + shows "poly_tuple_pullback n S fs \ carrier (R\<^bsup>n\<^esup>)" + using poly_tuple_pullback_def assms + by blast + +lemma poly_pullbackE': + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "S \ carrier (R\<^bsup>m\<^esup>)" + assumes "as \ poly_tuple_pullback n S fs" + shows "as \ carrier (R\<^bsup>n\<^esup>)" + "poly_tuple_eval fs as \ S" + using assms + apply (meson poly_pullbackE subsetD) +proof- + have "as \ poly_tuple_eval fs -` S" + using assms unfolding poly_tuple_pullback_def + by blast + then show "poly_tuple_eval fs as \ S" + by blast +qed + +lemma poly_pullbackI: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "S \ carrier (R\<^bsup>m\<^esup>)" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + assumes "poly_tuple_eval fs as \ S" + shows "as \ poly_tuple_pullback n S fs" + using assms + unfolding poly_tuple_pullback_def + by blast + + + +end + +text\coordinate permutations as pullbacks. The point here is to realize that permutations of +indices are just pullbacks (or pushforwards) by particular polynomial maps\ + +abbreviation pvar_list where +"pvar_list R n \ map (pvar R) (index_list n)" + +lemma pvar_list_elements: + assumes "i < n" + shows "pvar_list R n ! i = pvar R i" + by (simp add: assms index_list_indices index_list_length) + +lemma pvar_list_length: +"length (pvar_list R n) = n" + by (simp add: index_list_length) + +context cring_coord_rings +begin + +lemma pvar_list_is_poly_tuple: + shows "is_poly_tuple n (pvar_list R n)" + unfolding is_poly_tuple_def +proof fix x + assume A: "x \ set (pvar_list R n)" + have "set (index_list n) = {.. x = pvar R i" + using A pvar_list_elements[of _ n R] pvar_list_length[of R n] + by (metis in_set_conv_nth) + then show "x \ carrier (R[\\<^bsub>n\<^esub>])" + using pvar_closed + by blast +qed + +lemma permutation_of_poly_list_is_poly_list: + assumes "is_poly_tuple k fs" + assumes "\ permutes {..< length fs}" + shows "is_poly_tuple k (permute_list \ fs)" + unfolding is_poly_tuple_def +proof- + show "set (permute_list \ fs) \ carrier (coord_ring R k)" + using assms is_poly_tuple_def set_permute_list + by blast +qed + +lemma permutation_of_poly_listE: + assumes "is_poly_tuple k fs" + assumes "\ permutes {..< length fs}" + assumes "i < length fs" + shows "(permute_list \ fs) ! i = fs ! (\ i)" + using assms permute_list_nth + by blast + +lemma pushforward_by_permutation_of_poly_list: + assumes "is_poly_tuple k fs" + assumes "\ permutes {..< length fs}" + assumes "as \ carrier (R\<^bsup>k\<^esup>)" + shows "poly_tuple_eval (permute_list \ fs) as = permute_list \ (poly_tuple_eval fs as)" + using assms unfolding poly_tuple_eval_def + by (metis permute_list_map) + +lemma pushforward_by_pvar_list: + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_tuple_eval (pvar_list R n) as = as" + using assms pvar_list_elements[of _ n R] unfolding poly_tuple_eval_def using eval_pvar[of _ n as] + by (metis (mono_tags, lifting) cartesian_power_car_memE length_map nth_equalityI nth_map pvar_list_length) + +lemma pushforward_by_permuted_pvar_list: + assumes "\ permutes {..< n}" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_tuple_eval (permute_list \ (pvar_list R n)) as = permute_list \ as" + by (metis assms pushforward_by_permutation_of_poly_list + pushforward_by_pvar_list pvar_list_is_poly_tuple pvar_list_length) + +lemma pullback_by_permutation_of_poly_list: + assumes "\ permutes {..< n}" + assumes "S \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_tuple_pullback n S (permute_list \ (pvar_list R n)) = + permute_list (fun_inv \) ` S" +proof + show "poly_tuple_pullback n S (permute_list \ (pvar_list R n)) \ permute_list (fun_inv \) ` S" + proof fix x + assume A: " x \ poly_tuple_pullback n S (permute_list \ (pvar_list R n))" + then obtain y where y_def: "y \ S \ poly_tuple_eval (permute_list \ (pvar_list R n)) x = y" + by (metis assms length_permute_list + permutation_of_poly_list_is_poly_list poly_pullbackE'(2) pvar_list_is_poly_tuple + pvar_list_length) + then have 0: "y = permute_list \ x" + by (metis A assms length_permute_list + permutation_of_poly_list_is_poly_list poly_pullbackE'(1) pushforward_by_permuted_pvar_list pvar_list_is_poly_tuple pvar_list_length) + have 1: "length x = n" + using A + by (metis "0" length_permute_list poly_tuple_evalE' pvar_list_length y_def) + then have "{..) y = x" + using 0 permutes_inv_o(1)[of \ "{..< n}"] permute_list_id[of x] permutes_inv[of \ "{..)" x \ ] + unfolding fun_inv_def + by metis + then show " x \ permute_list (fun_inv \) ` S" + using y_def by blast + qed + show "permute_list (fun_inv \) ` S \ poly_tuple_pullback n S (permute_list \ (pvar_list R n))" + proof fix x assume A: "x \ permute_list (fun_inv \) ` S" + then obtain y where y_def: "y \ S \ x = permute_list (fun_inv \) y" + by blast + have 0: "(fun_inv \) permutes {.. x = permute_list \ (permute_list (fun_inv \) y)" + by (simp add: y_def) + have 2: "length y = n" + using y_def A assms cartesian_power_car_memE + by blast + have 3: "\ permutes {.. x = y" + using assms(2) permute_list_id[of y] permute_list_compose[of \ y "(fun_inv \)" ] + 3 2 1 0 permutes_inv_o(2)[of \ "{..< n}"] + unfolding fun_inv_def + by metis + have 5: "x \ carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI) + using A 0 assms + apply (metis "2" "4" length_permute_list) + using A 0 assms + by (smt "2" in_set_conv_nth neq0_conv poly_tuple_evalE pushforward_by_pvar_list + pvar_list_is_poly_tuple pvar_list_length set_permute_list subset_iff y_def) + then have 6: "poly_tuple_eval (permute_list \ (pvar_list R n)) x = y" + using 4 assms pushforward_by_permuted_pvar_list[of \ n x] + by blast + then show "x \ poly_tuple_pullback n S (permute_list \ (pvar_list R n))" + using 5 y_def unfolding poly_tuple_pullback_def + by blast + qed +qed + +lemma pullback_by_permutation_of_poly_list': + assumes "\ permutes {..< n}" + assumes "S \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list R n)) = + permute_list \ ` S" +proof- + have 0: "(fun_inv (fun_inv \)) = \" + using assms unfolding fun_inv_def + using permutes_inv_inv + by blast + have 1: "fun_inv \ permutes {.." n S] + by presburger +qed + + + (**********************************************************************) + (**********************************************************************) + subsection\Composing Polynomial Tuples With Polynomials\ + (**********************************************************************) + (**********************************************************************) + +text\composition of a multivaribale polynomial by a list of polynomials\ + +definition poly_compose :: + "nat \ nat \ ('a, nat) mvar_poly list \ ('a, nat) mvar_poly \ ('a, nat) mvar_poly" where +"poly_compose n m fs = indexed_poly_induced_morphism {.. s. R.indexed_const s) (\i. fs!i) " + +lemma poly_compose_var: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "j < n" + shows "poly_compose n m fs (pvar R j) = (fs!j)" +proof- + have 0: "cring (coord_ring R m)" + using R.Pring_is_cring R.is_cring + unfolding coord_ring_def by blast + have 1: "(!) fs \ {.. carrier (coord_ring R m)" + using assms is_poly_tuple_def + by auto + have 2: "ring_hom_ring R (coord_ring R m) coord_const" + using indexed_const_ring_hom coord_const_ring_hom by blast + have "\i\{..i. fs!i)" "{.. s. R.indexed_const s)" "poly_compose n m fs"] + poly_compose_def + by (metis var_to_IP_def) + then show ?thesis + using assms + unfolding poly_compose_def var_to_IP_def + by blast +qed + +lemma Pring_universal_prop_assms: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + shows "(\i. fs!i) \ {.. carrier (coord_ring R m)" + "ring_hom_ring R (coord_ring R m) coord_const" +proof + show "\x. x \ {.. fs ! x \ carrier (coord_ring R m)" + using assms is_poly_tupleE by blast + show "ring_hom_ring R (coord_ring R m) coord_const" + using R.indexed_const_ring_hom coord_const_ring_hom by blast +qed + +lemma poly_compose_ring_hom: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + shows "(ring_hom_ring (R[\\<^bsub>n\<^esub>]) (coord_ring R m) (poly_compose n m fs))" + using Pring_universal_prop_assms[of n fs] assms + R.Pring_universal_prop(1)[of "(coord_ring R m)" "(\i. fs!i)" "{.. carrier (R[\\<^bsub>n\<^esub>])" + shows "(poly_compose n m fs f) \ carrier (coord_ring R m)" +proof- + have "poly_compose n m fs \ carrier (R[\\<^bsub>n\<^esub>]) \ carrier (R [\\<^bsub>m\<^esub>])" + using poly_compose_ring_hom[of m fs n] assms + unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def + by blast + then show ?thesis using assms by blast +qed + +lemma poly_compose_add: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (R[\\<^bsub>n\<^esub>])" + shows "poly_compose n m fs (f \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> g) = (poly_compose n m fs f) \\<^bsub>coord_ring R m\<^esub> (poly_compose n m fs g)" + using assms poly_compose_ring_hom ring_hom_add + by (metis (mono_tags, lifting) ring_hom_ring.homh) + +lemma poly_compose_mult: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "g \ carrier (R[\\<^bsub>n\<^esub>])" + shows "poly_compose n m fs (f \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> g) = (poly_compose n m fs f) \\<^bsub>coord_ring R m\<^esub> (poly_compose n m fs g)" + using assms poly_compose_ring_hom ring_hom_mult + by (metis (mono_tags, lifting) ring_hom_ring.homh) + +lemma poly_compose_indexed_pmult: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "i < n" + shows "poly_compose n m fs (f \ i) = (poly_compose n m fs f) \\<^bsub>coord_ring R m\<^esub> (fs!i)" +proof- + have "(f \ i) = f \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i" + using assms pvar_indexed_pmult + by blast + then show ?thesis using poly_compose_mult poly_compose_var assms + by (metis pvar_closed) +qed + +lemma poly_compose_const: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "a \ carrier R" + shows "poly_compose n m fs (coord_const a) = coord_const a" + using R.Pring_universal_prop(3)[of "(coord_ring R m)" "(\i. fs!i)" "{..evaluating polynomial compositions\ + +lemma poly_compose_eval: + assumes "is_poly_tuple m fs" + assumes "length fs = n" + assumes "f \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "as \ carrier (R\<^bsup>m\<^esup>)" + shows "eval_at_point R (poly_tuple_eval fs as) f = eval_at_point R as (poly_compose n m fs f)" + apply(rule R.indexed_pset.induct[of f "{..k. k \ carrier R \ eval_at_poly R (coord_const k) (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs (coord_const k)) as" + using assms + by (metis (no_types, lifting) eval_at_point_factored poly_compose_const R.total_eval_const) + show " \p Q. p \ Pring_set R {.. + eval_at_poly R p (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs p) as \ + Q \ Pring_set R {.. + eval_at_poly R Q (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs Q) as \ + eval_at_poly R (p \ Q) (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs (p \ Q)) as" + proof- + fix p Q + assume A0: "p \ Pring_set R {.. Pring_set R {.. Q) (poly_tuple_eval fs as) = eval_at_poly R p (poly_tuple_eval fs as) \ eval_at_poly R Q (poly_tuple_eval fs as)" + using A0 A1 A2 A3 + eval_at_point_add[of "(poly_tuple_eval fs as)" n p Q] + by (metis R.Pring_add R.Pring_car assms(2) assms(3) assms(4) assms coord_ring_def neq0_conv poly_tuple_eval_closed) + have A5: "poly_compose n m fs (p \ Q) = poly_compose n m fs p \\<^bsub>coord_ring R m\<^esub> poly_compose n m fs Q" + using assms poly_compose_add + by (metis A0 A2 R.Pring_add R.Pring_car coord_ring_def) + have A6: " eval_at_poly R (poly_compose n m fs (p \ Q)) as = eval_at_poly R (poly_compose n m fs p) as \ eval_at_poly R (poly_compose n m fs Q) as" + proof- + have 0: " as \ carrier (R\<^bsup>m\<^esup>)" + by (simp add: assms) + have 1: "poly_compose n m fs p \ carrier (coord_ring R m)" + using A0 R.Pring_car assms(1) assms(2) assms(3) assms(4) poly_compose_closed + by (metis coord_ring_def) + have 2: "poly_compose n m fs Q \ carrier (coord_ring R m)" + using A2 R.Pring_car assms(1) assms(2) assms(3) assms(4) poly_compose_closed + by (metis coord_ring_def) + show ?thesis + using 0 1 2 eval_at_point_add[of as m "(poly_compose n m fs p)" "(poly_compose n m fs Q)"] + A5 + by presburger + qed + show "eval_at_poly R (p \ Q) (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs (p \ Q)) as" + using A5 A6 A3 A1 A4 + by presburger + qed + show "\p i. p \ Pring_set R {.. + eval_at_poly R p (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs p) as \ + i \ {.. eval_at_poly R (p \ i) (poly_tuple_eval fs as) = eval_at_poly R (poly_compose n m fs (p \ i)) as" + using assms poly_compose_indexed_pmult eval_at_point_indexed_pmult + by (smt R.Pring_car coord_ring_def eval_at_point_mult is_poly_tupleE lessThan_iff neq0_conv poly_compose_closed poly_tuple_evalE'' poly_tuple_eval_closed) +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Extensional Polynomial Maps\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\Polynomial Maps between powers of a ring\ + +definition poly_map :: "nat \ ('a, nat) mvar_poly list \ 'a list \ 'a list" where +"poly_map n fs = (\a \ carrier (R\<^bsup>n\<^esup>). poly_tuple_eval fs a)" + +lemma poly_map_is_struct_map: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + shows "poly_map n fs \ struct_maps (R\<^bsup>n\<^esup>) (R\<^bsup>m\<^esup>)" + apply(rule struct_maps_memI) + unfolding poly_map_def using assms + apply (metis poly_tuple_eval_closed restrict_apply') + by (meson restrict_apply) + +lemma poly_map_closed: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "as \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_map n fs as \ carrier (R\<^bsup>m\<^esup>)" + using assms + by (meson poly_map_is_struct_map struct_maps_memE(1)) + +definition poly_maps :: "nat \ nat \ ('a list \ 'a list) set" where +"poly_maps n m = {F. (\ fs. length fs = m \ is_poly_tuple n fs \ F = poly_map n fs)}" + +lemma poly_maps_memE: + assumes "F \ poly_maps n m" + obtains fs where "length fs = m \ is_poly_tuple n fs \ F = poly_map n fs" + using assms unfolding poly_maps_def by blast + +lemma poly_maps_memI: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "F = poly_map n fs" + shows "F \ poly_maps n m" + using assms unfolding poly_maps_def by blast + +lemma poly_map_compose_closed: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "is_poly_tuple k gs" + assumes "length gs = n" + shows "is_poly_tuple k (map (poly_compose n k gs) fs)" + unfolding is_poly_tuple_def +proof fix y assume A: "y \ set (map (poly_compose n k gs) fs)" + then obtain f where f_def: "f \ set fs \ y = poly_compose n k gs f" + by (smt in_set_conv_nth length_map nth_map) + then show "y \ carrier (coord_ring R k)" + using assms poly_compose_closed + by (metis in_set_conv_nth is_poly_tupleE ) +qed + +lemma poly_map_compose_closed': + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "is_poly_tuple k gs" + assumes "length gs = n" + shows "poly_map k (map (poly_compose n k gs) fs) \ poly_maps k m" + apply(rule poly_maps_memI[of _ "map (poly_compose n k gs) fs"]) + using poly_map_compose_closed[of n fs m k gs] assms apply blast + apply (simp add: assms) + by auto + +lemma poly_map_pullback_char: + assumes "is_poly_tuple n fs" + assumes "length fs = m" + assumes "is_poly_tuple k gs" + assumes "length gs = n" + shows "(pullback (R\<^bsup>k\<^esup>) (poly_map k gs) (poly_map n fs)) = + poly_map k (map (poly_compose n k gs) fs)" +proof(rule ext) + fix x + show "pullback (R\<^bsup>k\<^esup>) (poly_map k gs) (poly_map n fs) x = + poly_map k (map (poly_compose n k gs) fs) x" + proof(cases "x \ carrier (R\<^bsup>k\<^esup>)") + case True + have 0: "length (pullback (R\<^bsup>k\<^esup>) (poly_map k gs) (poly_map n fs) x) = m" + using True assms poly_map_closed cartesian_power_car_memE + unfolding pullback_def + by (metis (mono_tags, lifting) compose_eq) + have 1: "is_poly_tuple k (map (poly_compose n k gs) fs)" + by (simp add: assms poly_map_compose_closed) + have 2: "length (map (poly_compose n k gs) fs) = m" + using assms length_map by auto + have 3: "\i. i < m \ + (pullback (R\<^bsup>k\<^esup>) (poly_map k gs) (poly_map n fs) x)! i = + eval_at_point R (poly_map k gs x) (fs ! i)" + unfolding pullback_def poly_map_def poly_tuple_eval_def + using assms True + by (smt compose_eq nth_map poly_tuple_eval_closed poly_tuple_eval_def restrict_apply') + have 4: "\i. i < m \ + poly_map k (map (poly_compose n k gs) fs) x ! i = + eval_at_point R (poly_map k gs x) (fs ! i)" + unfolding poly_map_def poly_tuple_eval_def using True assms + by (smt "2" cring_coord_rings.is_poly_tuple_def cring_coord_rings_axioms neq0_conv + nth_map nth_mem poly_compose_eval poly_tuple_eval_def restrict_apply' subset_code(1)) + show ?thesis using 0 1 2 3 4 assms True + by (metis cartesian_power_car_memE nth_equalityI poly_map_closed) + next + case False + then show ?thesis + unfolding poly_map_def pullback_def + by (metis affine_alg_set_empty compose_extensional extensional_restrict poly_map_def restrict_def) + qed +qed + +lemma poly_map_pullback_closed: + assumes "F \ poly_maps n m" + assumes "G \ poly_maps k n" + shows "(pullback (R\<^bsup>k\<^esup>) G F) \ poly_maps k m" + by (metis assms poly_map_compose_closed' + poly_map_pullback_char poly_maps_memE) + +lemma poly_map_cons: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_map n (f#fs) a = (eval_at_point R a f)#poly_map n fs a" + unfolding poly_map_def poly_tuple_eval_def + by (metis (mono_tags, lifting) assms list.simps(9) restrict_apply') + +lemma poly_map_append: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_map n (fs@gs) a = (poly_map n fs a) @ (poly_map n gs a)" +proof(induction fs) + case Nil + then show ?case + using assms unfolding poly_map_def poly_tuple_eval_def + by (metis (no_types, lifting) map_append restrict_apply') +next + case (Cons f fs) + have "poly_map n ((f # fs) @ gs) a = (eval_at_point R a f)#(poly_map n (fs@gs) a)" + using poly_map_cons + by (metis append_Cons assms) + hence "poly_map n ((f # fs) @ gs) a = (eval_at_point R a f)#(poly_map n fs a)@(poly_map n gs a)" + using Cons.IH by metis + thus ?case + by (metis Cons_eq_appendI assms poly_map_cons) +qed + +(**************************************************************************************************) +(**************************************************************************************************) +section\Nesting of Polynomial Rings\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma poly_ring_car_mono: + assumes "n \ m" + shows "carrier (R[\\<^bsub>n\<^esub>]) \ carrier (coord_ring R m)" + using R.Pring_carrier_subset + unfolding coord_ring_def + by (simp add: R.Pring_car R.Pring_carrier_subset assms) + +lemma poly_ring_car_mono'[simp]: + shows "carrier (R[\\<^bsub>n\<^esub>]) \ carrier (R[\\<^bsub>Suc n\<^esub>])" + "carrier (R[\\<^bsub>n\<^esub>]) \ carrier (R[\\<^bsub>n+m\<^esub>])" + using poly_ring_car_mono + apply simp + using poly_ring_car_mono + by simp + +lemma poly_ring_add_mono: + assumes "n \ m" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (R[\\<^bsub>n\<^esub>])" + shows "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>coord_ring R m\<^esub> B" + using assms unfolding coord_ring_def + by (metis R.Pring_add_eq) + +lemma poly_ring_add_mono': + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (R[\\<^bsub>n\<^esub>])" + shows "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> B" + "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> B" + using assms unfolding coord_ring_def + apply (metis R.Pring_add_eq) + by (metis R.Pring_add_eq) + +lemma poly_ring_times_mono: + assumes "n \ m" + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (R[\\<^bsub>n\<^esub>])" + shows "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>coord_ring R m\<^esub> B" + using assms unfolding coord_ring_def + by (metis R.Pring_mult_eq) + +lemma poly_ring_times_mono': + assumes "A \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "B \ carrier (R[\\<^bsub>n\<^esub>])" + shows "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> B" + "A \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> B = A \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> B" + using assms unfolding coord_ring_def + apply (metis R.Pring_mult_eq) + by (metis R.Pring_mult_eq) + +lemma poly_ring_one_mono: + assumes "n \ m" + shows "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> = \\<^bsub>coord_ring R m\<^esub>" + by (metis R.Pring_one coord_ring_def) + +lemma poly_ring_zero_mono: + assumes "n \ m" + shows "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> = \\<^bsub>coord_ring R m\<^esub>" + using R.Pring_zero_eq + by (metis coord_ring_def) + +text\replacing the variables in a polynomial with new variables\ + +definition shift_vars :: "nat \ nat \ ('a, nat) mvar_poly \ ('a, nat) mvar_poly" where +"shift_vars n m p = indexed_poly_induced_morphism {..\<^bsub>n+m\<^esub>]) coord_const (\i. pvar R (i + m)) p" + +lemma shift_vars_char: + shows "(ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n+m\<^esub>]) (shift_vars n m))" + "(\i \ {..a \ carrier R. (shift_vars n m) (R.indexed_const a) = (coord_const a))" +proof- + have 1: "(\i. pvar R (i + m)) \ {.. carrier (R[\\<^bsub>n+m\<^esub>])" + proof fix x + assume "x \ {.. carrier (R[\\<^bsub>n+m\<^esub>])" + using pvar_closed by blast + qed + have 2: "ring_hom_ring R (R[\\<^bsub>n+m\<^esub>]) coord_const" + using R.indexed_const_ring_hom unfolding coord_ring_def + by blast + have 3: "shift_vars n m = indexed_poly_induced_morphism {..\<^bsub>n+m\<^esub>]) coord_const (\i. pvar R (i + m))" + unfolding shift_vars_def + by blast + show "(ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>n+m\<^esub>]) (shift_vars n m))" + using 1 2 3 R.Pring_universal_prop[of "(R[\\<^bsub>n+m\<^esub>])" "(\i. pvar R (i + m))" "{..i \ {..\<^bsub>n+m\<^esub>])" "(\i. pvar R (i + m))" "{..a \ carrier R. (shift_vars n m) (R.indexed_const a) = (coord_const a))" + using 1 2 3 R.Pring_universal_prop[of "(R[\\<^bsub>n+m\<^esub>])" "(\i. pvar R (i + m))" "{.. carrier R" + shows "shift_vars n m (coord_const a) = coord_const a" + using assms(1) shift_vars_char(3) by blast + +lemma shift_vars_pvar: + assumes "i \ {.. carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "shift_vars n m (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) = shift_vars n m p \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> shift_vars n m Q" + using assms shift_vars_char(1)[of n m] + unfolding ring_hom_ring_def ring_hom_ring_axioms_def + using ring_hom_add + by (metis (no_types, lifting)) + +lemma shift_vars_mult: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + shows "shift_vars n m (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) = shift_vars n m p \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> shift_vars n m Q" + using assms shift_vars_char(1)[of n m] + unfolding ring_hom_ring_def ring_hom_ring_axioms_def unfolding coord_ring_def + using ring_hom_mult + by metis + +lemma shift_vars_indexed_pmult: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "i \ {.. i) = (shift_vars n m p) \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> (pvar R (i + m))" +proof- + have "(p \ i) = p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (pvar R i)" + using assms pvar_indexed_pmult by blast + then show ?thesis + using shift_vars_mult shift_vars_pvar assms unfolding coord_ring_def + by (metis R.Pring_var_closed) +qed + +lemma shift_vars_closed: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "shift_vars n m p \ carrier (R[\\<^bsub>n+m\<^esub>])" + using assms shift_vars_char(1)[of n m] ring_hom_closed[of "shift_vars n m"] + unfolding ring_hom_ring_def ring_hom_ring_axioms_def + by blast + +lemma shift_vars_eval: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "a \ carrier (R\<^bsup>m\<^esup>)" + assumes "b \ carrier (R\<^bsup>n\<^esup>)" + shows "eval_at_point R (a@b) (shift_vars n m p) = eval_at_point R b p" + apply(rule R.indexed_pset.induct[of p "{..k. k \ carrier R \ eval_at_poly R (shift_vars n m (coord_const k)) (a @ b) = eval_at_poly R (coord_const k) b" + proof- + fix k + have 0: "(a @ b) \ carrier (R\<^bsup>n + m\<^esup>)" + using assms + by (metis add.commute cartesian_product_closed') + assume A: "k \ carrier R" + then show "eval_at_poly R (shift_vars n m (coord_const k)) (a @ b) = eval_at_poly R (coord_const k) b" + using assms shift_vars_constant + eval_at_point_const[of k "(a @ b)" "m + n"] + eval_at_point_const[of k "b" n] 0 + by (metis eval_at_point_const) + qed + show "\p Q. p \ Pring_set R {.. + eval_at_poly R (shift_vars n m p) (a @ b) = eval_at_poly R p b \ + Q \ Pring_set R {.. + eval_at_poly R (shift_vars n m Q) (a @ b) = eval_at_poly R Q b \ + eval_at_poly R (shift_vars n m (p \ Q)) (a @ b) = eval_at_poly R (p \ Q) b" + proof- fix p Q + assume A0: " p \ Pring_set R {.. Pring_set R {.. Q) b = eval_at_poly R p b \ eval_at_poly R Q b" + using A0 A2 assms eval_at_point_add[of b n p Q] + by (metis R.Pring_add R.Pring_car coord_ring_def) + have A5: "(shift_vars n m (p \ Q)) = (shift_vars n m p) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (shift_vars n m Q)" + using A0 A2 R.Pring_add R.Pring_car assms(1) shift_vars_add unfolding coord_ring_def + by metis + have A6: " eval_at_poly R (shift_vars n m (p \ Q)) (a @ b) = + eval_at_poly R (shift_vars n m p) (a @ b) \ eval_at_poly R (shift_vars n m Q) (a @ b) " + using A5 eval_at_point_add shift_vars_closed A0 A2 R.Pring_car add.commute + assms unfolding coord_ring_def + by (metis R.Pring_add cartesian_power_concat(1)) + have A7: " eval_at_poly R (shift_vars n m (p \ Q)) (a @ b) = + eval_at_poly R p b \ eval_at_poly R Q b " + using A6 A1 A3 by presburger + then show " eval_at_poly R (shift_vars n m (p \ Q)) (a @ b) = eval_at_poly R (p \ Q) b " + using A4 + by presburger + qed + show "\p i. p \ Pring_set R {.. + eval_at_poly R (shift_vars n m p) (a @ b) = eval_at_poly R p b \ + i \ {.. eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = eval_at_poly R (p \ i) b" + proof- fix p i + assume A0: "p \ Pring_set R {.. carrier (R[\\<^bsub>n\<^esub>])" + using R.Pring_car unfolding coord_ring_def + by blast + assume A1: " eval_at_poly R (shift_vars n m p) (a @ b) = eval_at_poly R p b" + assume A2: "i \ {.. i)) = (shift_vars n m p) \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> (pvar R (i + m))" + using A0' shift_vars_indexed_pmult A2 assms(1) + by blast + have A4: "eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = + eval_at_poly R ( (shift_vars n m p) \\<^bsub>R[\\<^bsub>n+m\<^esub>]\<^esub> (pvar R (i + m))) (a@b)" + using A3 + by presburger + have A5: "a@b \ carrier (R\<^bsup>n+m\<^esup>)" + using assms(2) assms(3) cartesian_power_concat(2) by blast + have A6: "eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = + eval_at_poly R p b \ eval_at_poly R (pvar R (i + m)) (a @ b)" + using A5 A0' eval_at_point_mult[of "a@b" "n+m" "shift_vars n m p" "pvar R (i + m)"] + unfolding A4 by (metis A1 A2 Groups.add_ac(2) lessThan_iff local.pvar_closed nat_add_left_cancel_less shift_vars_closed) + have A7: " eval_at_poly R (pvar R (i + m)) (a @ b) = (a@b)!(i+m)" + proof- + have "i < n" + using assms A2 by blast + then have "i + m < n + m " + using add_less_cancel_right + by blast + then show ?thesis + using A5 eval_pvar[of "i+m" "n+m" "a@b"] + by blast + qed + then have A8: "eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = eval_at_poly R p b \ ((a @ b)!(i+m))" + using A6 by presburger + have A9: "eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = eval_at_poly R p b \ (b!i)" + proof- + have "length a = m" + using assms cartesian_power_car_memE by blast + then have "(a @ b)!(i+m) = b!i" + by (metis add.commute nth_append_length_plus) + then show ?thesis + using A8 + by presburger + qed + show " eval_at_poly R (shift_vars n m (p \ i)) (a @ b) = eval_at_poly R (p \ i) b" + proof- + have "i < n" + using A2 assms + by blast + then have "eval_at_poly R (p \ i) b = eval_at_poly R p b \ (b!i)" + using assms A0' eval_at_point_indexed_pmult + by blast + then show ?thesis using A9 + by presburger + qed + qed +qed + + +text\Evaluating a polynomial from a lower poly ring in a higher power:\ + +lemma poly_eval_cartesian_prod: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "b \ carrier (R\<^bsup>m\<^esup>)" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "eval_at_point R a p = eval_at_point R (a@b) p" + apply(rule coord_ring_induct[of p n]) + using assms apply blast +proof- + have 0: "a@b \ carrier (R\<^bsup>n + m\<^esup>)" + using assms cartesian_product_closed' by blast + show "\aa. aa \ carrier R \ eval_at_poly R (coord_const aa) a = eval_at_poly R (coord_const aa) (a @ b)" + proof- fix c assume "c \ carrier R" + show "eval_at_poly R (coord_const c) a = eval_at_poly R (coord_const c) (a @ b)" + using eval_at_point_const[of c a n] eval_at_point_const[of c "a@b" "n+m"] 0 + \c \ carrier R\ assms(2) assms(1) by presburger + qed + show "\i Q. Q \ carrier (R[\\<^bsub>n\<^esub>]) \ + eval_at_poly R Q a = eval_at_poly R Q (a @ b) \ + i < n \ eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) a = eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) (a @ b)" + proof- + fix i Q + assume A0: "Q \ carrier (R[\\<^bsub>n\<^esub>])" + assume A1: "eval_at_poly R Q a = eval_at_poly R Q (a @ b)" + assume A2: "i < n" + have LHS: "eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) a = eval_at_poly R Q a \ (a!i)" + by (metis A0 A2 assms eval_at_point_indexed_pmult pvar_indexed_pmult) + have RHS: "eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) (a @ b) = eval_at_poly R Q (a@b) \ ((a@b)!i)" + by (smt "0" A0 A2 add.commute eval_at_point_indexed_pmult le_add1 poly_ring_car_mono + pvar_indexed_pmult subsetD trans_less_add2) + show "eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) a = eval_at_poly R (Q \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i) (a @ b)" + proof- + have "length a > i" using A2 assms + using cartesian_power_car_memE by blast + then have "a!i = (a@b)!i" + by (metis nth_append) + then show ?thesis + using LHS RHS A1 + by presburger + qed + qed + show "\Q0 Q1. + Q0 \ carrier (R[\\<^bsub>n\<^esub>]) \ + Q1 \ carrier (R[\\<^bsub>n\<^esub>]) \ + eval_at_poly R Q0 a = eval_at_poly R Q0 (a @ b) \ + eval_at_poly R Q1 a = eval_at_poly R Q1 (a @ b) \ + eval_at_poly R (Q0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q1) a = eval_at_poly R (Q0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q1) (a @ b)" + proof- + fix Q0 Q1 + assume A0: "eval_at_poly R Q0 a = eval_at_poly R Q0 (a @ b)" + assume A1: "eval_at_poly R Q1 a = eval_at_poly R Q1 (a @ b)" + assume A2: "Q0 \ carrier (R[\\<^bsub>n\<^esub>])" + assume A3: "Q1 \ carrier (R[\\<^bsub>n\<^esub>])" + show "eval_at_poly R (Q0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q1) a = eval_at_poly R (Q0 \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q1) (a @ b)" + using A0 A1 A2 A3 assms eval_at_point_add[of _ n Q0 Q1] 0 unfolding coord_ring_def + by (metis (no_types, lifting) R.Pring_add_eq basic_trans_rules(31) coord_ring_def eval_at_point_add le_add1 poly_ring_car_mono) + qed +qed + +text\Evaluating polynomials at points in higher powers:\ + +lemma eval_at_points_higher_pow: + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "k \ n" + assumes "a \ carrier (R\<^bsup>k\<^esup>)" + shows "eval_at_point R a p = eval_at_point R (take n a) p" + using poly_eval_cartesian_prod[of "take n a" n "drop n a" "k - n" p] assms + by (metis (no_types, lifting) append_take_drop_id cartesian_power_car_memE cartesian_power_car_memE'' + cartesian_power_car_memI length_drop set_drop_subset subset_trans take_closed) + + +subsection\ Diagonal sets in even powers of \R\\ + +text\ + In this section, by a diagonal set in $R^(2m)$ we will mean the set of points $(x,x)$, + where $x \in R^m$. This is slightly different from the standard definition. Introducing these + sets will be useful for reasoning about multiplicative inverses of functions later on. +\ + +definition diagonal :: "nat \ 'a list set" where +"diagonal m = {x \ carrier (R\<^bsup>m+m\<^esup>). take m x = drop m x}" + +lemma diagonalE: + assumes "x \ diagonal m" + shows "x = (take m x)@(take m x)" + "x \ carrier (R\<^bsup>m+m\<^esup>)" + "take m x \ carrier (R\<^bsup>m\<^esup>)" + "\i. i < m \ x!i = x!(i + m)" + apply (metis (mono_tags, lifting) append_take_drop_id assms(1) diagonal_def mem_Collect_eq ) + using assms diagonal_def + apply blast + apply(rule cartesian_power_car_memI) + using assms unfolding diagonal_def + apply (metis (no_types, lifting) cartesian_power_car_memE le_add2 mem_Collect_eq take_closed) +proof- + show "set (take m x) \ carrier R" + proof fix a + assume "a \ set (take m x)" + then have "a \ set x" + by (meson in_set_takeD) + then show "a \ carrier R" + using assms unfolding diagonal_def using cartesian_power_car_memE'[of x] + by (smt cartesian_power_car_memE in_set_conv_nth mem_Collect_eq) + qed + show "\i. i < m \ x!i = x!(i + m)" + proof- fix i + assume A: "i < m" + have 0: "x = (take m x)@(take m x)" + using assms diagonal_def[of m] + by (metis (mono_tags, lifting) append_take_drop_id mem_Collect_eq) + then have 1: "x!i = take m x ! i" + by (metis A nth_take) + have 2: "length x = m + m" + using assms(1) cartesian_power_car_memE diagonal_def by blast + have 3: "take m x = drop m x" + by (metis "0" append_take_drop_id same_append_eq) + have 4: "drop m x ! i = x ! (i + m)" + by (metis "2" add.commute le_add1 nth_drop) + then show "x!i = x!(i + m)" + using "1" "3" by presburger + qed +qed + +lemma diagonalI: + assumes "x = (take m x)@(take m x)" + assumes "take m x \ carrier (R\<^bsup>m\<^esup>)" + shows "x \ diagonal m" + unfolding diagonal_def using assms + by (metis (mono_tags, lifting) append_eq_conv_conj cartesian_power_car_memE + cartesian_power_car_memI'' length_append mem_Collect_eq) + +definition diag_def_poly :: "nat \ nat \('a, nat) mvar_poly" where +"diag_def_poly n i = pvar R i \\<^bsub>coord_ring R (n + n)\<^esub> pvar R (i + n)" + +lemma diag_def_poly_closed: + assumes "i < n" + shows "diag_def_poly n i \ carrier (coord_ring R (n + n))" + using assms unfolding diag_def_poly_def coord_ring_def + by (metis (no_types, lifting) MP.minus_closed add.assoc add_leD1 coord_ring_def less_add_eq_less local.pvar_closed nat_less_le not_add_less1) + +lemma diag_def_poly_eval: + assumes "i < n" + assumes "x \ carrier (R\<^bsup>n+n\<^esup>)" + shows "eval_at_point R x (diag_def_poly n i) = (x!i) \ (x!(i + n))" + + using assms diag_def_poly_def[of n i] + eval_at_point_subtract[of x "n + n" "pvar R i" "pvar R (i + n)"] eval_pvar[of i "n + n"] + eval_pvar[of "i+n" "n + n"] pvar_closed[of i "n + n"] pvar_closed[of "i + n" "n + n"] + by (metis add_less_cancel_right trans_less_add2) + +definition diag_def_poly_set :: "nat \ ('a, nat) mvar_poly set" where +"diag_def_poly_set n = diag_def_poly n ` {.. carrier (coord_ring R (n + n))" +proof fix x + assume "x \ diag_def_poly_set n" + then obtain i where i_def: "i < n \ x = diag_def_poly n i" + unfolding diag_def_poly_set_def + by blast + then show "x \ carrier (coord_ring R (n + n))" + using diag_def_poly_closed + by blast +qed + +lemma diag_def_poly_set_finite: +"finite (diag_def_poly_set n)" + unfolding diag_def_poly_set_def + by blast + +lemma diag_def_poly_eval_at_diagonal: + assumes "x \ diagonal n" + assumes "i < n" + shows "eval_at_point R x (diag_def_poly n i) = \" +proof- + have "x!i = x!(i + n)" + using assms diagonalE(4) by blast + then show ?thesis + by (metis assms(1) assms(2) cartesian_power_car_memE cartesian_power_car_memE' cring_coord_rings.diag_def_poly_eval cring_coord_rings_axioms diagonalE(2) point_to_polysE point_to_polysE' pvar_trans_eval trans_less_add2) +qed + +lemma diagonal_as_affine_alg_set: + shows "diagonal n = affine_alg_set R (n + n) (diag_def_poly_set n)" +proof + show "diagonal n \ affine_alg_set R (n + n) (diag_def_poly_set n)" + proof fix x assume A: "x \ diagonal n" + show " x \ affine_alg_set R (n + n) (diag_def_poly_set n)" + apply(rule affine_alg_set_memI) + using A diagonalE(2) apply blast + using diag_def_poly_eval_at_diagonal[of x] diag_def_poly_set_def[of n] + atLeastAtMost_iff[of _ 0 "n-1"] + by (metis (no_types, lifting) A image_iff lessThan_iff) + qed + show "affine_alg_set R (n + n) (diag_def_poly_set n) \ diagonal n" + proof fix x + assume A: "x \ affine_alg_set R (n + n) (diag_def_poly_set n)" + show "x \ diagonal n" + proof(rule diagonalI) + show "x = take n x @ take n x" + proof- + have 0: "x = take n x @ drop n x" + by (metis append_take_drop_id) + have "take n x = drop n x" + proof- + have 0: "length x = n + n" + using A unfolding affine_alg_set_def + using cartesian_power_car_memE by blast + then have 1: "length (take n x) = length (drop n x)" + using A + by (metis (no_types, lifting) \x = take n x @ drop n x\ + add.commute add_right_cancel affine_alg_set_closed cartesian_power_car_memE + le_add1 length_append subsetD take_closed) + have "\i::nat. i < n \ (take n x)!i = (drop n x) ! i" + proof- fix i::nat assume A0: "i < n" + then have "i \ {.. (diag_def_poly_set n)" + using diag_def_poly_set_def by blast + then have "eval_at_point R x (diag_def_poly n i) = \" + using A affine_alg_set_memE by blast + then have "x!i = x!(n + i)" + using A0 diag_def_poly_eval[of i n x] + by (metis (no_types, lifting) A add.commute affine_alg_set_closed + cartesian_power_car_memE' nat_add_left_cancel_less R.r_right_minus_eq subsetD trans_less_add2) + then show "take n x ! i =drop n x ! i" + by (metis "0" A0 le_add1 nth_drop nth_take) + qed + then show ?thesis using 0 + by (metis "1" \x = take n x @ drop n x\ add_less_mono + length_append less_not_refl linorder_neqE_nat nth_equalityI) + qed + then show ?thesis + using 0 by metis + qed + show "take n x \ carrier (R\<^bsup>n\<^esup>)" + using A unfolding affine_alg_set_def + by (meson A affine_alg_set_closed le_add2 subset_eq take_closed) + qed + qed +qed + +lemma diagonal_is_algebraic: + shows "is_algebraic R (n + n) (diagonal n)" + apply(rule is_algebraicI[of "diag_def_poly_set n"]) + apply (simp add: diag_def_poly_set_finite) + using diag_def_poly_set_in_coord_ring apply blast + by (simp add: diagonal_as_affine_alg_set) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Tuples of Functions\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition is_function_tuple :: "('a, 'b) ring_scheme \ nat \ ('a list \ 'a) list \ bool" where +"is_function_tuple R n fs = (set fs \ carrier (R\<^bsup>n\<^esup>) \ carrier R)" + +lemma is_function_tupleI: + assumes "(set fs \ carrier (R\<^bsup>n\<^esup>) \ carrier R)" + shows "is_function_tuple R n fs " + by (simp add: assms is_function_tuple_def) + +lemma is_function_tuple_append: + assumes "is_function_tuple R n fs" + assumes "is_function_tuple R n gs" + shows "is_function_tuple R n (fs@gs)" + using assms is_function_tupleI set_append + by (simp add: is_function_tuple_def) + +lemma is_function_tuple_Cons: + assumes "is_function_tuple R n fs" + assumes "f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" + shows "is_function_tuple R n (f#fs)" + using assms is_function_tupleI + by (simp add: assms(2) is_function_tuple_def) + +lemma is_function_tuple_snoc: + assumes "is_function_tuple R n fs" + assumes "f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" + shows "is_function_tuple R n (fs@[f])" + apply(rule is_function_tupleI) + by (metis (no_types) Un_subset_iff append_Nil assms(1) assms(2) is_function_tuple_Cons is_function_tuple_def set_append) + +lemma is_function_tuple_list_update: + assumes "is_function_tuple R n fs" + assumes "f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" + assumes "i < n" + shows "is_function_tuple R n (fs[i := f])" + apply(rule is_function_tupleI) + by (metis assms(1) assms(2) is_function_tuple_def set_update_subsetI) + +definition function_tuple_eval :: "'b \ 'c \ ('d \ 'a) list \ 'd \ 'a list" where +"function_tuple_eval R n fs x = map (\f. f x) fs" + +lemma function_tuple_eval_closed: + assumes "is_function_tuple R n fs" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "function_tuple_eval R n fs x \ carrier (R\<^bsup>length fs\<^esup>)" + apply(rule cartesian_power_car_memI') + apply (metis function_tuple_eval_def length_map) +proof- fix i assume "i < length fs" + then show "function_tuple_eval R n fs x ! i \ carrier R" + unfolding function_tuple_eval_def using assms unfolding is_function_tuple_def + by (metis funcset_carrier nth_map nth_mem subsetD) +qed + +definition coord_fun :: + "('a, 'c) ring_scheme \ nat \ ('a list \ 'b list) \ nat \ 'a list \ 'b" where +"coord_fun R n g i = (\x \ carrier (R\<^bsup>n\<^esup>). (g x) ! i)" + +lemma(in cring) map_is_coord_fun_tuple: + assumes "g \ carrier (R\<^bsup>n\<^esup>) \\<^sub>E carrier (R\<^bsup>m\<^esup>)" + shows "g = (\ x \ carrier (R\<^bsup>n\<^esup>). function_tuple_eval R n (map (coord_fun R n g) [0..n\<^esup>)) x" + proof(cases "x \ carrier (R\<^bsup>n\<^esup>)") + case True + then have T0: "restrict (function_tuple_eval R n (map (coord_fun R n g) [0..n\<^esup>)) x = + function_tuple_eval R n (map (coord_fun R n g) [0..i. i < m \ (g x) ! i = (function_tuple_eval R n (map (coord_fun R n g) [0.. ('a \ 'd) list \ ('d list \ 'b) \ 'a \ 'b" where +"function_tuple_comp R fs f = f \ (function_tuple_eval R (0::nat) fs)" + +lemma function_tuple_comp_closed: + assumes "f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" + assumes "length fs = n" + assumes "is_function_tuple R m fs" + shows "function_tuple_comp R fs f \ carrier (R\<^bsup>m\<^esup>) \ carrier R" + unfolding function_tuple_comp_def + using assms + by (smt Pi_iff comp_apply function_tuple_eval_closed function_tuple_eval_def) + +fun id_function_tuple where +"id_function_tuple (R::('a,'b) partial_object_scheme) 0 = []"| +"id_function_tuple R (Suc n) = id_function_tuple R n @ [(\(x::'a list). x! n)] " + +lemma id_function_tuple_is_function_tuple: +"\k. k \ n \ is_function_tuple R k (id_function_tuple R n)" + apply(induction n) + apply (simp add: is_function_tupleI) +proof- fix n k + assume IH: "(\k. n \ k \ is_function_tuple R k (id_function_tuple R n))" + + assume A: "Suc n \ k" + have 0: "(\a. a!n) \ carrier (R\<^bsup>k\<^esup>) \ carrier R" + using A cartesian_power_car_memE' + by (metis Pi_I Suc_le_lessD) + have 1: " is_function_tuple R k (id_function_tuple R n)" + using A IH Suc_leD by blast + then show "is_function_tuple R k (id_function_tuple R (Suc n))" + using A 0 id_function_tuple.simps(2)[of R n] + is_function_tuple_snoc[of R k "id_function_tuple R n" "\a. a!n" ] + by (simp add: "0") +qed + +lemma id_function_tuple_is_function_tuple': +"is_function_tuple R n (id_function_tuple R n)" +by (simp add: id_function_tuple_is_function_tuple) + +lemma id_function_tuple_eval_is_take: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "k \ n \ function_tuple_eval R n (id_function_tuple R k) a = take k a" + apply(induction k) + using assms + apply (simp add: assms function_tuple_eval_def) +proof- fix k + assume IH: "(k \ n \ function_tuple_eval R n (id_function_tuple R k) a = take k a) " + assume A: "Suc k \ n" + then have 0: "function_tuple_eval R n (id_function_tuple R k) a = take k a " + using IH Suc_leD + by blast + have "function_tuple_eval R n (id_function_tuple R (Suc k)) a + = function_tuple_eval R n (id_function_tuple R k) a @ [a!k]" + using id_function_tuple.simps(2)[of R k] + by (simp add: function_tuple_eval_def) + then show "function_tuple_eval R n (id_function_tuple R (Suc k)) a = take (Suc k) a" + by (metis "0" A Suc_le_lessD assms cartesian_power_car_memE take_Suc_conv_app_nth) +qed + +lemma id_function_tuple_eval_is_id: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "function_tuple_eval R n (id_function_tuple R n) a = a" + using assms id_function_tuple_eval_is_take[of a R n n] + by (metis cartesian_power_car_memE order_refl take_all) + +text\Composing a function tuple with a polynomial\ + +definition poly_function_tuple_comp :: + "('a, 'b) ring_scheme \ nat \ ('a list \ 'a) list \ ('a, nat) mvar_poly \ 'a list \ 'a" where +"poly_function_tuple_comp R n fs f = eval_at_poly R f \ function_tuple_eval R n fs" + +context cring_coord_rings +begin + +lemma poly_function_tuple_comp_closed: + assumes "is_function_tuple R n fs" + assumes "f \ carrier (coord_ring R (length fs))" + shows "poly_function_tuple_comp R n fs f \ carrier (R\<^bsup>n\<^esup>) \ carrier R" +proof fix x assume A: "x \ carrier (R\<^bsup>n\<^esup>)" + then show "poly_function_tuple_comp R n fs f x \ carrier R" + using assms function_tuple_eval_closed eval_at_point_closed + unfolding poly_function_tuple_comp_def + by (metis comp_apply) +qed + +lemma poly_function_tuple_comp_eq: + assumes "is_function_tuple R n fs" + assumes "f \ carrier (coord_ring R (length fs))" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_function_tuple_comp R n fs f a = eval_at_poly R f ( function_tuple_eval R n fs a)" + unfolding poly_function_tuple_comp_def + using comp_apply + by metis + +lemma poly_function_tuple_comp_constant: + assumes "is_function_tuple R n fs" + assumes "a \ carrier R" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_function_tuple_comp R n fs (coord_const a) x = a" + unfolding poly_function_tuple_comp_def + using assms comp_apply function_tuple_eval_closed + by (metis eval_at_point_const) + +lemma poly_function_tuple_comp_add: + assumes "is_function_tuple R n fs" + assumes "k \length fs" + assumes "p \ carrier (coord_ring R k)" + assumes "Q \ carrier (coord_ring R k)" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_function_tuple_comp R n fs (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) x = + (poly_function_tuple_comp R n fs p x) \ (poly_function_tuple_comp R n fs Q x)" +proof- + have 0: "p \ carrier (coord_ring R (length fs))" + using assms poly_ring_car_mono[of k "length fs"] + by blast + have 1: "Q \ carrier (coord_ring R (length fs))" + using assms poly_ring_car_mono[of k "length fs"] + by blast + show ?thesis + using assms(1) assms(5) 0 1 R.Pring_add_eq[of ] + poly_function_tuple_comp_eq + function_tuple_eval_closed[of R n fs x] + eval_at_point_add[of "function_tuple_eval R n fs x" "length fs" p Q] + unfolding coord_ring_def by (metis R.Pring_add_closed) +qed + +lemma poly_function_tuple_comp_mult: + assumes "is_function_tuple R n fs" + assumes "k \length fs" + assumes "p \ carrier (coord_ring R k)" + assumes "Q \ carrier (coord_ring R k)" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_function_tuple_comp R n fs (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q) x = + (poly_function_tuple_comp R n fs p x) \ (poly_function_tuple_comp R n fs Q x)" +proof- + have 0: "p \ carrier (coord_ring R (length fs))" + using assms poly_ring_car_mono[of k "length fs"] + by blast + have 1: "Q \ carrier (coord_ring R (length fs))" + using assms poly_ring_car_mono[of k "length fs"] + by blast + show ?thesis + using assms 0 1 + poly_function_tuple_comp_eq + function_tuple_eval_closed[of R n fs x] + eval_at_point_mult[of "function_tuple_eval R n fs x" "length fs" p Q] + unfolding coord_ring_def + by (metis MP.m_closed R.Pring_mult_eq coord_ring_def) +qed + +lemma poly_function_tuple_comp_pvar: + assumes "is_function_tuple R n fs" + assumes "k < length fs" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + shows "poly_function_tuple_comp R n fs (pvar R k) x = (fs ! k) x" +proof- + have "(map (\f. f x) fs) \ carrier (R\<^bsup>length fs\<^esup>)" + using function_tuple_eval_closed[of R n fs x] + unfolding function_tuple_eval_def + using assms(1) assms(3) by blast + then have "eval_at_poly R (pvar R k) (map (\f. f x) fs) = (fs! k) x" + using eval_pvar[of k "length fs" "(map (\f. f x) fs)"] + by (metis assms(2) nth_map) + then show ?thesis + by (metis (mono_tags, lifting) assms(1) assms(2) assms(3) function_tuple_eval_def + nth_map poly_function_tuple_comp_eq pvar_closed) +qed + +end +text\The coordinate ring of polynomials indexed by natural numbers\ + +definition Coord_ring :: "('a, 'b) ring_scheme \ ('a, ('a, nat) mvar_poly) module" where +"Coord_ring R = Pring R (UNIV :: nat set)" + + +text\Some general closure lemmas for coordinate rings\ +context cring_coord_rings +begin +lemma coord_ring_monom_term_closed: + assumes "a \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "b \ carrier (R[\\<^bsub>n\<^esub>])" + shows "a \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> b[^]\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub>(n::nat) \ carrier (R[\\<^bsub>n\<^esub>])" + using assms monoid.nat_pow_closed[of "(R[\\<^bsub>n\<^esub>])"] + unfolding coord_ring_def + by (meson R.Pring_is_monoid monoid.m_closed) + +lemma coord_ring_monom_term_plus_closed: + assumes "a \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "b \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "c \ carrier (R[\\<^bsub>n\<^esub>])" + shows "c \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> a \\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub> b[^]\<^bsub>(R[\\<^bsub>n\<^esub>])\<^esub>(n::nat) \ carrier (R[\\<^bsub>n\<^esub>])" + using assms coord_ring_monom_term_closed R.Pring_add_closed + by blast + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Generic Univariate Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + By a generic univariate polynomial, we mean a polynomial in one variable whose coefficients are + coordinate functions over a ring. That is, a polynomial of the form: + \[f(t) = x_0 + x_1t + \dots + x_nt^n\] + Such a polynomial can be construed as an element of $R[x_0,..,x_n](t)$, or as an element of + $R[x_0,..,x_n, x_n{n+1}]$. We will intially define the latter version, and show that it can + easily be cast to the former using the function ``\texttt{IP\_to\_UP"}. Such a polynomial can be + cast to a univariate polynomial over the ring $R$ by substituting a tuple of ring elements for + the coefficients. +\ +definition generic_poly_lt :: "('a, 'b) ring_scheme \ nat \ ('a, nat) mvar_poly" where +"generic_poly_lt R n = (pvar R (Suc n)) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> (pvar R 0)[^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n" + +fun generic_poly where +"generic_poly R (0::nat) = pvar R 1"| +"generic_poly R (Suc n) = (generic_poly R n) \\<^bsub>(coord_ring R (n+3))\<^esub> generic_poly_lt R (Suc n)" + + +context cring_coord_rings +begin + +lemma generic_poly_lt_closed: +"generic_poly_lt R n \ carrier (coord_ring R (Suc (Suc n)))" +proof- + have 0: "(pvar R (Suc n)) \ carrier (coord_ring R (Suc (Suc n)))" + using pvar_closed + by blast + have 1: " (pvar R 0) \ carrier (coord_ring R (Suc (Suc n)))" + using pvar_closed + by blast + then have "(pvar R 0)[^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n \ carrier (coord_ring R (Suc (Suc n)))" + using monoid.nat_pow_closed + unfolding coord_ring_def by (metis R.Pring_is_monoid) + then show ?thesis using 0 + unfolding coord_ring_def + by (metis R.Pring_mult_closed coord_ring_def generic_poly_lt_def) +qed + +lemma generic_poly_lt_eval: + assumes "a \ carrier (R\<^bsup>n+2\<^esup>)" + shows "eval_at_point R a (generic_poly_lt R n) = a!(Suc n) \ (a!0)[^]n " +proof- + have "(pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) \ carrier (coord_ring R (n + 2))" + using monoid.nat_pow_closed pvar_closed unfolding coord_ring_def + by (metis R.Pring_is_monoid add_2_eq_Suc' zero_less_Suc) + then have "eval_at_point R a (generic_poly_lt R n) = + eval_at_poly R (pvar R (Suc n)) a \ eval_at_poly R (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) a" + unfolding generic_poly_lt_def + using assms pvar_closed[of "(Suc n)" "n + 2"] eval_at_point_mult[of a "n + 2" "pvar R (Suc n)" "(pvar R 0)[^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n"] + by (metis add_2_eq_Suc' lessI) + then show ?thesis using assms + by (metis add_2_eq_Suc' add_gr_0 eval_at_point_nat_pow eval_pvar lessI pvar_closed zero_less_numeral) +qed + +lemma generic_poly_closed: +"generic_poly R n \ carrier (coord_ring R (Suc (Suc n)))" + apply(induction n) + using pvar_closed[of 1 "Suc (Suc n)"] + apply (metis One_nat_def generic_poly.simps(1) lessI pvar_closed) +proof- + fix n assume IH: "generic_poly R n \ carrier (coord_ring R (Suc (Suc n)))" + have "generic_poly R n \ carrier (coord_ring R (Suc (Suc (Suc n))))" + using IH poly_ring_car_mono'[of "Suc (Suc n)"] + by blast + then show " generic_poly R (Suc n) \ carrier (coord_ring R (Suc (Suc (Suc n))))" + unfolding coord_ring_def + using generic_poly.simps[of R] generic_poly_lt_closed[of n] + by (metis MP.add.m_closed R.Pring_add_eq coord_ring_def generic_poly_lt_closed) +qed + +lemma generic_poly_closed': + assumes "k \n" + shows "generic_poly R k \ carrier (coord_ring R (Suc (Suc n)))" +by (meson Suc_le_mono assms generic_poly_closed poly_ring_car_mono subsetD) + +lemma generic_poly_eval_at_point: + assumes "a \ carrier (R\<^bsup>n+3\<^esup>)" + shows "eval_at_point R a (generic_poly R (Suc n)) = (eval_at_point R a (generic_poly R n)) \ + (a!(n + 2)) \ (a!0)[^](Suc n)" +proof- + have 0: "(generic_poly R n) \ carrier (coord_ring R (n + 3))" + using generic_poly_closed' + by (metis Suc3_eq_add_3 add.commute eq_imp_le le_SucI) + then show ?thesis + using generic_poly.simps(2) + generic_poly_closed'[of n "n + 3"] + generic_poly_lt_eval eval_at_point_add[of a "(n + 3)" "generic_poly R n"] + by (metis (no_types, lifting) add.left_commute add_2_eq_Suc' assms + generic_poly_lt_closed numeral_2_eq_2 numeral_3_eq_3 plus_1_eq_Suc) +qed + +end + +text \ + We can turn points in $R^{n+1}$ into univariate polynomials with the associated coefficients + via partial evaluation of the generic polynomials of degree $n$. \ + +definition ring_cfs_to_poly :: +"('a, 'b) ring_scheme \ nat \ 'a list \ ('a, nat) mvar_poly" where +"ring_cfs_to_poly R n as = coord_partial_eval R {1..\<^bsub>R\<^esub>#as) (generic_poly R n)" + +context cring_coord_rings +begin + +lemma ring_cfs_to_poly_closed: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "ring_cfs_to_poly R n as \ carrier (coord_ring R 1)" +proof- + have 0: "\ # as \ carrier (R\<^bsup>n+2\<^esup>)" + apply(rule cartesian_power_car_memI) + using assms + apply (metis add_2_eq_Suc' cartesian_power_car_memE length_Cons) + using assms + by (metis cartesian_power_car_memE'' insert_subset list.simps(15) R.zero_closed) + then have 1: "coord_partial_eval R {1.. # as) \ ring_hom (coord_ring R (n + 2)) (Pring R ({.. # as) \ ring_hom (coord_ring R (n + 2)) (coord_ring R 1)" + using 1 unfolding coord_ring_def + by presburger + then show ?thesis + unfolding ring_cfs_to_poly_def coord_ring_def + by (metis "0" Diff_subset \{.. + add_2_eq_Suc' coord_partial_eval_closed generic_poly_closed + le_numeral_extra(4) lessThan_minus_lessThan lessThan_subset_iff) +qed + +text\Variant which maps to the univariate polynomial ring\ + +definition ring_cfs_to_univ_poly :: "nat \ 'a list \ nat \ 'a" where +"ring_cfs_to_univ_poly n as = IP_to_UP (0::nat) (ring_cfs_to_poly R n as)" + +lemma ring_cfs_to_univ_poly_closed: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "ring_cfs_to_univ_poly n as \ carrier (UP R)" + unfolding ring_cfs_to_univ_poly_def apply(rule R.IP_to_UP_closed, rule R.is_cring) + using ring_cfs_to_poly_closed unfolding coord_ring_def + using assms by (metis One_nat_def lessThan_0 lessThan_Suc) + +lemma ring_cfs_to_poly_eq: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + assumes "k \n" + shows "ring_cfs_to_poly R k as = ring_cfs_to_poly R k (take (Suc k) as) " + unfolding ring_cfs_to_poly_def coord_partial_eval_def + apply(rule R.poly_eval_eval_function_eq[of "(point_to_eval_map R (\ # as))" "(point_to_eval_map R (\ # take (Suc k) as))" "{1.. # as))" + apply(rule R.closed_funI) + using assms cartesian_power_car_memE[of as R "Suc n"] + by (metis cartesian_power_car_memE'' nth_mem set_ConsD subset_code(1) R.zero_closed) + show "closed_fun R (\i. if i < length (\ # take (Suc k) as) then (\ # take (Suc k) as) ! i else \)" + apply(rule R.closed_funI) + using assms + by (metis cartesian_power_car_memE'' in_set_takeD nth_mem set_ConsD subset_code(1) R.zero_closed) + have 0: "length (\ # as) \ k + 2" + using assms + by (metis Suc_le_mono add_2_eq_Suc' cartesian_power_car_memE length_Cons) + have 1: "length (\ # take (Suc k) as) \k + 2" + using 0 + by (metis add_2_eq_Suc' assms(1) cartesian_power_car_memE + impossible_Cons length_Cons not_less_eq_eq take_closed) + show "restrict (point_to_eval_map R (\ # as)) {1.. # take (Suc k) as)) {1.. # as)) {1.. # take (Suc k) as)) {1.. {1.. # as)) {1..#as)!x" + unfolding restrict_def + by (metis "0" True atLeastLessThan_iff le_Suc_ex trans_less_add1) + have 11: "restrict (point_to_eval_map R (\ # take (Suc k) as)) {1.. # take (Suc k) as)!x" + unfolding restrict_def + by (metis "1" True atLeastLessThan_iff le_Suc_ex trans_less_add1) + have 2: "(\ # as) ! x = (\ # take (Suc k) as) ! x" + proof- + obtain l where l_def: "Suc l = x" + using True + by (metis One_nat_def Suc_le_D atLeastLessThan_iff) + have P1: "(\ # as) ! x = as ! l" + using 0 True l_def + by (meson nth_Cons_Suc) + have P0: "(\ # take (Suc k) as) ! x = (take (Suc k) as) ! l" + using 1 True l_def + by (meson nth_Cons_Suc) + have "l < Suc k" + using True l_def + by (metis Suc_1 Suc_eq_plus1 Suc_less_SucD add_Suc_right atLeastLessThan_iff) + then have "(\ # take (Suc k) as) ! x = as ! l" + using P0 + by (metis nth_take) + then show ?thesis + using P1 by metis + qed + then show ?thesis using 00 11 True + by presburger + next + case False + then show ?thesis + unfolding restrict_def + by presburger + qed + qed + show " generic_poly R k \ Pring_set R {.. carrier (R\<^bsup>Suc n\<^esup>)" + shows "coord_partial_eval R {1..\<^bsub>R\<^esub>#as) (generic_poly_lt R n) = + poly_scalar_mult R (as!n) ((pvar R 0)[^]\<^bsub>coord_ring R (n+2)\<^esub> n)" +proof- + have 0: "\ # as \ carrier (R\<^bsup>Suc (Suc n)\<^esup>)" + using assms cartesian_power_cons + by (metis Suc_eq_plus1 R.zero_closed) + have 1: "pvar R (Suc n) \ Pring_set R {..coord_ring R (Suc (Suc n))\<^esub> n \ Pring_set R {.. # as)) + (pvar R (Suc n) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + (poly_eval R {1.. # as)) (pvar R (Suc n))) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> + (poly_eval R {1.. # as)) + (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n))" + using 0 1 2 R.poly_eval_mult[of "pvar R (Suc n)" "{..coord_ring R (Suc (Suc n))\<^esub> n" + "(point_to_eval_map R (\ # as))" "{1.. # as)) + (pvar R (Suc n) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + (coord_const ((\ # as)! (Suc n))) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> + (poly_eval R {1.. # as)) + (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n))" + using 0 3 point_to_eval_map_closed[of "(\ # as)" "Suc (Suc n)"] + R.poly_eval_index[of "(point_to_eval_map R (\ # as))" "{1..coord_ring R (Suc (Suc n))\<^esub> n \ Pring_set R ({.. {..Pring R ({.. n \ carrier (Pring R ({..k::nat. (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> k) = (pvar R 0 [^]\<^bsub>Pring R ({..k)" + proof- fix k::nat show "(pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> k) = (pvar R 0 [^]\<^bsub>Pring R ({..k)" + apply(induction k) + using R.Pring_var_closed[of 0 "{..<(Suc (Suc n))}"] R.Pring_var_closed[of 0 "{..pvar R 0 [^]\<^bsub>Pring R ({.. n \ carrier (Pring R ({..) + qed + have 6: "(poly_eval R {1.. # as)) + (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n)) = (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n)" + using 5 0 point_to_eval_map_closed[of "(\ # as)" "Suc (Suc n)"] + R.poly_eval_trivial[of "(point_to_eval_map R (\ # as))" "pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n" "{.. # as)) + (pvar R (Suc n) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + (coord_const ((\ # as)! (Suc n))) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> + (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n)" + using 4 6 + by presburger + have 8: "(\ # as) ! Suc n = as! n" + by (meson nth_Cons_Suc) + have 88: "(\ # as) ! Suc n \ carrier R" + by (metis "8" assms cartesian_power_car_memE' less_Suc_eq) + have 9: "poly_eval R {1.. # as)) + (pvar R (Suc n) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + coord_const ((\ # as) ! Suc n) \\<^sub>p pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n " + using R.poly_scalar_mult_eq[of "(\ # as) ! Suc n" "pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n"] + unfolding coord_ring_def + by (metis (no_types, lifting) "7" R.Pring_mult coord_ring_def) + have 10: "poly_scalar_mult R ((\ # as) ! Suc n) (pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + coord_const ((\ # as) ! Suc n) \\<^sub>p pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n" + using 9 8 88 0 5 R.poly_scalar_mult_eq[of "(\ # as) ! Suc n" "pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n" "({..coord_ring R (Suc (Suc n))\<^esub> n) = + coord_const ((\ # as) ! Suc n) \\<^sub>p pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n" + using 10 8 + by metis + have 12: "poly_eval R {1.. # as)) + (pvar R (Suc n) \\<^bsub>coord_ring R (Suc (Suc n))\<^esub> pvar R 0 [^]\<^bsub>coord_ring R (Suc (Suc n))\<^esub> n) = + poly_scalar_mult R (as ! n) ((pvar R 0) [^]\<^bsub>coord_ring R (n + 2)\<^esub> n)" + using 11 9 + by (metis add_2_eq_Suc') + then show ?thesis + unfolding coord_partial_eval_def generic_poly_lt_def + by blast +qed + +lemma coord_partial_eval_generic_poly_lt': + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "coord_partial_eval R {1..\<^bsub>R\<^esub>#as) (generic_poly_lt R n) = + poly_scalar_mult R (as!n) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> n)" +proof- + have 0: "coord_partial_eval R {1..\<^bsub>R\<^esub>#as) (generic_poly_lt R n) = + poly_scalar_mult R (as!n) ((pvar R 0)[^]\<^bsub>coord_ring R (n+2)\<^esub> n)" + using assms coord_partial_eval_generic_poly_lt by blast + have 1: "\k::nat. (pvar R 0)[^]\<^bsub>coord_ring R (n+2)\<^esub> k = (pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> k" + proof- fix k::nat show "(pvar R 0)[^]\<^bsub>coord_ring R (n+2)\<^esub> k = (pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> k" + apply(induction k) + unfolding coord_ring_def + apply (metis Group.nat_pow_0 R.Pring_one_eq) + using nat_pow_def + by (metis R.Pring_mult_eq R.Pring_one add_2_eq_Suc') + qed + then show ?thesis + using "0" by presburger +qed + +lemma ring_cfs_to_poly_decomp: + assumes "as \ carrier (R\<^bsup>Suc (Suc n)\<^esup>)" + shows "ring_cfs_to_poly R (Suc n) as = ring_cfs_to_poly R n as \\<^bsub>coord_ring R 1\<^esub> + poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))" +proof- + have LHS: "ring_cfs_to_poly R (Suc n) as = + coord_partial_eval R {1.. # as) (generic_poly R n \\<^bsub>coord_ring R (Suc (Suc (Suc n)))\<^esub> generic_poly_lt R (Suc n))" + by (smt add_2_eq_Suc' add_Suc_right generic_poly.simps(2) numeral_2_eq_2 numeral_3_eq_3 ring_cfs_to_poly_def) + have LHS': "ring_cfs_to_poly R (Suc n) as = + coord_partial_eval R {1.. # as) (generic_poly R n) \\<^bsub>coord_ring R (Suc (Suc (Suc n)))\<^esub> + coord_partial_eval R {1.. # as) (generic_poly_lt R (Suc n))" + using coord_partial_eval_add[of as "Suc n"] + by (metis LHS add_2_eq_Suc' add_Suc_shift assms cartesian_power_cons + coord_partial_eval_add generic_poly_closed' generic_poly_lt_closed le_add2 plus_1_eq_Suc R.zero_closed) + have LHS'': "ring_cfs_to_poly R (Suc n) as = + coord_partial_eval R {1.. # as) (generic_poly R n) \\<^bsub>coord_ring R (Suc (Suc (Suc n)))\<^esub> + coord_partial_eval R {1.. # as) (generic_poly_lt R (Suc n))" + using coord_partial_eval_add[of as "Suc n"] + by (metis LHS add_2_eq_Suc' add_Suc_shift assms cartesian_power_cons + coord_partial_eval_add generic_poly_closed' generic_poly_lt_closed le_add2 plus_1_eq_Suc R.zero_closed) + have LHS''': "ring_cfs_to_poly R (Suc n) as = + coord_partial_eval R {1.. # as) (generic_poly R n) \\<^bsub>coord_ring R (Suc (Suc (Suc n)))\<^esub> + poly_scalar_mult R (as! (Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))" + using LHS'' coord_partial_eval_generic_poly_lt'[of as "Suc n"] assms + by presburger + have 0: "coord_partial_eval R {1.. # as) (generic_poly R n) = ring_cfs_to_poly R n as" + proof- + have 00: "(generic_poly R n) \ carrier (coord_ring R (n + 2))" + using add_2_eq_Suc' generic_poly_closed by presburger + have 01: "\ \ \" + using one_neq_zero + by presburger + have 02: "(\ # as) \ carrier (R\<^bsup>Suc (Suc n) + 1\<^esup>)" + using cartesian_power_cons[of as R "Suc (Suc n)" \] assms + by blast + have 03: "closed_fun R (point_to_eval_map R (\ # as))" + using point_to_eval_map_closed[of "\#as" "Suc (Suc (Suc n))"] + by (metis "02" Suc_eq_plus1) + have 04: "{1.. {.. {.. # as)" "{1.. carrier (R\<^bsup>Suc (Suc n)\<^esup>)" + shows "ring_cfs_to_poly R (Suc n) as = + ring_cfs_to_poly R n (take (Suc n) as) \\<^bsub>coord_ring R 1\<^esub> + poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))" + using assms ring_cfs_to_poly_decomp[of as n] + ring_cfs_to_poly_eq[of as "Suc n" n] le_eq_less_or_eq less_Suc_eq + by presburger + +lemma ring_cfs_to_univ_poly_decomp': + assumes "as \ carrier (R\<^bsup>Suc (Suc n)\<^esup>)" + shows "ring_cfs_to_univ_poly (Suc n) as = + ring_cfs_to_univ_poly n (take (Suc n) as) \\<^bsub>UP R\<^esub> + (as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))" +proof- + have 00: "(pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n) \ carrier (Pring R {0})" + using pvar_closed[of 0 1] monoid.nat_pow_closed[of "coord_ring R 1" _ n ] + unfolding coord_ring_def + by (metis One_nat_def R.Pring_is_monoid lessThan_0 lessThan_Suc less_one monoid.nat_pow_closed) + have LHS: "ring_cfs_to_univ_poly (Suc n) as = + IP_to_UP 0 (ring_cfs_to_poly R n (take (Suc n) as) \\<^bsub>coord_ring R 1\<^esub> + poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n)))" + using assms ring_cfs_to_poly_decomp' + unfolding ring_cfs_to_univ_poly_def + by presburger + have LHS': "ring_cfs_to_univ_poly (Suc n) as = + IP_to_UP 0 (ring_cfs_to_poly R n (take (Suc n) as)) \\<^bsub>UP R\<^esub> + IP_to_UP 0 (poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n)))" + proof- + have 0: " ring_cfs_to_poly R n (take (Suc n) as) \ carrier (Pring R {0})" + by (metis One_nat_def assms coord_ring_def le_add2 lessThan_0 lessThan_Suc plus_1_eq_Suc ring_cfs_to_poly_closed take_closed) + have 1: "as ! Suc n \ carrier R" + using assms cartesian_power_car_memE'[of as R "Suc (Suc n)"] + by blast + have 2: "poly_scalar_mult R (as ! Suc n) (pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n) \ carrier (Pring R {0})" + using 1 00 R.Pring_car R.poly_scalar_mult_closed[of "(as ! Suc n)" "(pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n)" "{0}"] + by blast + then show ?thesis + using 0 1 2 UP_cring.IP_to_UP_add[of R "(ring_cfs_to_poly R n (take (Suc n) as))" "0" + "poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))"] + by (metis LHS One_nat_def UP_cring_def coord_ring_def R.is_cring lessThan_0 lessThan_Suc) + qed + have 0: "IP_to_UP 0 (ring_cfs_to_poly R n (take (Suc n) as)) = + ring_cfs_to_univ_poly n (take (Suc n) as)" + using ring_cfs_to_univ_poly_def + by presburger + have 1: "(mset_to_IP R (nat_to_mset 0 (Suc n))) = (pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n)" + unfolding coord_ring_def using lessThan_iff less_one + by (metis UP_cring.intro UP_cring.pvar_pow R.is_cring) + have 2: "as ! Suc n \ carrier R" + using cartesian_power_car_memE' assms + by blast + have 3: "IP_to_UP 0 (poly_scalar_mult R (as ! Suc n) (pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n)) = + as ! Suc n \\<^bsub>UP R\<^esub> IP_to_UP 0 (pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n)" + using UP_cring.IP_to_UP_scalar_mult[of R "as!(Suc n)" "((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))" 0] + "00" "2" unfolding coord_ring_def + by (metis R.Pring_smult UP_cring.intro R.is_cring) + have 4: "IP_to_UP 0 (poly_scalar_mult R (as!(Suc n)) ((pvar R 0)[^]\<^bsub>coord_ring R 1\<^esub> (Suc n))) + = (as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))" + proof - + have "as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> Suc n = IP_to_UP (0::nat) (Mt (as ! Suc n) (nat_to_mset 0 (Suc n)))" + using 3 1 UP_cring.IP_to_UP_monom + by (metis UP_cring.intro R.is_cring) + then show ?thesis + using \mset_to_IP R (nat_to_mset 0 (Suc n)) = pvar R 0 [^]\<^bsub>coord_ring R 1\<^esub> Suc n\ + by presburger + qed + then show ?thesis + using "0" LHS' + by presburger +qed + +lemma ring_cfs_to_univ_poly_decomp: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + assumes "k < n" + shows "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k)) as) = ring_cfs_to_univ_poly k (take (Suc k) as) + \\<^bsub>UP R\<^esub> (as!(Suc k)) \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc k)" +proof- + have 0: "(take (Suc (Suc k)) as) \ carrier (R\<^bsup>Suc (Suc k)\<^esup>)" + using assms + by (meson Suc_leI Suc_mono take_closed) + then show ?thesis using ring_cfs_to_univ_poly_decomp'[of "take (Suc (Suc k)) as" k] + by (metis (no_types, lifting) Suc_leI assms(1) assms(2) cartesian_power_car_memE + lessI less_SucI nth_take nth_take_lemma) +qed + +lemma ring_cfs_to_univ_poly_degree: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "deg R (ring_cfs_to_univ_poly n as) \ n" + "as!n \ \ \ deg R (ring_cfs_to_univ_poly n as) = n" +proof- + have 0:"\as. as \ carrier (R\<^bsup>Suc n\<^esup>) \ + deg R (ring_cfs_to_univ_poly n as) \ n \ (as!n \ \ \ deg R (ring_cfs_to_univ_poly n as) = n)" + proof(induction n) + case 0 + show "\as. as \ carrier (R\<^bsup>Suc 0\<^esup>) \ + deg R (ring_cfs_to_univ_poly 0 as) \ 0 \ + (as ! 0 \ \ \ deg R (ring_cfs_to_univ_poly 0 as) = 0)" + proof- + fix as assume A: "as \ carrier (R\<^bsup>Suc 0\<^esup>)" + have 0:"cring R" + by (simp add: R.is_cring) + have 1:"\ # as \ carrier (R\<^bsup>2\<^esup>)" + using A cartesian_power_cons[of as R "Suc 0" \] + by (metis numeral_1_eq_Suc_0 numeral_One one_add_one R.zero_closed) + have 2: "(\ # as) ! 1 = as!0" + using A + by (metis One_nat_def nth_Cons_Suc) + have 3: "1 \ {(1::nat)..<0 + 2} \ {..<2}" + by auto + have 4: "coord_partial_eval R {1::nat..<0 + 2} (\ # as) (pvar R (1::nat)) = + R.indexed_const (as!0)" + unfolding ring_cfs_to_univ_poly_def ring_cfs_to_poly_def + using 0 1 2 one_neq_zero UP_cring.IP_to_UP_indexed_const[of R "as!0" 0] generic_poly.simps(1)[of R] coord_partial_eval_pvar[of "\#as" 2 "1::nat" "{1..<0+2}" ] + unfolding UP_cring_def + using "3" by presburger + have 5: "ring_cfs_to_univ_poly 0 as = IP_to_UP (0::nat) (R.indexed_const (as ! 0))" + unfolding ring_cfs_to_univ_poly_def ring_cfs_to_poly_def + using 4 generic_poly.simps(1)[of R] + by presburger + hence "ring_cfs_to_univ_poly 0 as = to_polynomial R (as!0)" + by (metis A UP_cring.IP_to_UP_indexed_const UP_cring.intro + cartesian_power_car_memE' R.is_cring lessI) + assume B: "as \ carrier (R\<^bsup>Suc 0\<^esup>) " + have 0: "(point_to_eval_map R (\ # as) 1) = as!0" + by (metis B One_nat_def cartesian_power_car_memE impossible_Cons le_numeral_extra(4) + linorder_neqE_nat nat_less_le nth_Cons_Suc) + have 1: "closed_fun R ((point_to_eval_map R (\ # as)))" + apply(rule R.closed_funI) + by (metis "0" B One_nat_def cartesian_power_car_memE cartesian_power_car_memE' + length_Suc_conv less_Suc0 less_SucE nth_Cons_0 R.zero_closed) + have 2: "(1::nat) \ ({1..<0 + 2}::nat set)" + by simp + have 3: "poly_eval R {1..<0 + 2} (point_to_eval_map R (\ # as)) (mset_to_IP R {#1#}) = + coord_const (point_to_eval_map R (\ # as) 1)" + using generic_poly.simps(1)[of R] one_neq_zero + unfolding ring_cfs_to_poly_def coord_partial_eval_def var_to_IP_def + using 0 1 2 R.poly_eval_index[of "(point_to_eval_map R (\ # as))" "{1..<0 + 2}" 1] + by (metis (no_types, lifting)) + have 4: "(ring_cfs_to_poly R 0 as) = coord_const (as! 0)" + using 3 0 generic_poly.simps(1)[of R] + unfolding ring_cfs_to_poly_def coord_partial_eval_def var_to_IP_def + by presburger + have 5: "as! 0 \ carrier R" + using assms B cartesian_power_car_memE' by blast + have 6: "(ring_cfs_to_univ_poly 0 as) = to_polynomial R (as! 0)" + unfolding ring_cfs_to_univ_poly_def ring_cfs_to_poly_def + using 3 4 5 UP_cring.IP_to_UP_indexed_const[of R "as!0" "0::nat"] + unfolding coord_partial_eval_def + by (smt "0" \ring_cfs_to_univ_poly 0 as = to_polynomial R (as ! 0)\ generic_poly.simps(1) ring_cfs_to_univ_poly_def var_to_IP_def) + then show " deg R (ring_cfs_to_univ_poly 0 as) \ 0 \ (as ! 0 \ \ \ deg R (ring_cfs_to_univ_poly 0 as) = 0)" + using UP_cring.degree_to_poly[of R "as! 0"] 5 UP_cring_def[of R] + using R.is_cring by presburger + qed + next + case (Suc n) + have 0: "(ring_cfs_to_univ_poly (Suc n) as) = ring_cfs_to_univ_poly n (take (Suc n) as) \\<^bsub>UP R\<^esub> + (as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))" + using ring_cfs_to_univ_poly_decomp' Suc.prems by blast + have 1: "(take (Suc n) as) \ carrier (R\<^bsup>Suc n\<^esup>)" + using Suc.prems + by (meson le_Suc_eq take_closed) + have 2: "deg R (ring_cfs_to_univ_poly n (take (Suc n) as)) \ n" + using "1" Suc.IH + by blast + have 3: "deg R ((as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))) \ Suc n" + using UP_cring.degree_monom[of R "as!(Suc n)" "Suc n"] UP_cring_def[of R] + Suc.prems cartesian_power_car_memE' le_Suc_eq lessI less_imp_le_nat zero_less_Suc + by (metis R.is_cring) + have 4: "(X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) \ carrier (UP R)" + proof- + have 40: "Group.monoid (UP R)" + using UP_cring_def[of R] UP_domain_def cring.axioms(1) ring.is_monoid + using UP_cring.UP_cring R.is_cring by blast + have 41: "X_poly R \ carrier (UP R)" + using UP_cring.X_closed[of R] UP_cring_def[of R] R.is_cring + by blast + show ?thesis + using monoid.nat_pow_closed[of "UP R" "X_poly R" "Suc n"] 40 41 + by blast + qed + have 5: "deg R (ring_cfs_to_univ_poly (Suc n) as) \Suc n" + proof(cases "as!(Suc n) = \") + case True + then have T0: "(as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) = \\<^bsub>UP R\<^esub>" + using 4 UP_ring.UP_smult_zero[of R "X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)"] UP_ring_def[of R] R.ring_axioms + by presburger + then show ?thesis + using UP_ring.deg_zero[of R] UP_ring_def[of R] + by (metis "0" "1" "2" "3" UP_ring.UP_zero_closed UP_ring.bound_deg_sum le_SucI R.ring_axioms ring_cfs_to_univ_poly_closed) + next + case False + have F0 : "as!(Suc n) \ carrier R" + by (metis Suc.prems cartesian_power_car_memE le_simps(1) lessI not_less_eq_eq poly_tuple_evalE poly_tuple_evalE' pushforward_by_pvar_list pvar_list_is_poly_tuple zero_less_Suc) + have F1: "(as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) \ carrier (UP R)" + using F0 4 UP_ring.UP_smult_closed[of R "as!(Suc n)" "X_poly R [^]\<^bsub>UP R\<^esub> Suc n "] + UP_ring_def[of R] assms R.ring_axioms + by blast + have "deg R ((as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))) = Suc n" + using False UP_cring.degree_monom[of R "as!(Suc n)" "Suc n"] UP_cring_def[of R] + cartesian_power_car_memE' lessI + using F0 R.is_cring + by presburger + then show ?thesis + using UP_ring.degree_of_sum_diff_degree[of R "(as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))" + "ring_cfs_to_univ_poly n (take (Suc n) as)"] 1 2 4 UP_domain_def[of R] F1 + ring_cfs_to_univ_poly_closed[of "take (Suc n) as" "Suc n"] "0" "3" + UP_ring_def[of R] UP_cring_def[of R] + UP_ring.equal_deg_sum less_Suc_eq_le ring_cfs_to_univ_poly_closed + by (metis R.ring_axioms) + qed + have 6: "(as ! (Suc n) \ \ \ deg R (ring_cfs_to_univ_poly (Suc n) as) = Suc n)" + proof + assume F: "as ! (Suc n) \ \ " + have F0 : "as!(Suc n) \ carrier R" + by (metis Suc.prems cartesian_power_car_memE le_simps(1) lessI not_less_eq_eq poly_tuple_evalE poly_tuple_evalE' pushforward_by_pvar_list pvar_list_is_poly_tuple zero_less_Suc) + have F1: "(as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) \ carrier (UP R)" + using F0 4 UP_ring.UP_smult_closed[of R "as!(Suc n)" "X_poly R [^]\<^bsub>UP R\<^esub> Suc n "] + UP_ring_def[of R] assms R.ring_axioms + by blast + then have F2: "deg R ((as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc n))) = Suc n" + using F0 F UP_cring.degree_monom[of R "as!(Suc n)" "Suc n"] UP_cring_def[of R] R.is_cring + by presburger + have F3: "ring_cfs_to_univ_poly n (take (Suc n) as) \ carrier (UP R)" + using "1" ring_cfs_to_univ_poly_closed + by blast + show "deg R (ring_cfs_to_univ_poly (Suc n) as) = Suc n" + using UP_domain_def[of R] 0 F1 F2 F3 1 2 + UP_ring.degree_of_sum_diff_degree[of R "ring_cfs_to_univ_poly n (take (Suc n) as)" + "as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> Suc n"] + UP_ring.equal_deg_sum le_imp_less_Suc UP_ring_def[of R] UP_cring_def[of R] + by (metis R.ring_axioms) + qed + show ?case + using "5" "6" by blast + qed + show "deg R (ring_cfs_to_univ_poly n as) \ n" + using 0 assms + by blast + show "as ! n \ \ \ deg R (ring_cfs_to_univ_poly n as) = n" + using 0 assms + by blast +qed + +lemma ring_cfs_to_univ_poly_constant: + assumes "as \ carrier (R\<^bsup>1\<^esup>)" + shows "ring_cfs_to_univ_poly 0 as = to_polynomial R (as!0)" +proof- + have 0: "(1::nat) \ {1..<0 + 2}" + by simp + have 1: "closed_fun R (point_to_eval_map R (\ # as))" + using assms + by (smt cartesian_power_car_memE'' R.closed_funI nth_mem set_ConsD subset_code(1) R.zero_closed) + have 2: "(point_to_eval_map R (\ # as) (1::nat)) = as!0" + by (metis One_nat_def assms cartesian_power_car_memE impossible_Cons + le_numeral_extra(4) linorder_neqE_nat nat_less_le nth_Cons_Suc) + have 3: "as!0 \ carrier R" + using assms cartesian_power_car_memE' + by blast + have "(poly_eval R {1::nat..<0 + 2} (point_to_eval_map R (\ # as)) (generic_poly R 0)) = coord_const (point_to_eval_map R (\ # as) 1)" + using generic_poly.simps(1)[of R] 0 1 one_not_zero + cring.poly_eval_index[of R "point_to_eval_map R (\ # as)" "{1..<0 + 2}" 1] + unfolding var_to_IP_def + using R.is_cring local.one_neq_zero by presburger + then have "(poly_eval R {1..<0 + 2} (point_to_eval_map R (\ # as)) (generic_poly R 0)) = coord_const (as!0)" + using 2 + by presburger + then show ?thesis + using 3 + unfolding ring_cfs_to_univ_poly_def ring_cfs_to_poly_def coord_partial_eval_def + by (metis UP_cring.IP_to_UP_indexed_const UP_cring.intro R.is_cring) +qed + +lemma ring_cfs_to_univ_poly_top_coeff: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "(ring_cfs_to_univ_poly n as) n = as ! n" + proof(cases "n = 0") + case True + have 0: "as ! 0 \ carrier R" + using assms cartesian_power_car_memE' + by blast + have 1: "to_polynomial R (as ! 0) 0 = as ! 0" + using assms cartesian_power_car_memE'[of as R "Suc n"] UP_ring.cfs_monom[of R] + unfolding to_polynomial_def UP_ring_def + using "0" R.ring_axioms by presburger + have "ring_cfs_to_univ_poly 0 as = to_polynomial R (as ! 0)" + using One_nat_def True assms ring_cfs_to_univ_poly_constant by presburger + then show ?thesis + using True 1 + by presburger + next + case False + obtain k where k_def: "Suc k = n" + using False + by (metis lessI less_Suc_eq_0_disj) + have "ring_cfs_to_univ_poly (Suc k) as (Suc k) = as ! (Suc k)" + proof- + have 0: "ring_cfs_to_univ_poly (Suc k) as n = ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k))as) n" + by (metis assms(1) k_def le_Suc_eq ring_cfs_to_poly_eq ring_cfs_to_univ_poly_def) + have 1: "take (Suc (Suc k)) as \ carrier (R\<^bsup>Suc (Suc k)\<^esup>)" + using assms k_def take_closed + by blast + have 2: "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k))as) = + ring_cfs_to_univ_poly k (take (Suc k) (take (Suc (Suc k)) as)) \\<^bsub>UP R\<^esub> (as!(Suc k))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc k))" + using 1 ring_cfs_to_univ_poly_decomp'[of "take (Suc (Suc k))as" k] assms + by (metis cartesian_power_car_memE k_def nat_le_linear take_all) + have 3: "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k))as) = + ring_cfs_to_univ_poly k (take (Suc k) as) \\<^bsub>UP R\<^esub> (as!(Suc k))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc k))" + using 2 + by (metis assms(1) k_def le_Suc_eq ring_cfs_to_poly_eq ring_cfs_to_univ_poly_decomp' ring_cfs_to_univ_poly_def) + have 4: "deg R (ring_cfs_to_univ_poly k (take (Suc k) as)) \ k" + by (metis assms(1) dual_order.refl k_def le_SucI ring_cfs_to_univ_poly_degree(1) take_closed) + have 5: "(ring_cfs_to_univ_poly k (take (Suc k) as)) \ carrier (UP R)" + by (metis assms(1) k_def le_Suc_eq le_refl ring_cfs_to_univ_poly_closed take_closed) + have 6: "X_poly R [^]\<^bsub>UP R\<^esub> Suc k \ carrier (UP R)" + using monoid.nat_pow_closed[of "UP R" "X_poly R" "Suc k"] domain_def ring.is_monoid[of "UP R"] + UP_cring.X_closed[of R] UP_domain_def[of R] UP_cring_def[of R] + cring.axioms(1) UP_cring.UP_cring + using R.is_cring by blast + have 7: " (as!(Suc k))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc k)) \ carrier (UP R)" + using UP_ring.UP_smult_closed[of R "as!(Suc k)" " (X_poly R [^]\<^bsub>UP R\<^esub> (Suc k))"] + UP_ring_def[of R] domain_def 6 cartesian_power_car_memE'[of as R _ "Suc k"] + assms(1) k_def R.ring_axioms by blast + have 8: "ring_cfs_to_univ_poly (Suc k) as (Suc k) = ( (as!(Suc k))\\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> (Suc k))) (Suc k)" + using 3 4 k_def + "5" 7 UP_cring_def[of R] UP_ring_def[of R] add.r_cancel_one' assms(1) + cartesian_power_car_memE le_eq_less_or_eq + le_imp_less_Suc take_all R.zero_closed UP_ring.UP_a_comm UP_ring.coeff_of_sum_diff_degree0 R.ring_axioms + by (metis (no_types, lifting)) + then show ?thesis using UP_cring_def[of R] UP_cring.monom_coeff assms(1) cartesian_power_car_memE + k_def lessI point_to_eval_map_closed + by (metis (no_types, lifting) cartesian_power_car_memE' R.is_cring) + qed + then show ?thesis + using k_def False + by blast + qed + +lemma(in UP_cring) monom_plus_lower_degree_top_coeff: + assumes "degree p < n" + assumes "p \ carrier (UP R)" + assumes "a \ carrier R" + shows "(p \\<^bsub>UP R\<^esub> (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> n)) n = a" +proof- + have 0: "(a \\<^bsub>UP R\<^esub> (X_poly R [^]\<^bsub>UP R\<^esub> n)) \ carrier (UP R)" + using P.nat_pow_closed P_def X_closed assms(3) smult_closed + by blast + have 1: "( (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> n) \\<^bsub>UP R\<^esub> p) n = (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> n) n" + using "0" UP_ring.coeff_of_sum_diff_degree0[of R] UP_cring_def[of R] assms(1) assms(2) + using is_UP_ring by blast + then show ?thesis + using 0 assms P_def UP_a_comm UP_cring.monom_coeff UP_cring_def[of R] + by (metis R_cring) +qed + +lemma(in UP_cring) monom_closed: + assumes "a \ carrier R" + shows "a \\<^bsub>UP R\<^esub> ((X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat)) \ carrier (UP R)" + using P.nat_pow_closed P_def assms X_closed carrier_is_submodule submoduleE(4) + by blast + +lemma(in UP_cring) monom_bottom_coeff: + assumes "a \ carrier R" + assumes "n > 0" + shows "(a \\<^bsub>UP R\<^esub> ((X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat))) 0 = \" + using assms monom_coeff[of a n] P_def local.monom_coeff + by presburger + +lemma(in UP_cring) monom_plus_lower_degree_bottom_coeff: + assumes "0 < n" + assumes "p \ carrier (UP R)" + assumes "a \ carrier R" + shows "(p \\<^bsub>UP R\<^esub> (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat))) 0 = p 0" +proof- + have 0: "p 0 \ carrier R" + using assms(2) UP_ring_def is_UP_ring P_def cfs_closed by blast + have 1: "(p \\<^bsub>UP R\<^esub> (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat))) 0 = p 0 \ (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> n) 0" + using assms monom_closed[of a n] cfs_add[of p "(a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat))" 0] + unfolding P_def + by blast + then have "(a \\<^bsub>UP R\<^esub> ((X_poly R) [^]\<^bsub>UP R\<^esub> n)) 0 = \" + using monom_bottom_coeff[of a n] P_def assms(1) assms(3) local.monom_coeff + by blast + then have 2: "(p \\<^bsub>UP R\<^esub> (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (n::nat))) 0 = p 0 \ \" + using 1 by metis + then show ?thesis + using 0 R.add.l_cancel_one[of "p 0"] R.zero_closed + by presburger +qed + +lemma ring_cfs_to_univ_poly_bottom_coeff: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "(ring_cfs_to_univ_poly n as) 0 = as ! 0" +proof- + have "\as. as \ carrier (R\<^bsup>Suc n\<^esup>) \ (ring_cfs_to_univ_poly n as) 0 = as ! 0" + apply(induction n) + using ring_cfs_to_univ_poly_top_coeff apply blast + proof- + fix n as + assume IH: "\as. as \ carrier (R\<^bsup>Suc n\<^esup>) \ (ring_cfs_to_univ_poly n as) 0 = as ! 0" + assume A: "as \ carrier (R\<^bsup>Suc (Suc n)\<^esup>)" + show "ring_cfs_to_univ_poly (Suc n) as 0 = as ! 0" + proof- + have 0: "ring_cfs_to_univ_poly (Suc n) as = ring_cfs_to_univ_poly n (take (Suc n) as) \\<^bsub>UP R\<^esub> (as!(Suc n))\\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc n)" + using A ring_cfs_to_univ_poly_decomp'[of as n] + by blast + have 1: "ring_cfs_to_univ_poly n (take (Suc n) as) \ carrier (UP R)" + by (meson A ring_cfs_to_univ_poly_closed R.is_cring le_Suc_eq take_closed) + have 2:"as ! Suc n \ carrier R" + using assms cartesian_power_car_memE' A + by blast + have 3: "(ring_cfs_to_univ_poly n (take (Suc n) as) \\<^bsub>UP R\<^esub> as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) 0 = + ring_cfs_to_univ_poly n (take (Suc n) as) 0" + proof- + have "as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> Suc n \ carrier (UP R)" + by (meson "2" UP_cring.monom_closed UP_cring_def R.is_cring) + hence 30: "(ring_cfs_to_univ_poly n (take (Suc n) as) \\<^bsub>UP R\<^esub> as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) 0 = + (ring_cfs_to_univ_poly n (take (Suc n) as)) 0 \ (as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)) 0" + using A ring_cfs_to_univ_poly_closed[of "take (Suc n) as" "n"] take_closed[of "Suc n" "Suc (Suc n)" as R] + UP_ring.cfs_add[of R "ring_cfs_to_univ_poly n (take (Suc n) as)" "as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> (Suc n)" 0] + unfolding UP_ring_def + using "1" R.ring_axioms by blast + have 31: "(as ! Suc n \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> Suc n) 0 = \" + by (metis (no_types, lifting) "2" Suc_neq_Zero UP_cring.monom_coeff UP_cring_def R.is_cring) + thus ?thesis using 30 2 + by (simp add: "1" UP_car_memE(1)) + qed + have 4: "(take (Suc n) as) \ carrier (R\<^bsup>Suc n\<^esup>)" + by (meson A le_Suc_eq take_closed) + have 5: "ring_cfs_to_univ_poly n (take (Suc n) as) 0 = as!0" + using IH[of "(take (Suc n) as)"] 4 nth_take[of 0 "Suc n" as] less_Suc_eq_0_disj + by presburger + then show ?thesis + using 0 3 + by presburger + qed + qed + then show ?thesis + using assms + by blast +qed + +lemma ring_cfs_to_univ_poly_chain: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + assumes "l \ n" + shows "l \ k \ k \ n \ (ring_cfs_to_univ_poly k (take (Suc k) as)) l = (ring_cfs_to_univ_poly l (take (Suc l) as)) l" + apply( induction k) + apply blast +proof- + fix k + assume IH: "(l \ k \ k \ n \ ring_cfs_to_univ_poly k (take (Suc k) as) l = ring_cfs_to_univ_poly l (take (Suc l) as) l)" + assume A: "l \ Suc k \ Suc k \ n" + show "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k)) as) l = ring_cfs_to_univ_poly l (take (Suc l) as) l" + proof(cases "l = Suc k") + case True + then show ?thesis + by blast + next + case False + then have "l \ k \ k \ n " + using A le_Suc_eq + by blast + then have 0: " ring_cfs_to_univ_poly k (take (Suc k) as) l = ring_cfs_to_univ_poly l (take (Suc l) as) l" + using IH + by blast + have 1: "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k)) as) = ring_cfs_to_univ_poly k (take (Suc k) as) + \\<^bsub>UP R\<^esub> (as!(Suc k)) \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc k)" + using assms A ring_cfs_to_univ_poly_decomp[of as n k] Suc_le_lessD + by blast + have 2: "ring_cfs_to_univ_poly (Suc k) (take (Suc (Suc k)) as) l = ring_cfs_to_univ_poly k (take (Suc k) as) l + \( (as!(Suc k)) \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc k)) l" + proof- + have 21: "ring_cfs_to_univ_poly k (take (Suc k) as) \ carrier (UP R)" + by (meson A assms(1) le_SucI ring_cfs_to_univ_poly_closed take_closed) + have 22: "as ! Suc k \\<^bsub>UP R\<^esub> X_poly R [^]\<^bsub>UP R\<^esub> Suc k \ carrier (UP R)" + using UP_ring_def[of R] A UP_ring.monom_closed assms(1) cartesian_power_car_memE' less_Suc_eq_le + monoid.nat_pow_closed[of "UP R" "X_poly R" "Suc k"] + unfolding X_poly_def + by (metis UP_ring.UP_ring UP_ring.UP_smult_closed R.ring_axioms R.one_closed ring.is_monoid) + show ?thesis + using 1 21 22 UP_ring.cfs_add[of R "ring_cfs_to_univ_poly k (take (Suc k) as)" "( (as!(Suc k)) \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc k))" l] + UP_ring_def[of R] R.ring_axioms by presburger + qed + have 3: "( (as!(Suc k)) \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub> (Suc k)) l = \" + using UP_cring.monom_coeff[of R "as!(Suc k)"] A False UP_cring_def assms(1) cartesian_power_car_memE' + by (metis R.is_cring le_imp_less_Suc) + then show ?thesis + using 2 + by (metis "0" Suc_le_mono assms(1) assms(2) cartesian_power_car_memE' lessI R.r_zero + ring_cfs_to_univ_poly_top_coeff take_closed) + qed +qed + +lemma ring_cfs_to_univ_poly_coeffs: + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + assumes "l \ n" + shows "(ring_cfs_to_univ_poly n as) l = (ring_cfs_to_univ_poly l (take (Suc l) as)) l" +proof- + have "(take (Suc n) as) = as" + using assms + by (metis cartesian_power_car_memE le_refl take_all) + then show ?thesis + using ring_cfs_to_univ_poly_chain[of as n l n] + by (metis assms(1) assms(2) order_refl) +qed + +lemma ring_cfs_to_univ_poly_coeffs': + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + assumes "l \ n" + shows "(ring_cfs_to_univ_poly n as) l = as! l" +proof- + have 0: "(ring_cfs_to_univ_poly l (take (Suc l) as)) l = (take (Suc l) as) ! l" + by (meson Suc_le_mono assms(1) assms(2) ring_cfs_to_univ_poly_top_coeff take_closed) + have 1: "(take (Suc l) as) ! l = as! l" + using nth_take[of l "Suc l" as] + by blast + then show ?thesis + using 0 assms ring_cfs_to_univ_poly_coeffs[of as n l] + by presburger +qed + +lemma ring_cfs_to_univ_poly_coeffs'': + assumes "as \ carrier (R\<^bsup>Suc n\<^esup>)" + shows "(ring_cfs_to_univ_poly n as) l = (if l \ n then as! l else \)" + apply(cases "l \n") + apply (meson assms ring_cfs_to_univ_poly_coeffs') +proof- assume "\ l \ n " then + have A: "n < l" + by auto + have "deg R (ring_cfs_to_univ_poly n as) \ n" + using assms ring_cfs_to_univ_poly_degree(1) by blast + then show ?thesis + using A domain_def[of R] deg_leE assms le_less_trans ring_cfs_to_univ_poly_closed UP_car_memE(2) + by auto +qed +end + +definition fun_tuple_to_univ_poly where +"fun_tuple_to_univ_poly R n m fs x = cring_coord_rings.ring_cfs_to_univ_poly R m (function_tuple_eval R n fs x)" + +context cring_coord_rings +begin + +lemma fun_tuple_to_univ_poly_closed: + assumes "is_function_tuple R n fs" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + assumes "length fs = Suc m" + shows "fun_tuple_to_univ_poly R n m fs x \ carrier (UP R)" + unfolding fun_tuple_to_univ_poly_def + using assms + ring_cfs_to_univ_poly_closed[of "function_tuple_eval R n fs x" m] + function_tuple_eval_closed[of R n fs x] + by metis + +lemma fun_tuple_to_univ_poly_degree_bound: + assumes "is_function_tuple R n fs" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + assumes "length fs = Suc m" + shows "deg R (fun_tuple_to_univ_poly R n m fs x) \ m" + unfolding fun_tuple_to_univ_poly_def + using ring_cfs_to_univ_poly_degree assms + by (metis function_tuple_eval_closed) + +lemma fun_tuple_to_univ_poly_degree: + assumes "is_function_tuple R n fs" + assumes "x \ carrier (R\<^bsup>n\<^esup>)" + assumes "length fs = Suc m" + assumes "(fs!m) x \\" + shows "deg R (fun_tuple_to_univ_poly R n m fs x) = m" + unfolding fun_tuple_to_univ_poly_def + using ring_cfs_to_univ_poly_degree[of "function_tuple_eval R n fs x" m] + assms + function_tuple_eval_def + function_tuple_eval_closed[of R n fs x] + by (metis lessI nth_map) + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Factoring a Polynomial as a Univariate Polynomial over a Multivariable Polynomial Ring\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition pre_to_univ_poly_hom :: "nat \ nat \ ('a, (('a, nat) mvar_poly, nat) mvar_poly) ring_hom" where +"pre_to_univ_poly_hom n i= MP.indexed_const (n-1) \ + R.indexed_const" + +lemma pre_to_univ_poly_hom_is_hom: + assumes "i < n" + shows "ring_hom_ring R (Pring (coord_ring R (n-1)) {i}) (pre_to_univ_poly_hom n i)" + using ring_hom_trans[of R.indexed_const R "coord_ring R (n-1)" + "ring.indexed_const(Pring R ({.. nat \ nat \ (('a, nat) mvar_poly, nat) mvar_poly" where +"pre_to_univ_poly_var_ass n i j =(if j < i then MP.indexed_const (n-1) (pvar R j) else + (if j = i then pvar (coord_ring R (n-1)) i else + (if j < n then MP.indexed_const (n-1) (pvar R (j - 1)) else + \\<^bsub>Pring (coord_ring R (n-1)) {i}\<^esub>)))" + +lemma pre_to_univ_poly_var_ass_closed: + assumes "i < n" + shows "closed_fun (Pring (coord_ring R (n-1)) {i}) (pre_to_univ_poly_var_ass n i)" +proof fix j + show "pre_to_univ_poly_var_ass n i j \ carrier (Pring (coord_ring R (n - 1)) {i})" + unfolding pre_to_univ_poly_var_ass_def + apply(cases "j < i") + using pvar_closed[of j n] assms cring.indexed_const_closed + apply (metis (no_types, lifting) R.Pring_is_cring Suc_diff_1 Suc_le_eq coord_ring_def diff_diff_cancel R.is_cring less_imp_diff_less local.pvar_closed not_less0 not_less_eq_eq) + apply(cases "j = i") + using assms apply (meson pvar_closed R.Pring_is_cring R.is_cring singletonI) + apply(cases "j < n") + using pvar_closed[of "j-1" n] assms MP.indexed_const_closed R.Pring_is_cring Suc_diff_1 Suc_le_eq coord_ring_def R.is_cring pvar_closed neq0_conv not_le + apply (metis MP.Pring_var_closed singletonI) + using MP.Pring_is_ring[of "n-1" "{i}"] apply blast + by (smt MP.Pring_zero_closed MP.indexed_const_closed Suc_diff_1 Suc_le_eq le_eq_less_or_eq less_Suc_eq local.pvar_closed nat_induct) +qed + +lemma pre_to_univ_poly_var_ass_closed': + assumes "i < n" + shows "(pre_to_univ_poly_var_ass n i) \ {.. carrier (Pring (coord_ring R (n-1)) {i})" + by (metis (no_types, lifting) Pi_iff UNIV_I assms pre_to_univ_poly_var_ass_closed) + +definition pre_to_univ_poly :: + "nat \ nat \ (('a, nat) mvar_poly, (('a, nat) mvar_poly, nat) mvar_poly) ring_hom" where +"pre_to_univ_poly (n::nat) (i::nat) = indexed_poly_induced_morphism {.. = pre_to_univ_poly n i" + shows "ring_hom_ring (R[\\<^bsub>n\<^esub>]) (Pring (coord_ring R (n-1)) {i}) \ " + "\j. j < i \ \ (pvar R j) = MP.indexed_const (n-1) (pvar R j)" + "\ (pvar R i) = pvar (coord_ring R (n-1)) i" + "\j. i < j \ j < n \ \ (pvar R j) = MP.indexed_const (n-1) (pvar R (j - 1))" + "\a. a \ carrier R \ \ (coord_const a) = MP.indexed_const (n-1) (coord_const a)" + "\p. p \ carrier (R[\\<^bsub>n\<^esub>]) \ pre_to_univ_poly n i p \ carrier (Pring (coord_ring R (n-1)) {i})" +proof- + have 0: "cring (Pring (coord_ring R (n - 1)) {i})" + using MP.Pring_is_cring coord_cring_cring by blast + have 1: "pre_to_univ_poly_var_ass n i \ {.. carrier (Pring (coord_ring R (n - 1)) {i})" + using Pi_iff assms(1) pre_to_univ_poly_var_ass_closed[of i n] + by blast + have 2: "ring_hom_ring R (Pring (coord_ring R (n - 1)) {i}) (pre_to_univ_poly_hom n i)" + using assms(1) pre_to_univ_poly_hom_is_hom by auto + + show 3:"ring_hom_ring (R[\\<^bsub>n\<^esub>]) (Pring (coord_ring R (n-1)) {i}) \ " + using R.Pring_universal_prop(1)[of "(Pring (coord_ring R (n-1)) {i})" "pre_to_univ_poly_var_ass n i" + "{..] assms 0 1 2 + unfolding pre_to_univ_poly_def + by (metis coord_ring_def) + + show " \j. j < i \ + \ (pvar R j) = MP.indexed_const (n-1) (pvar R j)" + proof- + fix j assume A: "j < i" + then have 00: "MP.indexed_const (n - 1) (pvar R j) = pre_to_univ_poly_var_ass n i j " + unfolding pre_to_univ_poly_var_ass_def by auto + have 01: "j \ {.. (pvar R j) = MP.indexed_const (n-1) (pvar R j)" + using R.Pring_universal_prop(2)[of "(Pring (coord_ring R (n-1)) {i})" "pre_to_univ_poly_var_ass n i" + "{..] assms 0 1 2 01 + MP.is_cring + unfolding pre_to_univ_poly_def 00 unfolding coord_ring_def var_to_IP_def + by blast + qed + show "\ (pvar R i) = pvar (coord_ring R (n - 1)) i" + using R.Pring_universal_prop[of "(Pring (coord_ring R (n-1)) {i})" "pre_to_univ_poly_var_ass n i" + "{..] assms 0 1 2 + unfolding pre_to_univ_poly_def coord_ring_def + using lessThan_iff less_not_refl pre_to_univ_poly_var_ass_def var_to_IP_def + by (metis coord_ring_def) + show "\j. i < j \ j < n \ \ (pvar R j) = MP.indexed_const (n - 1) (pvar R (j - 1))" + using R.Pring_universal_prop[of "(Pring (coord_ring R (n-1)) {i})" "pre_to_univ_poly_var_ass n i" + "{..] assms 0 1 2 + unfolding pre_to_univ_poly_def + using add_diff_inverse_nat lessThan_iff less_diff_conv less_imp_add_positive + not_add_less1 pre_to_univ_poly_var_ass_def trans_less_add2 var_to_IP_def + by (metis (no_types, lifting) coord_ring_def) + show "\a. a \ carrier R \ \ (R.indexed_const a) = MP.indexed_const (n - 1) (R.indexed_const a)" + using R.Pring_universal_prop(3)[of "(Pring (coord_ring R (n-1)) {i})" "pre_to_univ_poly_var_ass n i" + "{..] assms 0 1 2 comp_apply + unfolding pre_to_univ_poly_def pre_to_univ_poly_hom_def + by metis + show "\p. p \ carrier (R[\\<^bsub>n\<^esub>]) \ pre_to_univ_poly n i p \ carrier (Pring (coord_ring R (n - 1)) {i})" + proof- + fix p assume A: "p \ carrier (R[\\<^bsub>n\<^esub>])" + have "\ \ carrier (R[\\<^bsub>n\<^esub>]) \ carrier (Pring (coord_ring R (n - 1)) {i})" + using 3 unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def by blast + then show " pre_to_univ_poly n i p \ carrier (Pring (coord_ring R (n - 1)) {i})" + using A assms + by blast + qed +qed + +lemma insert_at_index_closed: + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "x \ carrier R" + assumes "i \ n" + shows "insert_at_index a x i \ carrier (R\<^bsup>Suc n\<^esup>)" + apply(rule cartesian_power_car_memI') +proof- + have 0: "length (take i a) = i" + using assms(1) assms(3) cartesian_power_car_memE take_closed by blast + have 1: "length (drop i a) = (n - i)" + using assms cartesian_power_car_memE length_drop + by blast + then have "length (x # drop i a) = Suc (n - i)" + by (metis length_Cons) + then show "length (insert_at_index a x i) = Suc n" + using 0 1 assms + by (metis Suc_eq_plus1 cartesian_power_car_memE insert_at_index_length) + show "\ia. ia < Suc n \ insert_at_index a x i ! ia \ carrier R" + proof- fix j assume A: "j < Suc n" + show "insert_at_index a x i ! j \ carrier R" + apply(cases "j < i") + apply (metis A assms(1) assms(3) cartesian_power_car_memE cartesian_power_car_memE' insert_at_index_eq' le_imp_less_Suc less_Suc_eq not_less_eq) + apply(cases "j = i") + apply (metis assms(1) assms(2) assms(3) cartesian_power_car_memE insert_at_index_eq) + proof- assume A1: "\ j < i " "j \i" + then have "i < j" by auto + then have "(take i a @ x # drop i a) ! j = drop i a ! (j - (Suc i))" + by (metis "0" A1(1) Suc_diff_Suc nth_Cons_Suc nth_append) + then show "insert_at_index a x i ! j \ carrier R" + by (metis A \i < j\ assms(1) cartesian_power_car_memE cartesian_power_car_memE' insert_at_index_eq'' less_Suc_eq_0_disj less_Suc_eq_le not_less0) + qed + qed +qed + +lemma pre_to_univ_poly_eval: + assumes "i < Suc n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "x \ carrier R" + assumes "as = insert_at_index a x i" + shows "eval_at_point R as p = eval_at_point R a (total_eval (R[\\<^bsub>n\<^esub>]) (\ i. coord_const x) (pre_to_univ_poly (Suc n) i p))" + apply(rule R.Pring_car_induct''[of p "{.. carrier (R\<^bsup>Suc n\<^esup>)" + using assms insert_at_index_closed + by (meson less_Suc_eq_le) + show " \c. c \ carrier R \ + eval_at_point R as (R.indexed_const c) = + eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (R.indexed_const c)))" + proof- fix c assume "c \ carrier R" + have 00: "eval_at_poly R (coord_const c) as = c" + using assms eval_at_point_const[of c as "Suc n"] "0" \c \ carrier R\ + by blast + have 01: "closed_fun (R[\\<^bsub>n\<^esub>]) (\n. coord_const x)" + using assms(4) R.indexed_const_closed + by (metis Pi_I coord_ring_def) + have 02: "(pre_to_univ_poly (Suc n) i (coord_const c)) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) (coord_const c)" + using pre_to_univ_poly_is_hom(5)[of i "Suc n" _ c] \c \ carrier R\ assms(1) diff_Suc_1 + by (metis coord_ring_def) + have 03: "(total_eval (R[\\<^bsub>n\<^esub>]) (\ i. coord_const x) (pre_to_univ_poly (Suc n) i (coord_const c))) = + coord_const c" + using 01 cring.total_eval_const[of "R[\\<^bsub>n\<^esub>]" "coord_const c" ] + by (smt "02" MP.total_eval_const \c \ carrier R\ coord_ring_def cring.indexed_const_closed R.is_cring) + show " eval_at_point R as (R.indexed_const c) = + eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (R.indexed_const c))) " + using assms 00 02 03 + by (metis \c \ carrier R\ coord_ring_def eval_at_point_const) + qed + have 01: "closed_fun (R[\\<^bsub>n\<^esub>]) (\n. coord_const x)" + using assms(4) R.indexed_const_closed + by (metis Pi_I coord_ring_def) + have 02: "ring_hom_ring (R[\\<^bsub>Suc n\<^esub>]) (Pring (R[\\<^bsub>n\<^esub>]) {i}) (pre_to_univ_poly (Suc n) i)" + using pre_to_univ_poly_is_hom(1)[of i "Suc n" ] + by (simp add: assms) + show "\p q. p \ carrier (Pring R {.. + q \ carrier (Pring R {.. + eval_at_point R as p = eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i p)) \ + eval_at_point R as q = eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i q)) \ + eval_at_point R as (p \\<^bsub>Pring R {.. q) = + eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. q)))" + proof- fix p q assume A: "p \ carrier (Pring R {.. carrier (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i p))" + " eval_at_point R as q = eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i q))" + have 0: "eval_at_poly R (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> q) as = + eval_at_poly R p as \\<^bsub>R\<^esub> eval_at_poly R q as" + using "0" A(1) A(2) eval_at_point_add unfolding coord_ring_def + by blast + have 1: "(total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> q))) = + (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p)) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> + (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i q))" + proof- + have 10: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(6)[of i "Suc n" _ p] + unfolding coord_ring_def + by (metis A(1) assms(1) diff_Suc_1) + have 11: "pre_to_univ_poly (Suc n) i q \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(6)[of i "Suc n" _ q] + unfolding coord_ring_def + + by (metis A(2) assms(1) diff_Suc_1) + have 12: "(pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> q)) = + (pre_to_univ_poly (Suc n) i p \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> pre_to_univ_poly (Suc n) i q)" + using ring_hom_ring.homh A 02 ring_hom_add[of "pre_to_univ_poly (Suc n) i" "R[\\<^bsub>Suc n\<^esub>] " "Pring (R[\\<^bsub>n\<^esub>]) {i}" + p q ] + unfolding coord_ring_def + + by blast + + show ?thesis + + using 01 10 11 12 A cring.total_eval_add[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i p" "{i}" + "pre_to_univ_poly (Suc n) i q" "\i. coord_const x"] + coord_cring_cring + unfolding coord_ring_def + + by smt + qed + have 2: "eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> q))) a = + eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p)) a \ + eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i q)) a" + proof- + have 20: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i}) " + using A(1) 02 unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def unfolding coord_ring_def + by blast + have 21: "pre_to_univ_poly (Suc n) i q \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i}) " + using A(2) 02 unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def unfolding coord_ring_def + by blast + have 22: "(total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p)) \ + carrier (R[\\<^bsub>n\<^esub>])" + using 21 01 A cring.total_eval_closed[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i p" + "{i}" "\i. coord_const x"] "20" coord_cring_cring + by metis + have 23: "(total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i q)) \ + carrier (R[\\<^bsub>n\<^esub>])" + using cring.total_eval_closed[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i q" "{i}" + "\i. coord_const x"] + by (metis "01" "21" coord_cring_cring) + + show ?thesis + using "1" "22" "23" assms(3) eval_at_point_add by presburger + qed + show "eval_at_point R as (p \\<^bsub>Pring R {.. q) = + eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. q)))" + using eval_at_point_add A 0 1 2 + unfolding coord_ring_def + + by presburger + qed + fix p j + assume A: "p \ carrier (Pring R {.. {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i p))" + show "eval_at_point R as (p \\<^bsub>Pring R {.. pvar R j) = + eval_at_point R a + (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. pvar R j)))" + proof- + have A0: "eval_at_poly R (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pvar R j) as = + eval_at_poly R p as \ as!j" + proof- + have "eval_at_poly R (pvar R j) as = as!j" + using A(2) 0 eval_pvar + by blast + then show ?thesis using A eval_at_point_mult[of as "Suc n" p "pvar R j" ] 0 + by (metis R.Pring_var_closed coord_ring_def) + qed + have A1: "(pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pvar R j)) = + (pre_to_univ_poly (Suc n) i p) \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> pre_to_univ_poly (Suc n) i (pvar R j)" + using A 02 ring_hom_ring.homh ring_hom_mult[of _ "R[\\<^bsub>Suc n\<^esub>]" _ p "pvar R j"] R.Pring_var_closed[of j "{..< Suc n}"] + unfolding coord_ring_def + by blast + have A2: "(total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pvar R j))) = + (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p ))\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> + (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) ( pre_to_univ_poly (Suc n) i (pvar R j)))" + proof- + have A20: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using 02 A unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def + unfolding coord_ring_def + + by blast + have A21: "pre_to_univ_poly (Suc n) i (pvar R j) \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using 02 A unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_hom_def + using R.Pring_var_closed[of j "{..< Suc n}"] + unfolding coord_ring_def + + by blast + show ?thesis using A1 cring.total_eval_mult[of _ "pre_to_univ_poly (Suc n) i p"] + + by (smt A20 A21 MP.closed_funI MP.total_eval_mult assms(4) coord_ring_def cring.indexed_const_closed R.is_cring) + qed + have A3: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using 02 A ring_hom_ring.homh unfolding ring_hom_def + unfolding coord_ring_def + by blast + have A4: "pre_to_univ_poly (Suc n) i (pvar R j) \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using 02 A ring_hom_ring.homh R.Pring_var_closed[of j "{..< Suc n}"] unfolding ring_hom_def + unfolding coord_ring_def + + by blast + have A5: "total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p ) \ + carrier (R[\\<^bsub>n\<^esub>])" + using 01 cring.total_eval_closed[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i p " "{i}"] + A3 coord_cring_cring + unfolding coord_ring_def + by smt + have A6: "total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j) ) \ + carrier (R[\\<^bsub>n\<^esub>])" + using 01 cring.total_eval_closed[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i (pvar R j) " "{i}"] + A4 coord_cring_cring + unfolding coord_ring_def + by smt + have A7: " eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pvar R j))) a + = eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i p)) a \ + eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j))) a" + using eval_at_point_mult A5 A6 A2 assms(3) by presburger + have A8: "eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j))) a = + as!j" + proof(cases "j = i") + case True + then have "pre_to_univ_poly (Suc n) i (pvar R j) = pvar (R[\\<^bsub>n\<^esub>]) i" + using pre_to_univ_poly_is_hom(3)[of i "Suc n"] assms(1) diff_Suc_1 by presburger + then have "total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j)) = + coord_const x" + using cring.total_eval_var[of "R[\\<^bsub>n\<^esub>]" "\i. coord_const x"] + unfolding coord_ring_def + by (smt "01" \\i. \cring (R [\\<^bsub>n\<^esub>]); (\i. R.indexed_const x) \ UNIV \ carrier (R [\\<^bsub>n\<^esub>])\ \ total_eval (R [\\<^bsub>n\<^esub>]) (\i. R.indexed_const x) (mset_to_IP (R [\\<^bsub>n\<^esub>]) {#i#}) = R.indexed_const x\ coord_ring_def cring_coord_rings.coord_cring_cring cring_coord_rings_axioms var_to_IP_def) + then have T0: "eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j))) a + = x" + using eval_at_point_const + by (metis assms(3) assms(4)) + have T1: "as!j = x" + using assms + by (metis True assms(5) cartesian_power_car_memE insert_at_index_eq le_eq_less_or_eq nat_le_linear + not_less_eq) + then show ?thesis + using T0 by blast + next + case False + then show ?thesis + proof(cases "j < i") + case True + then have "pre_to_univ_poly (Suc n) i (pvar R j) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) (pvar R j)" + using pre_to_univ_poly_is_hom(2)[of i "Suc n"] assms(1) diff_Suc_1 + unfolding coord_ring_def + by presburger + then have "total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j)) = + pvar R j" + using cring.total_eval_const[of "R[\\<^bsub>n\<^esub>]"] + by (smt Suc_less_eq True assms(1) coord_cring_cring less_trans_Suc local.pvar_closed) + then have T0: "eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j))) a + = a!j" + using eval_pvar + by (metis Suc_less_eq True assms(1) assms(3) less_trans_Suc) + have T1: "as!j = a!j" + using assms + by (metis True assms(5) cartesian_power_car_memE insert_at_index_eq' less_Suc_eq_le) + then show ?thesis + using T0 by presburger + next + case F: False + then have "pre_to_univ_poly (Suc n) i (pvar R j) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) (pvar R (j-1))" + using pre_to_univ_poly_is_hom(4)[of i "Suc n"] assms(1) diff_Suc_1 + unfolding coord_ring_def + by (metis A(2) False lessThan_iff linorder_neqE_nat) + then have "total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j)) = + pvar R (j-1)" + using cring.total_eval_const[of "R[\\<^bsub>n\<^esub>]"] + by (smt A(2) F False Suc_less_SucD add_diff_inverse_nat coord_cring_cring + lessThan_iff less_one linorder_neqE_nat local.pvar_closed not_less0 plus_1_eq_Suc) + then have T0: "eval_at_poly R (total_eval (R[\\<^bsub>n\<^esub>]) (\i. coord_const x) (pre_to_univ_poly (Suc n) i (pvar R j))) a + = a!(j-1)" + using eval_pvar[of "j-1" n a] + by (metis A(2) F False One_nat_def Suc_diff_Suc Suc_less_eq assms(3) + lessThan_iff linorder_neqE_nat minus_nat.diff_0 not_less0) + have T1: "as!j = a!(j-1)" + proof- + obtain k where k_def: "j = i + 1 + k" + using False F + by (metis Nat.add_0_right less_imp_add_positive less_one + nat_neq_iff semiring_normalization_rules(25)) + show "as!j = a!(j-1)" + proof- + have "length (take i a) = i" + using assms + by (meson cartesian_power_car_memE less_Suc_eq_le take_closed) + then have "as!j = ( x # drop i a)!(k+1)" + using k_def assms + unfolding coord_ring_def + by (metis Suc_eq_plus1 add.assoc insert_at_index.simps nth_append_length_plus plus_1_eq_Suc) + have "length (drop i a) \ k" + proof- + have "length (drop i a) = n - i" + using assms cartesian_power_car_memE length_drop + by blast + then show ?thesis + using assms k_def A(2) + by (metis Suc_eq_plus1 add.commute diff_Suc_Suc lessThan_iff less_diff_conv less_imp_le_nat) + qed + then have "as!j = (drop i a)! k" + using assms k_def + by (metis Nat.add_0_right One_nat_def \as ! j = (x # drop i a) ! (k + 1)\ add_Suc_right nth_Cons_Suc) + then show ?thesis using k_def assms + by (metis Nat.add_diff_assoc2 add_diff_cancel_right' cartesian_power_car_memE le_add2 less_Suc_eq_le nth_drop) + qed + qed + then show ?thesis + using T0 by presburger + qed + qed + then show "eval_at_point R as (p \\<^bsub>Pring R {.. pvar R j) = + eval_at_point R a (total_eval (Pring R {..i. R.indexed_const x) (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. pvar R j)))" + using A(3) A0 A7 + unfolding coord_ring_def + by presburger + qed +qed + +definition pre_to_univ_poly_inv_hom :: + "nat \ nat \ (('a, nat) mvar_poly,('a, nat) mvar_poly) ring_hom" where +"pre_to_univ_poly_inv_hom n i = R.relabel_vars {..<(n-1)} {..j. if j < i then j else j + 1)" + +lemma pre_to_univ_poly_inv_hom_is_hom: + assumes "i < Suc n" + shows "ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>Suc n\<^esub>]) (pre_to_univ_poly_inv_hom (Suc n) i)" +proof- + have 0: "ring_hom_ring (R[\\<^bsub>n\<^esub>]) (R[\\<^bsub>Suc n\<^esub>]) (R.relabel_vars {..j. if j < i then j else j + 1))" + unfolding coord_ring_def + apply(rule R.relabel_vars_is_morphism) + using assms + by (smt Pi_I Suc_eq_plus1 add_less_cancel_right lessThan_iff less_Suc_eq) + then show ?thesis + unfolding pre_to_univ_poly_inv_hom_def + by simp +qed + +lemma pre_to_univ_poly_inv_hom_const: + assumes "i < Suc n" + assumes "k \ carrier R" + shows "(pre_to_univ_poly_inv_hom (Suc n) i) (R.indexed_const k) = R.indexed_const k" +proof- + have 0: "(R.relabel_vars {..j. if j < i then j else j + 1)) (R.indexed_const k) = R.indexed_const k" + unfolding coord_ring_def + apply(rule R.relabel_vars_is_morphism) + using assms + apply (smt Pi_I Suc_eq_plus1 add_less_cancel_right lessThan_iff less_Suc_eq) + using assms(2) by blast + then show ?thesis + unfolding pre_to_univ_poly_inv_hom_def + using diff_Suc_1 by presburger +qed + +lemma pre_to_univ_poly_inv_hom_pvar_0: + assumes "i < Suc n" + assumes "j < i" + shows "pre_to_univ_poly_inv_hom (Suc n) i (pvar R j) = + pvar R j" + unfolding pre_to_univ_poly_inv_hom_def coord_ring_def + using R.relabel_vars_is_morphism(2)[of "\j. if j < i then j else j + 1" "{.. j" + assumes "j < n" + shows "pre_to_univ_poly_inv_hom (Suc n) i (pvar R j) = + pvar R (j + 1)" + unfolding pre_to_univ_poly_inv_hom_def + using assms R.relabel_vars_is_morphism(2)[of "\j. if j < i then j else j + 1" "{.. nat \ nat \ ('a, nat) mvar_poly" where +"pre_to_univ_poly_inv_var_ass n i j = pvar R i" + +lemma pre_to_univ_poly_inv_var_ass_closed: + assumes "i < Suc n" + shows "pre_to_univ_poly_inv_var_ass (Suc n) i \ {i} \ carrier (R[\\<^bsub>Suc n\<^esub>])" + by (metis Pi_I assms local.pvar_closed pre_to_univ_poly_inv_var_ass_def) + +definition pre_to_univ_poly_inv :: + "nat \ nat \ ((('a, nat) mvar_poly, nat) mvar_poly,('a, nat) mvar_poly) ring_hom" where +"pre_to_univ_poly_inv n i = indexed_poly_induced_morphism {i} (R[\\<^bsub>n\<^esub>]) + (pre_to_univ_poly_inv_hom n i) (pre_to_univ_poly_inv_var_ass n i)" + +lemma pre_to_univ_poly_inv_is_hom: + assumes "i < Suc n" + shows "ring_hom_ring (Pring (R[\\<^bsub>n\<^esub>]) {i}) (R[\\<^bsub>Suc n\<^esub>]) (pre_to_univ_poly_inv (Suc n) i)" + apply(rule cring.Pring_universal_prop[of _ _ "pre_to_univ_poly_inv_var_ass (Suc n) i" "{i}" "pre_to_univ_poly_inv_hom (Suc n) i"]) + unfolding coord_ring_def + apply (simp add: R.Pring_is_cring R.is_cring) + apply (simp add: R.Pring_is_cring R.is_cring) + apply (metis Pi_I R.Pring_var_closed assms lessThan_iff pre_to_univ_poly_inv_var_ass_def) + apply (metis assms coord_ring_def pre_to_univ_poly_inv_hom_is_hom) + by (simp add: coord_ring_def pre_to_univ_poly_inv_def) + +lemma pre_to_univ_poly_inv_pvar: + assumes "i < Suc n" + shows "(pre_to_univ_poly_inv (Suc n) i) (pvar (R[\\<^bsub>n\<^esub>]) i) = pvar R i" + using assms cring.Pring_universal_prop[of "R[\\<^bsub>n\<^esub>]" "R[\\<^bsub>Suc n\<^esub>]" + "pre_to_univ_poly_inv_var_ass (Suc n) i" "{i}" "pre_to_univ_poly_inv_hom (Suc n) i"] + by (metis Pi_I coord_cring_cring cring_coord_rings.pre_to_univ_poly_inv_var_ass_def + cring_coord_rings_axioms local.pvar_closed pre_to_univ_poly_inv_def + pre_to_univ_poly_inv_hom_is_hom singletonI var_to_IP_def) + +lemma pre_to_univ_poly_inv_const: + assumes "i < Suc n" + assumes "p \ carrier (R[\\<^bsub>n\<^esub>])" + shows "(pre_to_univ_poly_inv (Suc n) i) (ring.indexed_const (R[\\<^bsub>n\<^esub>]) p) = pre_to_univ_poly_inv_hom (Suc n) i p " + using assms cring.Pring_universal_prop[of "R[\\<^bsub>n\<^esub>]" "R[\\<^bsub>Suc n\<^esub>]" + "pre_to_univ_poly_inv_var_ass (Suc n) i" "{i}" "pre_to_univ_poly_inv_hom (Suc n) i"] + by (metis Pi_I coord_cring_cring cring_coord_rings.pre_to_univ_poly_inv_var_ass_def + cring_coord_rings_axioms local.pvar_closed pre_to_univ_poly_inv_def pre_to_univ_poly_inv_hom_is_hom) + +lemma pre_to_univ_poly_inverse: + assumes "i < Suc n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + shows "pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i p) = p" + apply(rule R.Pring_car_induct''[of p "{..c. c \ carrier R \ pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i (coord_const c)) = coord_const c" + proof- + fix c assume A: "c \ carrier R" + have 0: "pre_to_univ_poly (Suc n) i (coord_const c) = + MP.indexed_const n (coord_const c)" + using A assms(1) diff_Suc_1 pre_to_univ_poly_is_hom(5) by presburger + have 1: "(\j. if j < i then j else j + 1) \ {.. {..j. if j < i then j else j + 1)" "{..p q. p \ carrier (Pring R {.. + q \ carrier (Pring R {.. + pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i p) = + p \ + pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i q) = + q \ + pre_to_univ_poly_inv (Suc n) i + (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. q)) = + p \\<^bsub>Pring R {.. q" + proof- fix p q assume A: "p \ carrier (Pring R {.. carrier (Pring R {..\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> q)) = + (pre_to_univ_poly (Suc n) i p) \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> pre_to_univ_poly (Suc n) i q" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] ring_hom_ring.homh ring_hom_add A + unfolding coord_ring_def + by (metis (mono_tags, lifting) assms(1) diff_Suc_1) + have 1: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] A + unfolding coord_ring_def + by (metis assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + have 2: "pre_to_univ_poly (Suc n) i q \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] A + unfolding coord_ring_def + by (metis assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + show "pre_to_univ_poly_inv (Suc n) i + (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. q)) = + p \\<^bsub>Pring R {.. q" + using 0 1 2 A pre_to_univ_poly_inv_is_hom[of i n] ring_hom_ring.homh ring_hom_add + unfolding coord_ring_def + by (smt assms(1)) + qed + show "\p ia. + p \ carrier (Pring R {.. + ia \ {.. + pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i p) = p \ + pre_to_univ_poly_inv (Suc n) i + (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. pvar R ia)) = + p \\<^bsub>Pring R {.. pvar R ia" + proof- fix p j + assume A: " p \ carrier (Pring R {.. {..\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pvar R j)) + = (pre_to_univ_poly (Suc n) i p) \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> pre_to_univ_poly (Suc n) i (pvar R j)" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] ring_hom_ring.homh ring_hom_mult A + unfolding coord_ring_def + by (metis R.Pring_var_closed assms(1) diff_Suc_1) + have 1: "pre_to_univ_poly (Suc n) i p \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] A + unfolding coord_ring_def + by (metis assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + have 1: "pre_to_univ_poly (Suc n) i (pvar R j) \ carrier (Pring (R[\\<^bsub>n\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] A + unfolding coord_ring_def + by (metis R.Pring_var_closed assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + have 2: "pre_to_univ_poly_inv (Suc n) i (pre_to_univ_poly (Suc n) i (pvar R j)) = pvar R j" + proof(cases "j = i") + case True + then have "(pre_to_univ_poly (Suc n) i (pvar R j)) = pvar (R[\\<^bsub>n\<^esub>]) j" + using pre_to_univ_poly_is_hom(3)[of i "Suc n"] assms(1) diff_Suc_1 by presburger + then show ?thesis + unfolding coord_ring_def + using True \pre_to_univ_poly (Suc n) i (pvar R j) = pvar (R[\\<^bsub>n\<^esub>]) j\ assms(1) pre_to_univ_poly_inv_pvar by presburger + next + case False + show ?thesis + proof(cases "j < i") + case True + then have "(pre_to_univ_poly (Suc n) i (pvar R j)) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) (pvar R j)" + using pre_to_univ_poly_is_hom(2) [of i "Suc n"] assms(1) diff_Suc_1 + unfolding coord_ring_def + + by presburger + then show ?thesis + using pre_to_univ_poly_inv_const[of i n "(pvar R j)"] + pre_to_univ_poly_inv_hom_pvar_0[of i n j] + by (metis Suc_less_eq True assms(1) less_trans_Suc local.pvar_closed) + next + case F: False + then have "(pre_to_univ_poly (Suc n) i (pvar R j)) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) (pvar R (j-1))" + using pre_to_univ_poly_is_hom(4)[of i "Suc n"] assms(1) diff_Suc_1 + unfolding coord_ring_def + by (metis A(2) False lessThan_iff linorder_neqE_nat) + then show ?thesis + using pre_to_univ_poly_inv_const[of i n "(pvar R (j-1))"] + pre_to_univ_poly_inv_hom_pvar_0[of i n "j-1"] + by (metis (no_types, lifting) A(2) F False One_nat_def Suc_eq_plus1 add_diff_inverse_nat + assms(1) le_neq_implies_less lessThan_iff less_one local.pvar_closed nat_le_linear + not_less_eq plus_1_eq_Suc pre_to_univ_poly_inv_hom_pvar_1) + qed + qed + show "pre_to_univ_poly_inv (Suc n) i + (pre_to_univ_poly (Suc n) i (p \\<^bsub>Pring R {.. pvar R j)) = + p \\<^bsub>Pring R {.. pvar R j" + using 0 1 2 A pre_to_univ_poly_inv_is_hom[of i n] + ring_hom_ring.homh[of _ _ "pre_to_univ_poly_inv (Suc n) i "] + ring_hom_mult[of "pre_to_univ_poly_inv (Suc n) i "] + unfolding coord_ring_def + by (smt assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + qed +qed + +lemma coord_ring_car_induct: + assumes "Q \ carrier (R[\\<^bsub>n\<^esub>])" + assumes "\c. c \ carrier R \ A (R.indexed_const c)" + assumes "\p q. p \ carrier (R[\\<^bsub>n\<^esub>]) \ q \ carrier (R[\\<^bsub>n\<^esub>]) \ A p \ A q \ A (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q)" + assumes "\p i. p \ carrier (R[\\<^bsub>n\<^esub>]) \ i < n \ A p \ A (p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> pvar R i)" + shows "A Q" + unfolding coord_ring_def apply(rule R.Pring_car_induct''[of _ "{.. carrier (R[\\<^bsub>n\<^esub>])" + shows "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) = MP.indexed_const n p" + apply(rule coord_ring_car_induct[of _ n]) + using assms(2) apply blast +proof- + show "\c. c \ carrier R \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (R.indexed_const c))) = + MP.indexed_const n (R.indexed_const c)" + proof- fix k assume A: "k \ carrier R" + have 0: "R.indexed_const k \ carrier (R [\\<^bsub>n\<^esub>])" + using A + by (metis coord_ring_def R.indexed_const_closed) + have 1: "pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (R.indexed_const k)) = pre_to_univ_poly_inv_hom (Suc n) i (R.indexed_const k)" + using 0 assms pre_to_univ_poly_inv_const[of i n "R.indexed_const k"] + by linarith + have "pre_to_univ_poly_inv_hom (Suc n) i (R.indexed_const k) = R.indexed_const k" + using A pre_to_univ_poly_inv_hom_const[of i n k] assms + by blast + thus "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (R.indexed_const k))) = MP.indexed_const n (R.indexed_const k) " + using 1 + by (metis A assms(1) coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(5)) + qed + show "\p q. p \ carrier (R [\\<^bsub>n\<^esub>]) \ + q \ carrier (R [\\<^bsub>n\<^esub>]) \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) = MP.indexed_const n p \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n q)) = MP.indexed_const n q \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> q))) = MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> q)" + proof- fix p Q + assume A: "p \ carrier (R [\\<^bsub>n\<^esub>]) " + "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) = MP.indexed_const n p" + "Q \ carrier (R [\\<^bsub>n\<^esub>])" + "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n Q)) = MP.indexed_const n Q " + have 0: "p \ Q = p \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> Q" + by (metis R.Pring_add coord_ring_def) + have 1: "MP.indexed_const n (p \ Q) = (MP.indexed_const n p) \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub>(MP.indexed_const n Q)" + by (metis "0" MP.Pring_add MP.indexed_padd_const) + have 2: "MP.indexed_const n p \ carrier (Pring (R [\\<^bsub>n\<^esub>]) {i})" + using A unfolding coord_ring_def + by (metis MP.indexed_const_closed R.Pring_car coord_ring_def) + have 3: "MP.indexed_const n Q \ carrier (Pring (R [\\<^bsub>n\<^esub>]) {i})" + using A(3) MP.indexed_const_closed by blast + have 4: "(pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> Q))) = + pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p) \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n Q)" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] + pre_to_univ_poly_inv_is_hom(1)[of i n] + ring_hom_add[of "pre_to_univ_poly_inv (Suc n) i" "(Pring (R [\\<^bsub>n\<^esub>]) {i})" + "(R [\\<^bsub>Suc n\<^esub>])" "MP.indexed_const n p" "MP.indexed_const n Q"] + ring_hom_ring.homh + MP.indexed_const_closed[of p n "{i}"] + MP.indexed_const_closed[of Q n "{i}"] A R.Pring_car[of "{..\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> Q))) = + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n Q))" + proof- + have 50: "pre_to_univ_poly (Suc n) i \ ring_hom (R [\\<^bsub>Suc n\<^esub>]) (Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i})" + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] ring_hom_ring.homh + by (metis assms(1)) + have 51: "pre_to_univ_poly_inv (Suc n) i \ ring_hom (Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}) (R [\\<^bsub>Suc n\<^esub>]) " + using pre_to_univ_poly_inv_is_hom ring_hom_ring.homh + by (metis assms(1) diff_Suc_1) + have 52: " pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p) \ carrier (R [\\<^bsub>Suc n\<^esub>])" + using 51 ring_hom_closed[of "pre_to_univ_poly_inv (Suc n) i" ] + by (smt "2" diff_Suc_1) + have 53: " pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n Q) \ carrier (R [\\<^bsub>Suc n\<^esub>]) " + using 51 ring_hom_closed[of "pre_to_univ_poly_inv (Suc n) i" ] + by (smt 3 diff_Suc_1) + show ?thesis using 50 51 52 53 + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] + ring_hom_add[of "pre_to_univ_poly (Suc n) i" "R [\\<^bsub>Suc n\<^esub>]" "Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}" + "pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)" + "pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n Q)"] 4 + by (metis diff_Suc_1) + qed + show "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> Q))) = MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> Q)" + using 5 A "0" "1" by metis + qed + show "\p ia. + p \ carrier (R [\\<^bsub>n\<^esub>]) \ + ia < n \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) = MP.indexed_const n p \ + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R ia))) = MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R ia)" + proof- fix p j + assume A: "p \ carrier (R [\\<^bsub>n\<^esub>])" "j < n" + "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) = MP.indexed_const n p" + show "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R j))) = MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R j)" + proof- + have 0: "pre_to_univ_poly_inv (Suc n) i \ ring_hom (Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}) (R [\\<^bsub>Suc n\<^esub>]) " + using pre_to_univ_poly_inv_is_hom(1)[of i n] ring_hom_ring.homh + by (metis assms(1) diff_Suc_1) + have 1: "MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> (pvar R j)) = MP.indexed_const n p \\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> MP.indexed_const n (pvar R j)" + by (metis A(1) A(2) MP.indexed_const_mult local.pvar_closed) + have 2: "(pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R j))) = + pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p) \\<^bsub>(R [\\<^bsub>Suc n\<^esub>])\<^esub> pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j))" + using 0 1 ring_hom_mult A + by (metis (no_types, lifting) MP.indexed_const_closed diff_Suc_1 local.pvar_closed) + have 3: "pre_to_univ_poly(Suc n) i \ ring_hom (R [\\<^bsub>Suc n\<^esub>]) (Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}) " + using assms(1) pre_to_univ_poly_is_hom(1) ring_hom_ring.homh by blast + have 4: "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R j))) = + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n p)) \\<^bsub>Pring (R [\\<^bsub>n\<^esub>]) {i}\<^esub> + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j)))" + using 2 3 ring_hom_mult + by (smt "0" A(1) A(2) MP.indexed_const_closed diff_Suc_1 local.pvar_closed ring_hom_closed) + have 5: "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (p \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> pvar R j))) = + MP.indexed_const n p \\<^bsub>Pring (R [\\<^bsub>n\<^esub>]) {i}\<^esub> + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j)))" + using A "4" by presburger + have 6: "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j))) = (MP.indexed_const n (pvar R j))" + proof- + have "(pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j))) = pre_to_univ_poly_inv_hom (Suc n) i (pvar R j)" + using A(2) assms(1) local.pvar_closed pre_to_univ_poly_inv_const by blast + hence "pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j))) = + pre_to_univ_poly (Suc n) i (pre_to_univ_poly_inv_hom (Suc n) i (pvar R j))" + by presburger + show ?thesis + proof(cases "j < i") + case True + then have "(pre_to_univ_poly_inv_hom (Suc n) i (pvar R j)) = (pvar R j)" + using pre_to_univ_poly_inv_hom_pvar_0[of i n j] assms(1) by blast + thus ?thesis using pre_to_univ_poly_is_hom + by (metis True \pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j)) = pre_to_univ_poly_inv_hom (Suc n) i (pvar R j)\ assms(1) coord_ring_def diff_Suc_1) + next + case False + have "pre_to_univ_poly_inv_hom (Suc n) i (pvar R j) = pvar R (j + 1)" + using pre_to_univ_poly_inv_hom_pvar_1[of i n j] A(2) False assms(1) not_le + by blast + thus ?thesis using pre_to_univ_poly_is_hom + by (metis A(2) False Suc_eq_plus1 \pre_to_univ_poly_inv (Suc n) i (MP.indexed_const n (pvar R j)) = pre_to_univ_poly_inv_hom (Suc n) i (pvar R j)\ + assms(1) coord_ring_def diff_Suc_1 not_less_eq) + qed + qed + show ?thesis using 6 A + using "1" "5" by presburger + qed + qed +qed + +definition to_univ_poly :: "nat \ nat \ + (('a, nat) mvar_poly , ('a, nat) mvar_poly u_poly) ring_hom" where +"to_univ_poly n i = IP_to_UP i \ (pre_to_univ_poly n i) " + +definition from_univ_poly :: "nat \ nat \ + (('a, nat) mvar_poly u_poly , ('a, nat) mvar_poly) ring_hom" where +"from_univ_poly n i = pre_to_univ_poly_inv n i \ (UP_to_IP (coord_ring R (n-1)) i)" + +lemma to_univ_poly_is_hom: + assumes "i \ n" + shows "(to_univ_poly (Suc n) i) \ ring_hom (R[\\<^bsub>Suc n\<^esub>]) (UP (R[\\<^bsub>n\<^esub>])) " + unfolding to_univ_poly_def + apply(rule ring_hom_trans[of _ _ "Pring (R[\\<^bsub>n\<^esub>]) {i}"]) + using assms pre_to_univ_poly_is_hom ring_hom_ring.homh + apply (metis diff_Suc_1 le_imp_less_Suc) + using UP_cring.IP_to_UP_ring_hom[of "(Pring R {.. n" + shows "(from_univ_poly (Suc n) i) \ ring_hom (UP (R[\\<^bsub>n\<^esub>])) (R[\\<^bsub>Suc n\<^esub>]) " + unfolding from_univ_poly_def + apply(rule ring_hom_trans[of _ _ "Pring (R[\\<^bsub>n\<^esub>]) {i}"]) + using assms UP_cring.UP_to_IP_ring_hom[of "coord_ring R (Suc n - 1)" i] + ring_hom_ring.homh[of "UP (coord_ring R (Suc n - 1))" "Pring (coord_ring R (Suc n - 1)) {i}" "UP_to_IP (coord_ring R (Suc n - 1)) i"] + unfolding coord_ring_def UP_cring_def + apply (metis R.Pring_is_cring diff_Suc_1 R.is_cring) + using assms ring_hom_ring.homh le_imp_less_Suc pre_to_univ_poly_inv_is_hom + unfolding coord_ring_def UP_cring_def + by blast + +lemma to_univ_poly_inverse: + assumes "i \ n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + shows "from_univ_poly (Suc n) i (to_univ_poly (Suc n) i p) = p" +proof- + have 0: "pre_to_univ_poly (Suc n) i p \ Pring_set (R[\\<^bsub>n\<^esub>]) {i}" + using pre_to_univ_poly_is_hom(6)[of i "Suc n" _ p] assms ring.Pring_car + unfolding coord_ring_def UP_domain_def + by (metis R.Pring_is_ring diff_Suc_1 le_imp_less_Suc) + have 1: "UP_to_IP (R[\\<^bsub>n\<^esub>]) i + (IP_to_UP i (pre_to_univ_poly (Suc n) i p)) = + pre_to_univ_poly (Suc n) i p" + using 0 UP_cring.UP_to_IP_inv[of "R[\\<^bsub>n\<^esub>]" "pre_to_univ_poly (Suc n) i p" i ] + R.Pring_is_cring + unfolding coord_ring_def UP_cring_def + using R.is_cring by blast + have 2: "from_univ_poly (Suc n) i (to_univ_poly (Suc n) i p) = + (pre_to_univ_poly_inv (Suc n) i ( + (UP_to_IP (coord_ring R (Suc n - 1)) i) ( + (IP_to_UP i ( + (pre_to_univ_poly (Suc n) i) p)))))" + unfolding from_univ_poly_def to_univ_poly_def + unfolding coord_ring_def + by (metis comp_eq_dest_lhs) + have 3: "from_univ_poly (Suc n) i (to_univ_poly (Suc n) i p) = + (pre_to_univ_poly_inv (Suc n) i ( + pre_to_univ_poly (Suc n) i p))" + using 0 1 2 + unfolding coord_ring_def + using diff_Suc_1 by presburger + then show ?thesis + using pre_to_univ_poly_inverse assms(1) assms(2) less_Suc_eq_le + by presburger +qed + +lemma to_univ_poly_closed: + assumes "i \ n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + shows "to_univ_poly (Suc n) i p \ carrier (UP (R[\\<^bsub>n\<^esub>]))" + using to_univ_poly_is_hom[of i n] assms unfolding ring_hom_def + by blast + +lemma to_univ_poly_add: + assumes "i \ n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + shows "to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>Q) = + to_univ_poly (Suc n) i p \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub> to_univ_poly (Suc n) i Q" + using to_univ_poly_is_hom ring_hom_add + by (metis assms(1) assms(2) assms(3)) + +lemma to_univ_poly_mult: + assumes "i \ n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "Q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + shows "to_univ_poly (Suc n) i (p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>Q) = + to_univ_poly (Suc n) i p \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub> to_univ_poly (Suc n) i Q" + using to_univ_poly_is_hom ring_hom_mult + by (metis assms(1) assms(2) assms(3)) + +lemma from_univ_poly_closed: + assumes "i \ n" + assumes "p \ carrier (UP (R[\\<^bsub>n\<^esub>])) " + shows "from_univ_poly (Suc n) i p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + using from_univ_poly_is_hom[of i n] assms unfolding ring_hom_def + by blast + +lemma from_univ_poly_add: + assumes "i \ n" + assumes "p \ carrier (UP (R[\\<^bsub>n\<^esub>])) " + assumes "Q \ carrier (UP (R[\\<^bsub>n\<^esub>])) " + shows "from_univ_poly (Suc n) i (p \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub>Q) = + from_univ_poly (Suc n) i p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> from_univ_poly (Suc n) i Q" + using from_univ_poly_is_hom ring_hom_add + by (metis assms(1) assms(2) assms(3)) + +lemma from_univ_poly_mult: + assumes "i \ n" + assumes "p \ carrier (UP (R[\\<^bsub>n\<^esub>])) " + assumes "Q \ carrier (UP (R[\\<^bsub>n\<^esub>])) " + shows "from_univ_poly (Suc n) i (p \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub>Q) = + from_univ_poly (Suc n) i p \\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> from_univ_poly (Suc n) i Q" + using from_univ_poly_is_hom ring_hom_mult + by (metis assms(1) assms(2) assms(3)) + +lemma(in UP_cring) monom_as_mult: + assumes "a \ carrier R" + shows "up_ring.monom (UP R) a n = to_poly a \\<^bsub> UP R\<^esub> up_ring.monom (UP R) \ n" + by (metis One_nat_def P_def R.one_closed R.r_one UP_cring.poly_shift_monom add_Suc assms is_UP_cring local.monom_mult plus_1_eq_Suc to_polynomial_def) + +lemma cring_coord_rings_coord_ring: +"cring_coord_rings (R[\\<^bsub>n\<^esub>])" + unfolding cring_coord_rings_def + cring_coord_rings_axioms_def coord_ring_def + apply(rule conjI) + unfolding UP_cring_def + apply (metis coord_cring_cring coord_ring_def) + using cring_coord_rings_axioms + unfolding cring_coord_rings_def cring_coord_rings_axioms_def + by (metis coord_ring_def coord_ring_one coord_ring_zero) + +lemma from_univ_poly_monom_inverse: + assumes "i < Suc n" + assumes "a \ carrier (R[\\<^bsub>n\<^esub>])" + shows "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a m)) = up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a m" +proof- + have 0: "up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a m = (to_polynomial (R[\\<^bsub>n\<^esub>]) a) \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub> (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m)" + using UP_cring.monom_as_mult[of "R[\\<^bsub>n\<^esub>]" a m] unfolding UP_ring_def + using UP_cring_def assms coord_cring_cring by blast + have 1 : "(UP_to_IP (R [\\<^bsub>Suc n - 1\<^esub>]) i) (to_polynomial (R[\\<^bsub>n\<^esub>]) a) = ring.indexed_const (R[\\<^bsub>n\<^esub>]) a" + using UP_cring.UP_to_IP_const[of "R [\\<^bsub>Suc n - 1\<^esub>]" a i] unfolding UP_cring_def + by (simp add: assms coord_cring_cring) + have 2: "(from_univ_poly (Suc n) i (to_polynomial (R[\\<^bsub>n\<^esub>]) a)) + = pre_to_univ_poly_inv (Suc n) i (ring.indexed_const (R[\\<^bsub>n\<^esub>]) a)" + unfolding from_univ_poly_def using 1 + by (metis comp_apply) + have 3: "from_univ_poly (Suc n) i (to_polynomial (R [\\<^bsub>n\<^esub>]) a) = pre_to_univ_poly_inv_hom (Suc n) i a" + using pre_to_univ_poly_inv_const[of i n a] assms 2 + by presburger + have 4: "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (to_polynomial (R [\\<^bsub>n\<^esub>]) a)) = + IP_to_UP i ((pre_to_univ_poly (Suc n) i) (pre_to_univ_poly_inv_hom (Suc n) i a))" + using 3 unfolding to_univ_poly_def from_univ_poly_def + by (metis comp_apply) + have 5: "(pre_to_univ_poly (Suc n) i) (pre_to_univ_poly_inv (Suc n) i (ring.indexed_const (R[\\<^bsub>n\<^esub>]) a)) = (ring.indexed_const (R[\\<^bsub>n\<^esub>]) a)" + using assms(1) assms(2) pre_to_univ_poly_inverse' by blast + have "(to_univ_poly (Suc n) i) (from_univ_poly (Suc n) i (to_polynomial (R[\\<^bsub>n\<^esub>]) a)) = IP_to_UP i (ring.indexed_const (R[\\<^bsub>n\<^esub>]) a)" + unfolding to_univ_poly_def + by (metis "2" "5" comp_apply) + hence 6: "(to_univ_poly (Suc n) i) (from_univ_poly (Suc n) i (to_polynomial (R[\\<^bsub>n\<^esub>]) a)) = to_polynomial (R[\\<^bsub>n\<^esub>]) a" + using UP_cring.IP_to_UP_indexed_const[of "R[\\<^bsub>n\<^esub>]"] + by (smt UP_cring_def assms(2) coord_cring_cring) + have 7: "(to_univ_poly (Suc n) i) (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m)) = up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m" + proof- + have 70: "pvar (R [\\<^bsub>n\<^esub>]) i [^]\<^bsub>Pring (R [\\<^bsub>n\<^esub>]) {i}\<^esub> m \ carrier (Pring (R [\\<^bsub>n\<^esub>]) {i})" + using Cring_Multivariable_Poly.pvar_closed[of "R[\\<^bsub>n\<^esub>]" i "{i}"] monoid.nat_pow_closed[of "R[\\<^bsub>n\<^esub>]"] + by (meson MP.Pring_is_monoid coord_cring_cring equalityD2 insert_subset monoid.nat_pow_closed) + have 71: "(UP_to_IP (R [\\<^bsub>Suc n - 1\<^esub>]) i) (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> m) = + (pvar (R[\\<^bsub>n\<^esub>]) i)[^]\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub>m" + using 70 UP_cring.UP_to_IP_monom[of "R[\\<^bsub>n\<^esub>]" "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>" i m ] cring.Pring_smult_one[of "R[\\<^bsub>n\<^esub>]" "pvar (R [\\<^bsub>n\<^esub>]) i [^]\<^bsub>Pring (R [\\<^bsub>n\<^esub>]) {i}\<^esub> m" "{i}"] + unfolding UP_cring_def + using MP.one_closed coord_cring_cring diff_Suc_1 by presburger + hence 72: "from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m) = + pre_to_univ_poly_inv (Suc n) i ((pvar (R[\\<^bsub>n\<^esub>]) i)[^]\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub>m)" + unfolding from_univ_poly_def + using comp_apply[of "pre_to_univ_poly_inv (Suc n) i" "UP_to_IP (R [\\<^bsub>Suc n - 1\<^esub>]) i" "up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> m"] + by presburger + have 73: " pre_to_univ_poly_inv (Suc n) i \ ring_hom (Pring (R [\\<^bsub>n\<^esub>]) {i}) (R [\\<^bsub>Suc n\<^esub>]) " + using pre_to_univ_poly_inv_is_hom[of i n] assms(1) ring_hom_ring.homh by blast + hence 74: "from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m) = (pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>m" + unfolding from_univ_poly_def + using 70 71 72 pre_to_univ_poly_inv_pvar[of i n] + ring_hom_nat_pow[of "(Pring (R [\\<^bsub>n\<^esub>]) {i})" "R [\\<^bsub>Suc n\<^esub>]" "pre_to_univ_poly_inv (Suc n) i" "(pvar (R[\\<^bsub>n\<^esub>]) i)" m] + by (metis MP.Pring_is_ring MP.Pring_var_closed MP.ring_axioms assms(1) from_univ_poly_def singletonI) + hence 75: "(to_univ_poly (Suc n) i) (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m)) + = (to_univ_poly (Suc n) i) ((pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>m)" + by metis + have 76: "pre_to_univ_poly (Suc n) i (pvar R i) = pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i" + using pre_to_univ_poly_is_hom(3)[of i "Suc n" ] assms(1) by blast + have "pre_to_univ_poly (Suc n) i \ ring_hom (R [\\<^bsub>Suc n\<^esub>]) (Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}) " + apply(rule ring_hom_ring.homh) + using pre_to_univ_poly_is_hom(1)[of i "Suc n"] + using assms(1) by blast + hence "pre_to_univ_poly (Suc n) i (pvar R i [^]\<^bsub>R [\\<^bsub>Suc n\<^esub>]\<^esub> m) = pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i [^]\<^bsub>Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}\<^esub> m" + using 76 ring_hom_nat_pow[of "R[\\<^bsub>Suc n\<^esub>]" "Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}" "pre_to_univ_poly (Suc n) i" "pvar R i" m] + by (metis MP.Pring_is_ring MP.ring_axioms assms(1) local.pvar_closed) + hence 77: "(to_univ_poly (Suc n) i) (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m)) + =IP_to_UP i (pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i [^]\<^bsub>Pring (R [\\<^bsub>Suc n - 1\<^esub>]) {i}\<^esub> m)" + unfolding to_univ_poly_def using comp_apply[of "IP_to_UP i" " pre_to_univ_poly (Suc n) i"] + using "74" by presburger + have 78: "IP_to_UP i (pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i) = X_poly (R[\\<^bsub>n\<^esub>])" + using cring.IP_to_UP_var[of "R[\\<^bsub>n\<^esub>]"] + by (simp add: MP.IP_to_UP_var var_to_IP_def) + have 79: "IP_to_UP i \ ring_hom (Pring (R [\\<^bsub>n\<^esub>]) {i}) (UP (R [\\<^bsub>n\<^esub>]))" + using UP_cring.IP_to_UP_ring_hom[of "R[\\<^bsub>n\<^esub>]" i] ring_hom_ring.homh[of "Pring (R [\\<^bsub>n\<^esub>]) {i}"] + unfolding UP_cring_def + using coord_cring_cring by blast + have 80: "pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i \ carrier (Pring (R [\\<^bsub>n\<^esub>]) {i})" + by (metis "76" assms(1) diff_Suc_1 local.pvar_closed pre_to_univ_poly_is_hom(6)) + have 81: "ring (UP (R[\\<^bsub>n\<^esub>]))" + using UP_ring.UP_ring[of "R[\\<^bsub>n\<^esub>]"] unfolding UP_ring_def + using MP.ring_axioms by blast + hence 82: "IP_to_UP i (pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i [^]\<^bsub>Pring (R[\\<^bsub>n\<^esub>]) {i}\<^esub> m) = X_poly (R[\\<^bsub>n\<^esub>]) [^]\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> m" + + using 78 79 80 ring_hom_nat_pow[of "Pring (R [\\<^bsub>n\<^esub>]) {i}" "UP (R [\\<^bsub>n\<^esub>])" "IP_to_UP i" "pvar (R [\\<^bsub>Suc n - 1\<^esub>]) i" m] + by (metis MP.Pring_is_ring) + have 83: "\\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> \\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> X_poly (R [\\<^bsub>n\<^esub>]) [^]\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> m = X_poly (R [\\<^bsub>n\<^esub>]) [^]\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> m" + using UP_ring.UP_smult_one[of "R[\\<^bsub>n\<^esub>]" "X_poly (R [\\<^bsub>n\<^esub>]) [^]\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> m"] + UP_cring.X_closed[of "R[\\<^bsub>n\<^esub>]"] monoid.nat_pow_closed[of "UP (R[\\<^bsub>n\<^esub>])" "X_poly (R[\\<^bsub>n\<^esub>])" m] + unfolding UP_ring_def UP_cring_def + using 81 MP.ring_axioms coord_cring_cring ring.is_monoid by blast + have 84: "X_poly (R[\\<^bsub>n\<^esub>]) [^]\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> m = up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> m" + using 83 UP_cring.monom_rep_X_pow[of "R[\\<^bsub>n\<^esub>]" "\\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>" m] + monoid.nat_pow_closed[of "UP (R[\\<^bsub>n\<^esub>])" "X_poly (R[\\<^bsub>n\<^esub>])" m] 81 + unfolding UP_cring_def + using MP.one_closed coord_cring_cring by presburger + thus ?thesis using 77 + by (metis "82" diff_Suc_1) + qed + have 8: "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a m)) = + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (to_polynomial (R[\\<^bsub>n\<^esub>]) a)) \\<^bsub>UP (R[\\<^bsub>n\<^esub>])\<^esub> + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m))" + proof- + have 80: "to_polynomial (R [\\<^bsub>n\<^esub>]) a \ carrier (UP (R [\\<^bsub>n\<^esub>]))" + using UP_cring.to_poly_closed[of "R[\\<^bsub>n\<^esub>]" a] UP_cring_def assms(2) coord_cring_cring + by blast + have 81: "up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R [\\<^bsub>n\<^esub>]\<^esub> m \ carrier (UP (R [\\<^bsub>n\<^esub>])) " + apply(rule UP_ring.monom_closed[of "R[\\<^bsub>n\<^esub>]"]) unfolding UP_ring_def using MP.one_closed + apply (simp add: MP.ring_axioms) + using MP.one_closed by blast + have 82: "(from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a m)) = + (from_univ_poly (Suc n) i (to_polynomial (R[\\<^bsub>n\<^esub>]) a)) \\<^bsub>(R[\\<^bsub>Suc n\<^esub>])\<^esub> + (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m))" + using 80 81 from_univ_poly_mult[of i n "to_polynomial (R [\\<^bsub>n\<^esub>]) a" "(up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> m)"] 0 + by (metis assms(1) less_Suc_eq_le) + thus ?thesis using to_univ_poly_mult 80 81 + by (metis assms(1) from_univ_poly_closed less_Suc_eq_le) + qed + thus ?thesis + using "0" "6" "7" by metis +qed + +lemma from_univ_poly_inverse: + assumes "i \ n" + assumes "p \ carrier (UP (R[\\<^bsub>n\<^esub>]))" + shows "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i p) = p" +proof(rule UP_ring.poly_induct3[of "R[\\<^bsub>n\<^esub>]"]) + show "UP_ring (R [\\<^bsub>n\<^esub>])" + unfolding UP_ring_def + by (simp add: MP.ring_axioms) + show "p \ carrier (UP (R [\\<^bsub>n\<^esub>]))" + using assms by blast + show "\p q. q \ carrier (UP (R [\\<^bsub>n\<^esub>])) \ + p \ carrier (UP (R [\\<^bsub>n\<^esub>])) \ + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i p) = p \ + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i q) = q \ + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (p \\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> q)) = p \\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> q" + proof- fix p q + assume A: "q \ carrier (UP (R [\\<^bsub>n\<^esub>]))" "p \ carrier (UP (R [\\<^bsub>n\<^esub>]))" + "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i p) = p" + "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i q) = q" + show "to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (p \\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> q)) = p \\<^bsub>UP (R [\\<^bsub>n\<^esub>])\<^esub> q" + using A assms + from_univ_poly_add[of i n p q] + to_univ_poly_add[of i n "from_univ_poly (Suc n) i p" "from_univ_poly (Suc n) i q"] + from_univ_poly_closed[of i n p] from_univ_poly_closed[of i n q] + by presburger + qed + show "\a na. a \ carrier (R [\\<^bsub>n\<^esub>]) \ + to_univ_poly (Suc n) i (from_univ_poly (Suc n) i (up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a na)) = up_ring.monom (UP (R [\\<^bsub>n\<^esub>])) a na" + using from_univ_poly_monom_inverse[of i ] assms(1) le_imp_less_Suc by presburger +qed + +lemma to_univ_poly_eval: + assumes "i < Suc n" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "x \ carrier R" + assumes "as = insert_at_index a x i" + shows "eval_at_point R as p = eval_at_point R a (to_function (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p) (coord_const x))" +proof- + have 0: "pre_to_univ_poly (Suc n) i p \ Pring_set (R[\\<^bsub>n\<^esub>]) {i}" + using assms pre_to_univ_poly_is_hom(1)[of i "Suc n"] unfolding ring_hom_ring_def + unfolding coord_ring_def UP_domain_def coord_ring_def UP_domain_def + by (metis MP.Pring_car coord_ring_def diff_Suc_1 pre_to_univ_poly_is_hom(6)) + have 1: " closed_fun (R[\\<^bsub>n\<^esub>]) (\n. coord_const x)" + using assms(4) R.indexed_const_closed + unfolding coord_ring_def UP_domain_def + by blast + have "(to_function (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p) (coord_const x)) = + to_function (R[\\<^bsub>n\<^esub>]) (IP_to_UP i ((pre_to_univ_poly (Suc n) i) p)) (coord_const x)" + unfolding to_univ_poly_def + unfolding coord_ring_def UP_domain_def + by (metis comp_apply) + then have 2: "(to_function (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p) (coord_const x)) = + (total_eval (R[\\<^bsub>n\<^esub>]) (\ i. coord_const x) (pre_to_univ_poly (Suc n) i p))" + using 0 1 UP_cring.IP_to_UP_poly_eval[of "R[\\<^bsub>n\<^esub>]" + "(pre_to_univ_poly (Suc n) i) p" i "\ i. coord_const x"] + unfolding coord_ring_def UP_cring_def + using assms(4) cring.indexed_const_closed R.Pring_is_cring R.cring_axioms + by smt + then show ?thesis using pre_to_univ_poly_eval[of i n p a x as] + using assms(1) assms(2) assms(3) assms(4) assms(5) by presburger +qed + +text\ + The function \texttt{one\_over\_poly}, introduced in the theory \texttt{Cring\_Poly}, maps a + polynomial $p(x)$ to the unique polynomial $q(x)$ which satisfies the relation + $q(x) = x^n p(1/x)$. This will be used later to show that the function $f(x) = 1/x$ is + semialgebraic over the field $\mathbb{Q}_p$.\ +lemma to_univ_poly_one_over_poly: + assumes "field R" + assumes "i < (Suc n)" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "Q = from_univ_poly (Suc n) i (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p))" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "x \ carrier R" + assumes "x \ \" + assumes "b = insert_at_index a x i" + assumes "c = insert_at_index a (inv x) i" + assumes "N = UP_ring.degree (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)" + shows "Q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + "eval_at_point R b Q = (x[^]N) \ (eval_at_point R c p)" +proof- + have 0: "(to_univ_poly (Suc n) i p) \ carrier (UP (R[\\<^bsub>n\<^esub>]))" + using assms(2) assms(3) less_Suc_eq_le to_univ_poly_closed by blast + have 1: "(UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)) \ carrier (UP (R[\\<^bsub>n\<^esub>]))" + using 0 assms UP_domain_def UP_cring.one_over_poly_closed UP_cring_def coord_cring_cring by blast + show "Q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + using 1 assms from_univ_poly_closed[of i n] less_Suc_eq_le + by blast + have 2: "coord_const x \ Units (R[\\<^bsub>n\<^esub>])" + proof- + + have 20: "inv x \ carrier R" + using assms(1) assms(6) assms(7) field.field_Units by blast + have 21: "x \ (inv x) = \ " + using assms field.field_Units R.Units_r_inv + by blast + have 22: "coord_const x \ carrier (R[\\<^bsub>n\<^esub>])" + using assms(6) R.indexed_const_closed + unfolding coord_ring_def + by blast + have 23: "coord_const (inv x) \ carrier (R[\\<^bsub>n\<^esub>])" + using "20" R.indexed_const_closed + unfolding coord_ring_def +by blast + have 24: "coord_const x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const (inv x) = coord_const (x \ (inv x))" + using assms(6) 20 R.indexed_const_mult unfolding coord_ring_def + by blast + have 25: "coord_const x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const (inv x) = \\<^bsub>coord_ring R n\<^esub>" + unfolding coord_ring_def + by (metis "20" "21" R.Pring_one assms(6) R.indexed_const_mult) + have 26: "coord_const (inv x) \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const x = \\<^bsub>coord_ring R n\<^esub>" + unfolding coord_ring_def + by (metis "21" "22" "23" "24" MP.m_comm R.Pring_one coord_ring_def) + then show ?thesis + using 23 Units_def[of "R[\\<^bsub>n\<^esub>]"] "22" "25" + by blast + qed + have 3: "inv\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (coord_const x) = coord_const (inv x)" + proof- + have 20: "inv x \ carrier R" + using assms(1) assms(6) assms(7) field.field_Units by blast + have 21: "x \ (inv x) = \ " + using assms field.field_Units R.Units_r_inv + by blast + have 22: "coord_const x \ carrier (R[\\<^bsub>n\<^esub>])" + using assms(6) R.indexed_const_closed + unfolding coord_ring_def +by blast + have 23: "coord_const (inv x) \ carrier (R[\\<^bsub>n\<^esub>])" + using "20" R.indexed_const_closed unfolding coord_ring_def + by blast + have 24: "coord_const x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const (inv x) = coord_const (x \ (inv x))" + using assms(6) 20 R.indexed_const_mult unfolding coord_ring_def + by blast + have 25: "coord_const x \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> coord_const (inv x) = \\<^bsub>coord_ring R n\<^esub>" + unfolding coord_ring_def + by (metis "20" "21" R.Pring_one assms(6) R.indexed_const_mult) + show ?thesis + using 22 23 25 R.Pring_is_cring[of "{..\<^bsub>n\<^esub>]"] + unfolding coord_ring_def + + by (metis R.Pring_is_monoid R.Pring_mult_comm R.is_cring) + qed + have 4: "to_function (R[\\<^bsub>n\<^esub>]) (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)) + (coord_const x) = (coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> + (to_function (R[\\<^bsub>n\<^esub>]) ( (to_univ_poly (Suc n) i p)) (coord_const (inv\<^bsub>R\<^esub> x)))" + using 3 assms UP_cring_def UP_cring.one_over_poly_eval[of "R[\\<^bsub>n\<^esub>]" " (to_univ_poly (Suc n) i p)" "coord_const x"] + unfolding coord_ring_def + by (metis "0" "2" MP.Units_closed R.Pring_is_cring UP_cring.to_fun_def coord_ring_def R.is_cring) + have 5: "eval_at_point R a (to_function (R[\\<^bsub>n\<^esub>]) (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)) + (coord_const x)) + = eval_at_point R a ((coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> + (to_function (R[\\<^bsub>n\<^esub>]) ( (to_univ_poly (Suc n) i p)) (coord_const (inv\<^bsub>R\<^esub> x))) ) " + using 4 + by presburger + have 6: "to_univ_poly (Suc n) i Q = (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p))" + using assms from_univ_poly_inverse + by (meson "1" less_Suc_eq_le) + have 7: "eval_at_point R a (to_function (R[\\<^bsub>n\<^esub>]) (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)) + (coord_const x)) = eval_at_point R b Q" + using 6 to_univ_poly_eval[of i n Q a x b] assms \Q \ carrier (R[\\<^bsub>Suc n\<^esub>])\ + by smt + have 8: "(coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N \ carrier (R[\\<^bsub>n\<^esub>])" + using monoid.nat_pow_closed[of "R[\\<^bsub>n\<^esub>]"] + unfolding coord_ring_def + using R.Pring_is_monoid assms(6) R.indexed_const_closed by blast + have 9: "to_function (R[\\<^bsub>n\<^esub>]) ( (to_univ_poly (Suc n) i p)) (coord_const (inv\<^bsub>R\<^esub> x)) + \ carrier (R[\\<^bsub>n\<^esub>])" + proof- + have 91: "to_univ_poly (Suc n) i p \ carrier (UP (R[\\<^bsub>n\<^esub>]))" + by (simp add: "0") + have " coord_const (inv x) \ carrier (R[\\<^bsub>n\<^esub>])" + proof- + have "inv x \ carrier R" + using assms(1) assms(6) assms(7) field.field_Units by blast + then show ?thesis + using R.indexed_const_closed[of "inv x"] assms + unfolding coord_ring_def + + by blast + qed + then show ?thesis + using 91 UP_cring_def[of "R[\\<^bsub>n\<^esub>]" ] UP_cring.to_fun_closed[of "R[\\<^bsub>n\<^esub>]" "to_univ_poly (Suc n) i p" "coord_const (inv\<^bsub>R\<^esub> x)"] + to_univ_poly_closed[of i n p] UP_domain_def[of "R[\\<^bsub>n\<^esub>]"] + unfolding coord_ring_def + using R.Pring_is_cring R.is_cring + by (metis UP_cring.to_fun_def) + qed + have 10: " eval_at_point R b Q = (eval_at_point R a ((coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N)) \ + (eval_at_point R a (to_function (R[\\<^bsub>n\<^esub>]) ( (to_univ_poly (Suc n) i p)) (coord_const (inv\<^bsub>R\<^esub> x))))" + using 7 5 eval_at_point_mult[of a n "(coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N" + "(to_function (R[\\<^bsub>n\<^esub>]) ( (to_univ_poly (Suc n) i p)) (coord_const (inv\<^bsub>R\<^esub> x)))"] + "8" "9" assms(5) + by presburger + have 11: "inv x \ carrier R" + using assms(1) assms(6) assms(7) field.field_Units by blast + have 12: " eval_at_point R b Q = (eval_at_point R a ((coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N)) \ + (eval_at_point R c p)" + using 10 11 to_univ_poly_eval[of i n p a "inv x" c] assms(2) assms(3) assms(5) assms(9) + by presburger + show 12: " eval_at_point R b Q = (x[^]N) \ + (eval_at_point R c p)" + proof- + have 0: "(coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N = coord_const (x[^]N)" + proof(induction N) + case 0 + have 00: "coord_const x [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> (0::nat) = \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>" + using nat_pow_def[of "R[\\<^bsub>n\<^esub>]" _ "(0::nat)"] + unfolding coord_ring_def + by (meson Group.nat_pow_0) + then show ?case + unfolding coord_ring_def + by (metis Group.nat_pow_0 R.Pring_one) + next + case (Suc N) fix N::nat assume IH: "coord_const x [^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> N = coord_const (x [^] N)" + then show ?case + using R.indexed_const_mult Group.nat_pow_Suc Suc.IH assms(6) R.nat_pow_closed + unfolding coord_ring_def + by (metis ) + qed + have 1: "(eval_at_point R a ((coord_const x)[^]\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub>N)) = x[^]N" + using 0 + by (metis assms(5) assms(6) eval_at_point_const R.nat_pow_closed) + show ?thesis using 0 1 "12" + by presburger + qed +qed + +lemma to_univ_poly_one_over_poly': + assumes "field R" + assumes "i < (Suc n)" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "Q = from_univ_poly (Suc n) i (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p))" + assumes "a \ carrier (R\<^bsup>n\<^esup>)" + assumes "x \ carrier R" + assumes "x \ \" + assumes "b = insert_at_index a x i" + assumes "c = insert_at_index a (inv x) i" + assumes "N = UP_ring.degree (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)" + assumes "q = (pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>(k::nat)\\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> Q" + shows "q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + "eval_at_point R b q = (x[^](N + k)) \ (eval_at_point R c p)" +proof- + have 0: "(pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>k \ carrier (R[\\<^bsub>Suc n\<^esub>])" + using pvar_closed[of i "Suc n"] monoid.nat_pow_closed[] + unfolding coord_ring_def + by (metis R.Pring_is_monoid assms(2)) + have 1: "b \ carrier (R\<^bsup>Suc n\<^esup>)" + using assms(2) assms(5) assms(6) assms(8) insert_at_index_closed less_Suc_eq_le + by blast + have 11 : "c \ carrier (R\<^bsup>Suc n\<^esup>)" + proof- + have "inv x \ carrier R" + using assms field.field_Units + by blast + then show ?thesis + using assms insert_at_index_closed less_Suc_eq_le + by blast + qed + have 2: "eval_at_point R b q = eval_at_point R b ((pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>(k::nat)) + \ eval_at_point R b Q" + using assms 0 1 unfolding coord_ring_def + by (metis R.Pring_mult coord_ring_def eval_at_point_mult to_univ_poly_one_over_poly(1)) + have 3: "eval_at_point R b ((pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>(k::nat)) = + x[^](k::nat)" + proof(induction k) + case 0 + have T0: "eval_at_point R b ((pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>(0::nat)) = + eval_at_point R b (\\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>)" + using nat_pow_def[of "R[\\<^bsub>Suc n\<^esub>]" "pvar R i" "0::nat"] + by (metis Group.nat_pow_0) + then show ?case + by (metis "1" assms(2) eval_at_point_nat_pow R.nat_pow_0 local.pvar_closed) + next + case (Suc k) fix k::nat + assume IH: "eval_at_poly R (pvar R i [^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> k) b = x [^] k " + have 0: "eval_at_poly R (pvar R i) b = b!i" + using eval_pvar[of i "Suc n"] assms "1" + by blast + have "length a = n" + using assms(5) cartesian_power_car_memE by blast + then have "eval_at_poly R (pvar R i) b = x" + using 0 assms(8) insert_at_index_eq[of i a x] + by (metis assms(2) less_Suc_eq_le) + then show ?case + using "1" assms(2) eval_at_point_nat_pow local.pvar_closed + by blast + qed + have 4: "eval_at_point R b Q = (x[^]N) \ (eval_at_point R c p)" + using to_univ_poly_one_over_poly(2)[of i n p Q a x b c N] assms(1) assms(10) assms(2) + assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9) + by blast + have 5: "eval_at_point R b q = x[^](k::nat) \ ((x[^]N) \ (eval_at_point R c p))" + using 4 3 2 + by presburger + show 6: "eval_at_point R b q = x[^](N + k) \ (eval_at_point R c p)" + proof- + + have 60: "x[^](k::nat) \ carrier R" + using assms(6) by blast + have 61: "x[^]N \ carrier R" + using assms(6) by blast + have 62: "eval_at_point R c p \ carrier R" + using eval_at_point_closed[of c "Suc n" p] \c \ carrier (R\<^bsup>Suc n\<^esup>)\ assms(3) + by blast + show ?thesis using 5 60 61 62 + by (metis assms(6) R.m_assoc R.m_comm R.nat_pow_mult) + qed + show "q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + using assms + unfolding coord_ring_def + using 0 R.Pring_mult_closed to_univ_poly_one_over_poly(1) + by (metis coord_ring_def) + +qed + +lemma to_univ_poly_one_over_poly'': + assumes "field R" + assumes "i < (Suc n)" + assumes "p \ carrier (R[\\<^bsub>Suc n\<^esub>])" + assumes "N \ UP_ring.degree (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)" + shows "\ q \ carrier (R[\\<^bsub>Suc n\<^esub>]). ( \ x \ carrier R - {\}. ( \ a \ carrier (R\<^bsup>n\<^esup>). + eval_at_point R (insert_at_index a x i) q = (x[^]N) \ (eval_at_point R (insert_at_index a (inv x) i) p)))" +proof- + obtain Q where Q_def: + "Q = from_univ_poly (Suc n) i (UP_cring.one_over_poly (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p))" + by blast + obtain k where k_def: "k = (N - UP_ring.degree (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p))" + by blast + obtain q where q_def: "q = (pvar R i)[^]\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub>(k::nat)\\<^bsub>R[\\<^bsub>Suc n\<^esub>]\<^esub> Q" + by blast + have 0: " ( \ x \ carrier R - {\}.( \ a \ carrier (R\<^bsup>n\<^esup>). + eval_at_point R (insert_at_index a x i) q = (x[^]N) \ (eval_at_point R (insert_at_index a (inv x) i) p)))" + proof fix x + assume A0: " x \ carrier R - {\}" + show " \a\carrier (R\<^bsup>n\<^esup>). eval_at_poly R q (insert_at_index a x i) = x [^] N \ eval_at_poly R p (insert_at_index a (inv x) i)" + proof fix a assume A1: "a \ carrier (R\<^bsup>n\<^esup>)" + obtain l where l_def: "l = UP_ring.degree (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)" + by blast + have "eval_at_poly R q (insert_at_index a x i) = x [^] (l + k) \ eval_at_poly R p (insert_at_index a (inv x) i)" + using assms A1 A0 to_univ_poly_one_over_poly'(2)[of i n p Q a x "insert_at_index a x i" "insert_at_index a (inv x) i" l q k] + Q_def l_def q_def + by blast + then show " eval_at_poly R q (insert_at_index a x i) = x [^] N \ eval_at_poly R p (insert_at_index a (inv x) i)" + using k_def assms l_def add_diff_inverse_nat less_Suc_eq not_less_eq + by (metis diff_diff_cancel diff_less_Suc) + qed + qed + have 1: "q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + proof- + obtain a where a_def: "a = map (\i. \) [(0::nat).. carrier (R\<^bsup>n\<^esup>)" + apply(rule cartesian_power_car_memI') + using a_def + apply (metis Ex_list_of_length coeff_list length_map length_rev) + proof- fix i assume A: "i < n" + then have "a!i = \" + using a_def + by (metis R_list_length length_map map_nth nth_map) + then show "a ! i \ carrier R" + using a_def R.one_closed + by metis + qed + then show "q \ carrier (R[\\<^bsub>Suc n\<^esub>])" + using assms q_def k_def Q_def to_univ_poly_one_over_poly'(1)[of i n p Q a \ _ _ "deg (R[\\<^bsub>n\<^esub>]) + (to_univ_poly (Suc n) i p)" q "N -deg (R[\\<^bsub>n\<^esub>]) (to_univ_poly (Suc n) i p)" ] + using one_closed local.one_neq_zero by blast + qed + show ?thesis + using 0 1 by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +section\Restricted Inverse Images and Complements\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + This section introduces some versions of basic set operations for extensional functions and sets. + We would like a version of the inverse image which intersects the inverse image of a function + with the set \texttt{carrier }$(R^n)$, and a version of the complement of a set which takes the + comeplement relative to \texttt{carrier }$(R^n)$. These will have to be defined in parametrized + families, with one such object for each natural number $n$.\ +definition evimage (infixr "\\" 90) where +"evimage n f S = ((f -` S) \ carrier (R\<^bsup>n\<^esup>))" + +definition euminus_set :: "nat \ 'a list set \ 'a list set" ("_ \<^sup>c\" 70) where +"S\<^sup>c\<^bsub>n\<^esub> = carrier (R\<^bsup>n\<^esup>) - S" + +lemma extensional_vimage_closed: +"f \\<^bsub>n\<^esub> S \ carrier (R\<^bsup>n\<^esup>)" + unfolding evimage_def by blast + +subsection \Inverse image of a function\ + +lemma evimage_eq [simp]: "a \ f \\<^bsub>n\<^esub> B \ a \ carrier (R\<^bsup>n\<^esup>) \ f a \ B" + unfolding evimage_def + by blast + +lemma evimage_singleton_eq: "a \ f \\<^bsub>n\<^esub> {b} \ a \ carrier (R\<^bsup>n\<^esup>) \ f a = b" + unfolding evimage_def + by blast + +lemma evimageI [intro]: "a \ carrier (R\<^bsup>n\<^esup>) \ f a = b \ b \ B \ a \ f \\<^bsub>n\<^esub> B" + unfolding vimage_def + using evimage_eq by blast + +lemma evimageI2: "a \ carrier (R\<^bsup>n\<^esup>) \ f a \ A \ a \ f \\<^bsub>n\<^esub> A" + unfolding vimage_def by fast + +lemma evimageE [elim!]: "a \ f \\<^bsub>n\<^esub> B \ (\x. f a = x \ x \ B \ p) \ p" + unfolding evimage_def + by blast + +lemma evimageD: "a \ f\\<^bsub>n\<^esub> A \ f a \ A" + unfolding vimage_def by fast + +lemma evimage_empty [simp]: "f \\<^bsub>n\<^esub> {} = {}" + by blast + +lemma evimage_Compl: + assumes "f \ carrier (R\<^bsup>n\<^esup>) \ carrier (R\<^bsup>m\<^esup>)" + shows "(f \\<^bsub>n\<^esub>(A\<^sup>c\<^bsub>m\<^esub>)) = ((f -` A)\<^sup>c\<^bsub>n\<^esub>) " +proof- + have "(f \\<^bsub>n\<^esub>(A\<^sup>c\<^bsub>m\<^esub>)) = ((f -` (carrier (R\<^bsup>m\<^esup>)) - (f -` A))) \ carrier (R\<^bsup>n\<^esup>)" + unfolding evimage_def euminus_set_def by blast + hence 0: "(f \\<^bsub>n\<^esub>(A\<^sup>c\<^bsub>m\<^esub>)) = (f -` (carrier (R\<^bsup>m\<^esup>)) \ carrier (R\<^bsup>n\<^esup>)) - (f -` A)" + by (simp add: Int_Diff Int_commute) + have "(f -` (carrier (R\<^bsup>m\<^esup>)) \ carrier (R\<^bsup>n\<^esup>)) = carrier (R\<^bsup>n\<^esup>)" + proof + show "f -` carrier (R\<^bsup>m\<^esup>) \ carrier (R\<^bsup>n\<^esup>) \ carrier (R\<^bsup>n\<^esup>)" + by auto + show "carrier (R\<^bsup>n\<^esup>) \ f -` carrier (R\<^bsup>m\<^esup>) \ carrier (R\<^bsup>n\<^esup>)" + using assms by blast + qed + thus ?thesis using 0 + by (simp add: euminus_set_def) +qed + +lemma evimage_Un [simp]: "f \\<^bsub>n\<^esub> (A \ B) = (f \\<^bsub>n\<^esub> A) \ (f \\<^bsub>n\<^esub> B)" + unfolding evimage_def by blast + +lemma evimage_Int [simp]: "f \\<^bsub>n\<^esub> (A \ B) = (f \\<^bsub>n\<^esub> A) \ (f \\<^bsub>n\<^esub> B)" + unfolding evimage_def by blast + +lemma evimage_Collect_eq [simp]: "f \\<^bsub>n\<^esub> Collect p = {y \ carrier (R\<^bsup>n\<^esup>). p (f y)}" + unfolding evimage_def by blast + +lemma evimage_Collect: "(\x. x \ carrier (R\<^bsup>n\<^esup>) \ p (f x) = Q x) \ f \\<^bsub>n\<^esub> (Collect p) = Collect Q \ carrier (R\<^bsup>n\<^esup>)" + unfolding evimage_def by blast + +lemma evimage_insert: "f \\<^bsub>n\<^esub> (insert a B) = (f \\<^bsub>n\<^esub> {a}) \ (f \\<^bsub>n\<^esub> B)" + \ \NOT suitable for rewriting because of the recurrence of \{a}\.\ + unfolding evimage_def by blast + +lemma evimage_Diff: "f \\<^bsub>n\<^esub> (A - B) = (f \\<^bsub>n\<^esub> A) - (f \\<^bsub>n\<^esub> B)" + unfolding evimage_def by blast + +lemma evimage_UNIV [simp]: "f \\<^bsub>n\<^esub> UNIV = carrier (R\<^bsup>n\<^esup>)" + unfolding evimage_def by blast + +lemma evimage_mono: "A \ B \ f \\<^bsub>n\<^esub> A \ f \\<^bsub>n\<^esub> B" + \ \monotonicity\ + unfolding evimage_def by blast + +lemma evimage_image_eq: "(f \\<^bsub>n\<^esub> (f ` A)) = {y \ carrier (R\<^bsup>n\<^esup>). \x\A. f x = f y}" + unfolding evimage_def by (blast intro: sym) + +lemma image_evimage_subset: "f ` (f \\<^bsub>n\<^esub> A) \ A" + by blast + +lemma image_evimage_eq [simp]: "f ` (f \\<^bsub>n\<^esub> A) = A \ (f ` carrier (R\<^bsup>n\<^esup>))" + unfolding evimage_def by blast + +lemma image_subset_iff_subset_evimage: "A \ carrier (R\<^bsup>n\<^esup>) \ f ` A \ B \ A \ f \\<^bsub>n\<^esub> B" + by blast + +lemma evimage_const [simp]: "((\x. c) \\<^bsub>n\<^esub> A) = (if c \ A then carrier (R\<^bsup>n\<^esup>) else {})" + unfolding evimage_def using vimage_const[of c A] + by (smt Int_commute inf_bot_right inf_top.right_neutral) + +lemma evimage_if [simp]: "((\x. if x \ B then c else d) \\<^bsub>n\<^esub> A) = + (if c \ A then (if d \ A then carrier (R\<^bsup>n\<^esup>) else B \ carrier (R\<^bsup>n\<^esup>) ) + else if d \ A then B\<^sup>c\<^bsub>n\<^esub> else {})" +unfolding evimage_def euminus_set_def using vimage_if[of B c d A] + by (metis Diff_Compl Diff_UNIV Diff_empty Int_commute double_compl) + +lemma evimage_inter_cong: "(\ w. w \ S \ f w = g w) \ f \\<^bsub>n\<^esub> y \ S = g \\<^bsub>n\<^esub> y \ S" +unfolding evimage_def + by (smt Int_assoc Int_commute vimage_inter_cong) + +lemma evimage_ident [simp]: "(\x. x) \\<^bsub>n\<^esub> Y = Y \ carrier (R\<^bsup>n\<^esup>)" +unfolding evimage_def + by blast + + +end + + + + +end diff --git a/thys/Padic_Field/document/root.bib b/thys/Padic_Field/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Padic_Field/document/root.bib @@ -0,0 +1,62 @@ + +@Book{dummit2004abstract, + author = {Dummit, David}, + title = {Abstract algebra}, + publisher = {John Wiley \& Sons, Inc}, + year = {2004}, + address = {Hoboken, NJ}, + isbn = {0471433349} + } + +@Book{engler2005valued, + author = {Engler, Antonio}, + title = {Valued fields}, + publisher = {Springer}, + year = {2005}, + address = {Berlin New York}, + isbn = {354024221X} + } + +@misc{keithconrad, title={Hensel's Lemma}, url={https://kconrad.math.uconn.edu/blurbs/gradnumthy/hensel.pdf}, author={Conrad, Keith}} + +@inproceedings{Thi, +author = {Lewis, Robert Y.}, +title = {A Formal Proof of Hensel's Lemma over the p-Adic Integers}, +year = {2019}, +isbn = {9781450362221}, +publisher = {Association for Computing Machinery}, +address = {New York, NY, USA}, +url = {https://doi.org/10.1145/3293880.3294089}, +doi = {10.1145/3293880.3294089}, +abstract = {The field of p-adic numbers ℚp and the ring of p-adic integers ℤp are essential constructions of modern number theory. Hensel’s lemma, described by Gouv\^{e}a as the “most important algebraic property of the p-adic numbers,” shows the existence of roots of polynomials over ℤp provided an initial seed point. The theorem can be proved for the p-adics with significantly weaker hypotheses than for general rings. We construct ℚp and ℤp in the Lean proof assistant, with various associated algebraic properties, and formally prove a strong form of Hensel’s lemma. The proof lies at the intersection of algebraic and analytic reasoning and demonstrates how the Lean mathematical library handles such a heterogeneous topic.}, +booktitle = {Proceedings of the 8th ACM SIGPLAN International Conference on Certified Programs and Proofs}, +pages = {15–26}, +numpages = {12}, +keywords = {Hensel's lemma, formal proof, Lean, p-adic}, +location = {Cascais, Portugal}, +series = {CPP 2019} +} + +@article{10.2307/2274477, + ISSN = {00224812}, + URL = {http://www.jstor.org/stable/2274477}, + author = {Johan Pas}, + journal = {The Journal of Symbolic Logic}, + number = {3}, + pages = {1125--1129}, + publisher = {Association for Symbolic Logic}, + title = {On the Angular Component Map Modulo P}, + volume = {55}, + year = {1990} +} + +@article{denef1986, +author = {Denef, Jan}, +journal = {Journal für die reine und angewandte Mathematik}, +keywords = {rationality of Poincaré series; Macintyre's theorem; elimination of quantifiers; p-adic fields; cell decomposition theorem}, +pages = {154-166}, +title = {p-adic Semi-Algebraic Sets and Cell Decomposition.}, +url = {http://eudml.org/doc/152854}, +volume = {369}, +year = {1986}, +} diff --git a/thys/Padic_Field/document/root.tex b/thys/Padic_Field/document/root.tex new file mode 100755 --- /dev/null +++ b/thys/Padic_Field/document/root.tex @@ -0,0 +1,64 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym, amsmath, amssymb, amsfonts} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{$p$-adic Fields} +\author{Aaron Crighton} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + + +\begin{abstract} +We formalize the fields $\mathbb{Q}_p$ of $p$-adic numbers within the framework of the HOL-Algebra library. The $p$-adic field is defined simply as the fraction field of the ring of $p$-adic integers. The valuation, and basic topological properties of $\mathbb{Q}_p$ are developed, including deducing Hensel's Lemma for $\mathbb{Q}_p$ from the same theorem for $\mathbb{Z}_p$. The theory of semialgebraic subsets of $\mathbb{Q}_p^n$ and semialgebraic functions is also developed, as outlined in \cite{denef1986}. In order to formulate these results, general theory about multivariable polynomial rings and cartesian powers of a ring must also be developed. This work is done with a view to formalizing the proof in \cite{denef1986} of Macintyre's quantifier elimination theorem for semialgebraic subsets of $\mathbb{Q}_p^n$. +\end{abstract} +% generated text of all theories + +\input{session} + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,704 +1,705 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRYSTALS-Kyber CRDT CSP_RefTK CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK DPRM_Theorem DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier FO_Theory_Rewriting Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaNet Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp LP_Duality Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multiset_Ordering_NPC Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker Package_logic PAL PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints +Padic_Field Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quasi_Borel_Spaces Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions SCC_Bloemen_Sequential Schutz_Spacetime Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sophomores_Dream Solidity Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith VYDRA_MDL WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL