diff --git a/thys/Padic_Ints/Cring_Poly.thy b/thys/Padic_Ints/Cring_Poly.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Cring_Poly.thy @@ -0,0 +1,7052 @@ +theory Cring_Poly + imports "HOL-Algebra.UnivPoly" "HOL-Algebra.Subrings" Function_Ring +begin + +text\ + This theory extends the material in \<^theory>\HOL-Algebra.UnivPoly\. The main additions are + material on Taylor expansions of polynomials and polynomial derivatives, and various applications + of the universal property of polynomial evaluation. These include construing polynomials as + functions from the base ring to itself, composing one polynomial with another, and extending + homomorphisms between rings to homomoprhisms of their polynomial rings. These formalizations + are necessary components of the proof of Hensel's lemma for $p$-adic integers, and for the + proof of $p$-adic quantifier elimination. \ + +lemma(in ring) ring_hom_finsum: + assumes "h \ ring_hom R S" + assumes "ring S" + assumes "finite I" + assumes "F \ I \ carrier R" + shows "h (finsum R F I) = finsum S (h \ F) I" +proof- + have I: "(h \ ring_hom R S \ F \ I \ carrier R) \ h (finsum R F I) = finsum S (h \ F) I" + apply(rule finite_induct, rule assms) + using assms ring_hom_zero[of h R S] + apply (metis abelian_group_def abelian_monoid.finsum_empty is_ring ring_def) + proof(rule) + fix A a + assume A: "finite A" "a \ A" "h \ ring_hom R S \ F \ A \ carrier R \ + h (finsum R F A) = finsum S (h \ F) A" "h \ ring_hom R S \ F \ insert a A \ carrier R" + have 0: "h \ ring_hom R S \ F \ A \ carrier R " + using A by auto + have 1: "h (finsum R F A) = finsum S (h \ F) A" + using A 0 by auto + have 2: "abelian_monoid S" + using assms ring_def abelian_group_def by auto + have 3: "h (F a \ finsum R F A) = h (F a) \\<^bsub>S\<^esub> (finsum S (h \ F) A) " + using ring_hom_add assms finsum_closed 1 A(4) by fastforce + have 4: "finsum R F (insert a A) = F a \ finsum R F A" + using finsum_insert[of A a F] A assms by auto + have 5: "finsum S (h \ F) (insert a A) = (h \ F) a \\<^bsub>S\<^esub> finsum S (h \ F) A" + apply(rule abelian_monoid.finsum_insert[of S A a "h \ F"]) + apply (simp add: "2") + apply(rule A) + apply(rule A) + using ring_hom_closed A "0" apply fastforce + using A ring_hom_closed by auto + show "h (finsum R F (insert a A)) = + finsum S (h \ F) (insert a A)" + unfolding 4 5 3 by auto + qed + thus ?thesis using assms by blast +qed + +lemma(in ring) ring_hom_a_inv: + assumes "ring S" + assumes "h \ ring_hom R S" + assumes "b \ carrier R" + shows "h (\ b) = \\<^bsub>S\<^esub> h b" +proof- + have "h b \\<^bsub>S\<^esub> h (\ b) = \\<^bsub>S\<^esub>" + by (metis (no_types, hide_lams) abelian_group.a_inv_closed assms(1) assms(2) assms(3) + is_abelian_group local.ring_axioms r_neg ring_hom_add ring_hom_zero) + then show ?thesis + by (metis (no_types, lifting) abelian_group.minus_equality add.inv_closed assms(1) + assms(2) assms(3) ring.is_abelian_group ring.ring_simprules(10) ring_hom_closed) +qed + +lemma(in ring) ring_hom_minus: + assumes "ring S" + assumes "h \ ring_hom R S" + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "h (a \ b) = h a \\<^bsub>S\<^esub> h b" + using assms ring_hom_add[of h R S a "\\<^bsub>R\<^esub> b"] + unfolding a_minus_def + using ring_hom_a_inv[of S h b] by auto + +lemma ring_hom_nat_pow: + assumes "ring R" + assumes "ring S" + assumes "h \ ring_hom R S" + assumes "a \ carrier R" + shows "h (a[^]\<^bsub>R\<^esub>(n::nat)) = (h a)[^]\<^bsub>S\<^esub>(n::nat)" + using assms by (simp add: ring_hom_ring.hom_nat_pow ring_hom_ringI2) + +lemma (in ring) Units_not_right_zero_divisor: + assumes "a \ Units R" + assumes "b \ carrier R" + assumes "a \ b = \" + shows "b = \" +proof- + have "inv a \ a \ b = \ " + using assms Units_closed Units_inv_closed r_null m_assoc[of "inv a" a b] by presburger + thus ?thesis using assms + by (metis Units_l_inv l_one) +qed + +lemma (in ring) Units_not_left_zero_divisor: + assumes "a \ Units R" + assumes "b \ carrier R" + assumes "b \ a = \" + shows "b = \" +proof- + have "b \ (a \ inv a) = \ " + using assms Units_closed Units_inv_closed l_null m_assoc[of b a"inv a"] by presburger + thus ?thesis using assms + by (metis Units_r_inv r_one) +qed + +lemma (in cring) finsum_remove: + assumes "\i. i \ Y \ f i \ carrier R" + assumes "finite Y" + assumes "i \ Y" + shows "finsum R f Y = f i \ finsum R f (Y - {i})" +proof- + have "finsum R f (insert i (Y - {i})) = f i \ finsum R f (Y - {i})" + apply(rule finsum_insert) + using assms apply blast apply blast using assms apply blast + using assms by blast + thus ?thesis using assms + by (metis insert_Diff) +qed + + +type_synonym degree = nat +text\The composition of two ring homomorphisms is a ring homomorphism\ +lemma ring_hom_compose: + assumes "ring R" + assumes "ring S" + assumes "ring T" + assumes "h \ ring_hom R S" + assumes "g \ ring_hom S T" + assumes "\c. c \ carrier R \ f c = g (h c)" + shows "f \ ring_hom R T" +proof(rule ring_hom_memI) + show "\x. x \ carrier R \ f x \ carrier T" + using assms by (metis ring_hom_closed) + show " \x y. x \ carrier R \ y \ carrier R \ f (x \\<^bsub>R\<^esub> y) = f x \\<^bsub>T\<^esub> f y" + proof- + fix x y + assume A: "x \ carrier R" "y \ carrier R" + show "f (x \\<^bsub>R\<^esub> y) = f x \\<^bsub>T\<^esub> f y" + proof- + have "f (x \\<^bsub>R\<^esub> y) = g (h (x \\<^bsub>R\<^esub> y))" + by (simp add: A(1) A(2) assms(1) assms(6) ring.ring_simprules(5)) + then have "f (x \\<^bsub>R\<^esub> y) = g ((h x) \\<^bsub>S\<^esub> (h y))" + using A(1) A(2) assms(4) ring_hom_mult by fastforce + then have "f (x \\<^bsub>R\<^esub> y) = g (h x) \\<^bsub>T\<^esub> g (h y)" + using A(1) A(2) assms(4) assms(5) ring_hom_closed ring_hom_mult by fastforce + then show ?thesis + by (simp add: A(1) A(2) assms(6)) + qed + qed + show "\x y. x \ carrier R \ y \ carrier R \ f (x \\<^bsub>R\<^esub> y) = f x \\<^bsub>T\<^esub> f y" + proof- + fix x y + assume A: "x \ carrier R" "y \ carrier R" + show "f (x \\<^bsub>R\<^esub> y) = f x \\<^bsub>T\<^esub> f y" + proof- + have "f (x \\<^bsub>R\<^esub> y) = g (h (x \\<^bsub>R\<^esub> y))" + by (simp add: A(1) A(2) assms(1) assms(6) ring.ring_simprules(1)) + then have "f (x \\<^bsub>R\<^esub> y) = g ((h x) \\<^bsub>S\<^esub> (h y))" + using A(1) A(2) assms(4) ring_hom_add by fastforce + then have "f (x \\<^bsub>R\<^esub> y) = g (h x) \\<^bsub>T\<^esub> g (h y)" + by (metis (no_types, hide_lams) A(1) A(2) assms(4) assms(5) ring_hom_add ring_hom_closed) + then show ?thesis + by (simp add: A(1) A(2) assms(6)) + qed + qed + show "f \\<^bsub>R\<^esub> = \\<^bsub>T\<^esub>" + by (metis assms(1) assms(4) assms(5) assms(6) ring.ring_simprules(6) ring_hom_one) +qed + + +(**************************************************************************************************) +(**************************************************************************************************) +section\Basic Notions about Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) +context UP_ring +begin + + +text\rings are closed under monomial terms\ +lemma monom_term_car: + assumes "c \ carrier R" + assumes "x \ carrier R" + shows "c \ x[^](n::nat) \ carrier R" + using assms monoid.nat_pow_closed + by blast + +text\Univariate polynomial ring over R\ + +lemma P_is_UP_ring: +"UP_ring R" + by (simp add: UP_ring_axioms) + +text\Degree function\ +abbreviation(input) degree where +"degree f \ deg R f" + +lemma UP_car_memI: + assumes "\n. n > k \ p n = \" + assumes "\n. p n \ carrier R" + shows "p \ carrier P" +proof- + have "bound \ k p" + by (simp add: assms(1) bound.intro) + then show ?thesis + by (metis (no_types, lifting) P_def UP_def assms(2) mem_upI partial_object.select_convs(1)) +qed + +lemma(in UP_cring) UP_car_memI': + assumes "\x. g x \ carrier R" + assumes "\x. x > k \ g x = \" + shows "g \ carrier (UP R)" +proof- + have "bound \ k g" + using assms unfolding bound_def by blast + then show ?thesis + using P_def UP_car_memI assms(1) by blast +qed + +lemma(in UP_cring) UP_car_memE: + assumes "g \ carrier (UP R)" + shows "\x. g x \ carrier R" + "\x. x > (deg R g) \ g x = \" + using P_def assms UP_def[of R] apply (simp add: mem_upD) + using assms UP_def[of R] up_def[of R] + by (smt R.ring_axioms UP_ring.deg_aboveD UP_ring.intro partial_object.select_convs(1) restrict_apply up_ring.simps(2)) + +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Lemmas About Coefficients\ + (**************************************************************************************************) + (**************************************************************************************************) + +context UP_ring +begin +text\The goal here is to reduce dependence on the function coeff from Univ\_Poly, in favour of using +a polynomial itself as its coefficient function.\ + +lemma coeff_simp: + assumes "f \ carrier P" + shows "coeff (UP R) f = f " +proof fix x show "coeff (UP R) f x = f x" + using assms P_def UP_def[of R] by auto +qed + +text\Coefficients are in R\ + +lemma cfs_closed: + assumes "f \ carrier P" + shows "f n \ carrier R" + using assms coeff_simp[of f] P_def coeff_closed + by fastforce + +lemma cfs_monom: + "a \ carrier R ==> (monom P a m) n = (if m=n then a else \)" +using coeff_simp P_def coeff_monom monom_closed by auto + +lemma cfs_zero [simp]: "\\<^bsub>P\<^esub> n = \" + using P_def UP_zero_closed coeff_simp coeff_zero by auto + +lemma cfs_one [simp]: "\\<^bsub>P\<^esub> n = (if n=0 then \ else \)" + by (metis P_def R.one_closed UP_ring.cfs_monom UP_ring_axioms monom_one) + +lemma cfs_smult [simp]: + "[| a \ carrier R; p \ carrier P |] ==> (a \\<^bsub>P\<^esub> p) n = a \ p n" + using P_def UP_ring.coeff_simp UP_ring_axioms UP_smult_closed coeff_smult by fastforce + +lemma cfs_add [simp]: + "[| p \ carrier P; q \ carrier P |] ==> (p \\<^bsub>P\<^esub> q) n = p n \ q n" + by (metis P.add.m_closed P_def UP_ring.coeff_add UP_ring.coeff_simp UP_ring_axioms) + +lemma cfs_a_inv [simp]: + assumes R: "p \ carrier P" + shows "(\\<^bsub>P\<^esub> p) n = \ (p n)" + using P.add.inv_closed P_def UP_ring.coeff_a_inv UP_ring.coeff_simp UP_ring_axioms assms + by fastforce + +lemma cfs_minus [simp]: + "[| p \ carrier P; q \ carrier P |] ==> (p \\<^bsub>P\<^esub> q) n = p n \ q n" + using P.minus_closed P_def coeff_minus coeff_simp by auto + +lemma cfs_monom_mult_r: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "(monom P a n \\<^bsub>P\<^esub> p) (k + n) = a \ p k" + using coeff_monom_mult assms P.m_closed P_def coeff_simp monom_closed by auto + +lemma(in UP_cring) cfs_monom_mult_l: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "(p \\<^bsub>P\<^esub> monom P a n) (k + n) = a \ p k" + using UP_m_comm assms(1) assms(2) cfs_monom_mult_r by auto + +lemma(in UP_cring) cfs_monom_mult_l': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "m \ n" + shows "(f \\<^bsub>P\<^esub> (monom P a n)) m = a \ (f (m - n))" + using cfs_monom_mult_l[of f a n "m-n"] assms + by simp + +lemma(in UP_cring) cfs_monom_mult_r': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "m \ n" + shows "((monom P a n) \\<^bsub>P\<^esub> f) m = a \ (f (m - n))" + using cfs_monom_mult_r[of f a n "m-n"] assms + by simp +end + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Degree Bound Lemmas\ + (**************************************************************************************************) + (**************************************************************************************************) + +context UP_ring +begin + +lemma bound_deg_sum: + assumes " f \ carrier P" + assumes "g \ carrier P" + assumes "degree f \ n" + assumes "degree g \ n" + shows "degree (f \\<^bsub>P\<^esub> g) \ n" + using P_def UP_ring_axioms assms(1) assms(2) assms(3) assms(4) + by (meson deg_add max.boundedI order_trans) + +lemma bound_deg_sum': + assumes " f \ carrier P" + assumes "g \ carrier P" + assumes "degree f < n" + assumes "degree g < n" + shows "degree (f \\<^bsub>P\<^esub> g) < n" + using P_def UP_ring_axioms assms(1) assms(2) + assms(3) assms(4) + by (metis bound_deg_sum le_neq_implies_less less_imp_le_nat not_less) + +lemma equal_deg_sum: + assumes " f \ carrier P" + assumes "g \ carrier P" + assumes "degree f < n" + assumes "degree g = n" + shows "degree (f \\<^bsub>P\<^esub> g) = n" +proof- + have 0: "degree (f \\<^bsub>P\<^esub> g) \n" + using assms bound_deg_sum + P_def UP_ring_axioms by auto + show "degree (f \\<^bsub>P\<^esub> g) = n" + proof(rule ccontr) + assume "degree (f \\<^bsub>P\<^esub> g) \ n " + then have 1: "degree (f \\<^bsub>P\<^esub> g) < n" + using 0 by auto + have 2: "degree (\\<^bsub>P\<^esub> f) < n" + using assms by simp + have 3: "g = (f \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> (\\<^bsub>P\<^esub> f)" + using assms + by (simp add: P.add.m_comm P.r_neg1) + then show False using 1 2 3 assms + by (metis UP_a_closed UP_a_inv_closed deg_add leD le_max_iff_disj) + qed +qed + +lemma equal_deg_sum': + assumes "f \ carrier P" + assumes "g \ carrier P" + assumes "degree g < n" + assumes "degree f = n" + shows "degree (f \\<^bsub>P\<^esub> g) = n" + using P_def UP_a_comm UP_ring.equal_deg_sum UP_ring_axioms assms(1) assms(2) assms(3) assms(4) + by fastforce + +lemma degree_of_sum_diff_degree: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree q < degree p" + shows "degree (p \\<^bsub>P\<^esub> q) = degree p" + by(rule equal_deg_sum', auto simp: assms) + +lemma degree_of_difference_diff_degree: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree q < degree p" + shows "degree (p \\<^bsub>P\<^esub> q) = degree p" +proof- + have A: "(p \\<^bsub>P\<^esub> q) = p \\<^bsub>P\<^esub> (\\<^bsub>P\<^esub> q)" + by (simp add: P.minus_eq) + have "degree (\\<^bsub>P\<^esub> q) = degree q " + by (simp add: assms(2)) + then show ?thesis + using assms A + by (simp add: degree_of_sum_diff_degree) +qed + +lemma (in UP_ring) deg_diff_by_const: + assumes "g \ carrier (UP R)" + assumes "a \ carrier R" + assumes "h = g \\<^bsub>UP R\<^esub> up_ring.monom (UP R) a 0" + shows "deg R g = deg R h" + unfolding assms using assms + by (metis P_def UP_ring.bound_deg_sum UP_ring.deg_monom_le UP_ring.monom_closed UP_ring_axioms degree_of_sum_diff_degree gr_zeroI not_less) + +lemma (in UP_ring) deg_diff_by_const': + assumes "g \ carrier (UP R)" + assumes "a \ carrier R" + assumes "h = g \\<^bsub>UP R\<^esub> up_ring.monom (UP R) a 0" + shows "deg R g = deg R h" + apply(rule deg_diff_by_const[of _ "\ a"]) + using assms apply blast + using assms apply blast + by (metis P.minus_eq P_def assms(2) assms(3) monom_a_inv) + +lemma(in UP_ring) deg_gtE: + assumes "p \ carrier P" + assumes "i > deg R p" + shows "p i = \" + using assms P_def coeff_simp deg_aboveD by metis +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Leading Term Function\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition leading_term where +"leading_term R f = monom (UP R) (f (deg R f)) (deg R f)" + +context UP_ring +begin + +abbreviation(input) ltrm where +"ltrm f \ monom P (f (deg R f)) (deg R f)" + +text\leading term is a polynomial\ +lemma ltrm_closed: + assumes "f \ carrier P" + shows "ltrm f \ carrier P" + using assms + by (simp add: cfs_closed) + +text\Simplified coefficient function description for leading term\ +lemma ltrm_coeff: + assumes "f \ carrier P" + shows "coeff P (ltrm f) n = (if (n = degree f) then (f (degree f)) else \)" + using assms + by (simp add: cfs_closed) + +lemma ltrm_cfs: + assumes "f \ carrier P" + shows "(ltrm f) n = (if (n = degree f) then (f (degree f)) else \)" + using assms + by (simp add: cfs_closed cfs_monom) + +lemma ltrm_cfs_above_deg: + assumes "f \ carrier P" + assumes "n > degree f" + shows "ltrm f n = \" + using assms + by (simp add: ltrm_cfs) + +text\The leading term of f has the same degree as f\ + +lemma deg_ltrm: + assumes "f \ carrier P" + shows "degree (ltrm f) = degree f" + using assms + by (metis P_def UP_ring.lcoeff_nonzero_deg UP_ring_axioms cfs_closed coeff_simp deg_const deg_monom) + +text\Subtracting the leading term yields a drop in degree\ + +lemma minus_ltrm_degree_drop: + assumes "f \ carrier P" + assumes "degree f = Suc n" + shows "degree (f \\<^bsub>P\<^esub> (ltrm f)) \ n" +proof(rule UP_ring.deg_aboveI) + show C0: "UP_ring R" + by (simp add: UP_ring_axioms) + show C1: "f \\<^bsub>P\<^esub> ltrm f \ carrier (UP R)" + using assms ltrm_closed P.minus_closed P_def + by blast + show C2: "\m. n < m \ coeff (UP R) (f \\<^bsub>P\<^esub> ltrm f) m = \" + proof- + fix m + assume A: "n\<^bsub>P\<^esub> ltrm f) m = \" + proof(cases " m = Suc n") + case True + have B: "f m \ carrier R" + using UP.coeff_closed P_def assms(1) cfs_closed by blast + have "m = degree f" + using True by (simp add: assms(2)) + then have "f m = (ltrm f) m" + using ltrm_cfs assms(1) by auto + then have "(f m) \\<^bsub>R\<^esub>( ltrm f) m = \" + using B UP_ring_def P_is_UP_ring + B R.add.r_inv R.is_abelian_group abelian_group.minus_eq by fastforce + then have "(f \\<^bsub>UP R\<^esub> ltrm f) m = \" + by (metis C1 ltrm_closed P_def assms(1) coeff_minus coeff_simp) + then show ?thesis + using C1 P_def UP_ring.coeff_simp UP_ring_axioms by fastforce + next + case False + have D0: "m > degree f" using False + using A assms(2) by linarith + have B: "f m \ carrier R" + using UP.coeff_closed P_def assms(1) cfs_closed + by blast + have "f m = (ltrm f) m" + using D0 ltrm_cfs_above_deg P_def assms(1) coeff_simp deg_aboveD + by auto + then show ?thesis + by (metis B ltrm_closed P_def R.r_neg UP_ring.coeff_simp UP_ring_axioms a_minus_def assms(1) coeff_minus) + qed + qed +qed + +lemma ltrm_decomp: + assumes "f \ carrier P" + assumes "degree f >(0::nat)" + obtains g where "g \ carrier P \ f = g \\<^bsub>P\<^esub> (ltrm f) \ degree g < degree f" +proof- + have 0: "f \\<^bsub>P\<^esub> (ltrm f) \ carrier P" + using ltrm_closed assms(1) by blast + have 1: "f = (f \\<^bsub>P\<^esub> (ltrm f)) \\<^bsub>P\<^esub> (ltrm f)" + using assms + by (metis "0" ltrm_closed P.add.inv_solve_right P.minus_eq) + show ?thesis using assms 0 1 minus_ltrm_degree_drop[of f] + by (metis ltrm_closed Suc_diff_1 Suc_n_not_le_n deg_ltrm equal_deg_sum' linorder_neqE_nat that) +qed + +text\leading term of a sum\ +lemma coeff_of_sum_diff_degree0: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree q < n" + shows "(p \\<^bsub>P\<^esub> q) n = p n" + using assms P_def UP_ring.deg_aboveD UP_ring_axioms cfs_add coeff_simp cfs_closed deg_aboveD + by auto + +lemma coeff_of_sum_diff_degree1: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree q < degree p" + shows "(p \\<^bsub>P\<^esub> q) (degree p) = p (degree p)" + using assms(1) assms(2) assms(3) coeff_of_sum_diff_degree0 by blast + + + +lemma ltrm_of_sum_diff_degree: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree p > degree q" + shows "ltrm (p \\<^bsub>P\<^esub> q) = ltrm p" + unfolding leading_term_def + using assms(1) assms(2) assms(3) coeff_of_sum_diff_degree1 degree_of_sum_diff_degree + by presburger + +text\leading term of a monomial\ + +lemma ltrm_monom: + assumes "a \ carrier R" + assumes "f = monom P a n" + shows "ltrm f = f" + unfolding leading_term_def + by (metis P_def UP_ring.cfs_monom UP_ring.monom_zero UP_ring_axioms assms(1) assms(2) deg_monom) + +lemma ltrm_monom_simp: + assumes "a \ carrier R" + shows "ltrm (monom P a n) = monom P a n" + using assms ltrm_monom by auto + +lemma ltrm_inv_simp[simp]: + assumes "f \ carrier P" + shows "ltrm (ltrm f) = ltrm f" + by (metis assms deg_ltrm ltrm_cfs) + +lemma ltrm_deg_0: + assumes "p \ carrier P" + assumes "degree p = 0" + shows "ltrm p = p" + using ltrm_monom assms P_def UP_ring.deg_zero_impl_monom UP_ring_axioms coeff_simp + by fastforce + +lemma ltrm_prod_ltrm: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "ltrm ((ltrm p) \\<^bsub>P\<^esub> (ltrm q)) = (ltrm p) \\<^bsub>P\<^esub> (ltrm q)" + using ltrm_monom R.m_closed assms(1) assms(2) cfs_closed monom_mult + by metis + +text\lead coefficient function\ + +abbreviation(input) lcf where +"lcf p \ p (deg R p)" + +lemma(in UP_ring) lcf_ltrm: +"ltrm p = monom P (lcf p) (degree p)" + by auto + +lemma lcf_closed: + assumes "f \ carrier P" + shows "lcf f \ carrier R" + by (simp add: assms cfs_closed) + +lemma(in UP_cring) lcf_monom: + assumes "a \ carrier R" + shows "lcf (monom P a n) = a" "lcf (monom (UP R) a n) = a" + using assms deg_monom cfs_monom apply fastforce + by (metis UP_ring.cfs_monom UP_ring.deg_monom UP_ring_axioms assms) + + +end + +text\Function which truncates a polynomial by removing the leading term\ + +definition truncate where +"truncate R f = f \\<^bsub>(UP R)\<^esub> (leading_term R f)" + +context UP_ring +begin + +abbreviation(input) trunc where +"trunc \ truncate R" + +lemma trunc_closed: + assumes "f \ carrier P" + shows "trunc f \ carrier P" + using assms unfolding truncate_def + by (metis ltrm_closed P_def UP_ring.UP_ring UP_ring_axioms leading_term_def ring.ring_simprules(4)) + +lemma trunc_simps: + assumes "f \ carrier P" + shows "f = (trunc f) \\<^bsub>P\<^esub> (ltrm f)" + "f \\<^bsub>P\<^esub> (trunc f) = ltrm f" + apply (metis ltrm_closed P.add.inv_solve_right P.minus_closed P_def a_minus_def assms Cring_Poly.truncate_def leading_term_def) + using trunc_closed[of f] ltrm_closed[of f] P_def P.add.inv_solve_right[of "ltrm f" f "trunc f"] + assms unfolding UP_cring_def + by (metis P.add.inv_closed P.add.m_lcomm P.add.r_inv_ex P.minus_eq P.minus_minus + P.r_neg2 P.r_zero Cring_Poly.truncate_def leading_term_def) + +lemma trunc_zero: + assumes "f \ carrier P" + assumes "degree f = 0" + shows "trunc f = \\<^bsub>P\<^esub>" + unfolding truncate_def + using assms ltrm_deg_0[of f] + by (metis P.r_neg P_def a_minus_def leading_term_def) + +lemma trunc_degree: + assumes "f \ carrier P" + assumes "degree f > 0" + shows "degree (trunc f) < degree f" + unfolding truncate_def using assms + by (metis ltrm_closed ltrm_decomp P.add.right_cancel Cring_Poly.truncate_def trunc_closed trunc_simps(1)) + +text\The coefficients of trunc agree with f for small degree\ + +lemma trunc_cfs: + assumes "p \ carrier P" + assumes "n < degree p" + shows " (trunc p) n = p n" + using P_def assms(1) assms(2) unfolding truncate_def + by (smt ltrm_closed ltrm_cfs R.minus_zero R.ring_axioms UP_ring.cfs_minus + UP_ring_axioms a_minus_def cfs_closed leading_term_def nat_neq_iff ring.ring_simprules(15)) + +text\monomial predicate\ + +definition is_UP_monom where +"is_UP_monom = (\f. f \ carrier (UP R) \ f = ltrm f)" + +lemma is_UP_monomI: + assumes "a \ carrier R" + assumes "p = monom P a n" + shows "is_UP_monom p" + using assms(1) assms(2) is_UP_monom_def ltrm_monom P_def monom_closed + by auto + +lemma is_UP_monomI': + assumes "f \ carrier (UP R)" + assumes "f = ltrm f" + shows "is_UP_monom f" + using assms P_def unfolding is_UP_monom_def by blast + +lemma monom_is_UP_monom: + assumes "a \ carrier R" + shows "is_UP_monom (monom P a n)" "is_UP_monom (monom (UP R) a n)" + using assms P_def ltrm_monom_simp monom_closed + unfolding is_UP_monom_def + by auto + +lemma is_UP_monomE: + assumes "is_UP_monom f" + shows "f \ carrier P" "f = monom P (lcf f) (degree f)" "f = monom (UP R) (lcf f) (degree f)" + using assms unfolding is_UP_monom_def + by(auto simp: P_def ) + +lemma ltrm_is_UP_monom: + assumes "p \ carrier P" + shows "is_UP_monom (ltrm p)" + using assms + by (simp add: cfs_closed monom_is_UP_monom(1)) + +lemma is_UP_monom_mult: + assumes "is_UP_monom p" + assumes "is_UP_monom q" + shows "is_UP_monom (p \\<^bsub>P\<^esub> q)" + apply(rule is_UP_monomI') + using assms is_UP_monomE P_def UP_mult_closed + apply simp + using assms is_UP_monomE[of p] is_UP_monomE[of q] + P_def monom_mult + by (metis lcf_closed ltrm_monom R.m_closed) +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Properties of Leading Terms and Leading Coefficients in Commutative Rings and Domains\ + (**************************************************************************************************) + (**************************************************************************************************) + +context UP_cring +begin + +lemma cring_deg_mult: + assumes "q \ carrier P" + assumes "p \ carrier P" + assumes "lcf q \ lcf p \\" + shows "degree (q \\<^bsub>P\<^esub> p) = degree p + degree q" +proof- + have "q \\<^bsub>P\<^esub> p = (trunc q \\<^bsub>P\<^esub> ltrm q) \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)" + using assms(1) assms(2) trunc_simps(1) by auto + then have "q \\<^bsub>P\<^esub> p = (trunc q \\<^bsub>P\<^esub> ltrm q) \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)" + by linarith + then have 0: "q \\<^bsub>P\<^esub> p = (trunc q \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)) \\<^bsub>P\<^esub> ( ltrm q \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p))" + by (simp add: P.l_distr assms(1) assms(2) ltrm_closed trunc_closed) + have 1: "(trunc q \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)) (degree p + degree q) = \" + proof(cases "degree q = 0") + case True + then show ?thesis + using assms(1) assms(2) trunc_simps(1) trunc_zero by auto + next + case False + have "degree ((trunc q) \\<^bsub>P\<^esub> p) \ degree (trunc q) + degree p" + using assms trunc_simps[of q] deg_mult_ring[of "trunc q" p] trunc_closed + by blast + then have "degree (trunc q \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)) < degree q + degree p" + using False assms(1) assms(2) trunc_degree trunc_simps(1) by fastforce + then show ?thesis + by (metis P_def UP_mult_closed UP_ring.coeff_simp UP_ring_axioms + add.commute assms(1) assms(2) deg_belowI not_less trunc_closed trunc_simps(1)) + qed + have 2: "(q \\<^bsub>P\<^esub> p) (degree p + degree q) = + ( ltrm q \\<^bsub>P\<^esub> (trunc p \\<^bsub>P\<^esub> ltrm p)) (degree p + degree q)" + using 0 1 assms cfs_closed trunc_closed by auto + have 3: "(q \\<^bsub>P\<^esub> p) (degree p + degree q) = + ( ltrm q \\<^bsub>P\<^esub> trunc p) (degree p + degree q) \ ( ltrm q \\<^bsub>P\<^esub> ltrm p) (degree p + degree q)" + by (simp add: "2" ltrm_closed UP_r_distr assms(1) assms(2) trunc_closed) + have 4: "( ltrm q \\<^bsub>P\<^esub> trunc p) (degree p + degree q) = \" + proof(cases "degree p = 0") + case True + then show ?thesis + using "2" "3" assms(1) assms(2) cfs_closed ltrm_closed trunc_zero by auto + next + case False + have "degree ( ltrm q \\<^bsub>P\<^esub> trunc p) \ degree (ltrm q) + degree (trunc p)" + using assms trunc_simps deg_mult_ring ltrm_closed trunc_closed by presburger + then have "degree ( ltrm q \\<^bsub>P\<^esub> trunc p) < degree q + degree p" + using False assms(1) assms(2) trunc_degree trunc_simps(1) deg_ltrm by fastforce + then show ?thesis + by (metis ltrm_closed P_def UP_mult_closed UP_ring.coeff_simp UP_ring_axioms add.commute assms(1) assms(2) deg_belowI not_less trunc_closed) + qed + have 5: "(q \\<^bsub>P\<^esub> p) (degree p + degree q) = ( ltrm q \\<^bsub>P\<^esub> ltrm p) (degree p + degree q)" + by (simp add: "3" "4" assms(1) assms(2) cfs_closed) + have 6: "ltrm q \\<^bsub>P\<^esub> ltrm p = monom P (lcf q \ lcf p) (degree p + degree q)" + unfolding leading_term_def + by (metis P_def UP_ring.monom_mult UP_ring_axioms add.commute assms(1) assms(2) cfs_closed) + have 7: "( ltrm q \\<^bsub>P\<^esub> ltrm p) (degree p + degree q) \\" + using 5 6 assms + by (metis R.m_closed cfs_closed cfs_monom) + have 8: "degree (q \\<^bsub>P\<^esub> p) \degree p + degree q" + using 5 6 7 P_def UP_mult_closed assms(1) assms(2) + by (simp add: UP_ring.coeff_simp UP_ring_axioms deg_belowI) + then show ?thesis + using assms(1) assms(2) deg_mult_ring by fastforce +qed + +text\leading term is multiplicative\ + +lemma ltrm_of_sum_diff_deg: + assumes "q \ carrier P" + assumes "a \ carrier R" + assumes "a \\" + assumes "degree q < n" + assumes "p = q \\<^bsub>P\<^esub> (monom P a n)" + shows "ltrm p = (monom P a n)" +proof- + have 0: "degree (monom P a n) = n" + by (simp add: assms(2) assms(3)) + have 1: "(monom P a n) \ carrier P" + using assms(2) by auto + have 2: "ltrm ((monom P a n) \\<^bsub>P\<^esub> q) = ltrm (monom P a n)" + using assms ltrm_of_sum_diff_degree[of "(monom P a n)" q] 1 "0" by linarith + then show ?thesis + using UP_a_comm assms(1) assms(2) assms(5) ltrm_monom by auto +qed + +lemma(in UP_cring) ltrm_smult_cring: + assumes "p \ carrier P" + assumes "a \ carrier R" + assumes "lcf p \ a \ \" + shows "ltrm (a \\<^bsub>P\<^esub>p) = a\\<^bsub>P\<^esub>(ltrm p)" + using assms + by (smt lcf_monom(1) P_def R.m_closed R.m_comm cfs_closed cfs_smult coeff_simp + cring_deg_mult deg_monom deg_ltrm monom_closed monom_mult_is_smult monom_mult_smult) + +lemma(in UP_cring) deg_zero_ltrm_smult_cring: + assumes "p \ carrier P" + assumes "a \ carrier R" + assumes "degree p = 0" + shows "ltrm (a \\<^bsub>P\<^esub>p) = a\\<^bsub>P\<^esub>(ltrm p)" + by (metis ltrm_deg_0 assms(1) assms(2) assms(3) deg_smult_decr le_0_eq module.smult_closed module_axioms) + +lemma(in UP_domain) ltrm_smult: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "ltrm (a \\<^bsub>P\<^esub>p) = a\\<^bsub>P\<^esub>(ltrm p)" + by (metis lcf_closed ltrm_closed ltrm_smult_cring P_def R.integral_iff UP_ring.deg_ltrm + UP_ring_axioms UP_smult_zero assms(1) assms(2) cfs_zero deg_nzero_nzero deg_zero_ltrm_smult_cring monom_zero) + +lemma(in UP_cring) cring_ltrm_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "lcf p \ lcf q \ \" + shows "ltrm (p \\<^bsub>P\<^esub> q) = (ltrm p) \\<^bsub>P\<^esub> (ltrm q)" +proof(cases "degree p = 0 \ degree q = 0") + case True + then show ?thesis + by (smt ltrm_closed ltrm_deg_0 ltrm_smult_cring R.m_comm UP_m_comm assms(1) assms(2) assms(3) cfs_closed monom_mult_is_smult) +next + case False + obtain q0 where q0_def: "q0 = trunc q" + by simp + obtain p0 where p0_def: "p0 = trunc p" + by simp + have Pq: "degree q0 < degree q" + using False P_def assms(2) q0_def trunc_degree by blast + have Pp: "degree p0 < degree p" + using False P_def assms(1) p0_def trunc_degree by blast + have "p \\<^bsub>P\<^esub> q = (p0 \\<^bsub>P\<^esub> ltrm(p)) \\<^bsub>P \<^esub>(q0 \\<^bsub>P\<^esub> ltrm(q))" + using assms(1) assms(2) p0_def q0_def trunc_simps(1) by auto + then have P0: "p \\<^bsub>P\<^esub> q = ((p0 \\<^bsub>P\<^esub> ltrm(p)) \\<^bsub>P \<^esub>q0) \\<^bsub>P\<^esub> ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q))" + by (simp add: P.r_distr assms(1) assms(2) ltrm_closed p0_def q0_def trunc_closed) + have P1: "degree ((p0 \\<^bsub>P\<^esub> ltrm(p)) \\<^bsub>P \<^esub>q0) < degree ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q))" + proof- + have LHS: "degree ((p0 \\<^bsub>P\<^esub> ltrm(p)) \\<^bsub>P \<^esub>q0) \ degree p + degree q0 " + proof(cases "q0 = \\<^bsub>P\<^esub>") + case True + then show ?thesis + using assms(1) p0_def trunc_simps(1) by auto + next + case False + then show ?thesis + using assms(1) assms(2) deg_mult_ring p0_def + q0_def trunc_simps(1) trunc_closed by auto + qed + have RHS: "degree ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q)) = degree p + degree q" + using assms(1) assms(2) deg_mult_ring ltrm_closed p0_def trunc_simps(1) + by (smt P_def UP_cring.lcf_monom(1) UP_cring.cring_deg_mult UP_cring_axioms add.commute assms(3) cfs_closed deg_ltrm) + then show ?thesis + using RHS LHS Pq + by linarith + qed + then have P2: "ltrm (p \\<^bsub>P\<^esub> q) = ltrm ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q))" + using P0 P1 + by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm + UP_mult_closed assms(1) assms(2) p0_def q0_def trunc_closed trunc_simps(1)) + have P3: " ltrm ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q)) = ltrm p \\<^bsub>P\<^esub> ltrm q" + proof- + have Q0: "((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q)) = (p0 \\<^bsub>P \<^esub>ltrm(q)) \\<^bsub>P\<^esub> (ltrm(p))\\<^bsub>P \<^esub>ltrm(q)" + by (simp add: P.l_distr assms(1) assms(2) ltrm_closed p0_def trunc_closed) + have Q1: "degree ((p0 \\<^bsub>P \<^esub>ltrm(q)) ) < degree ((ltrm(p))\\<^bsub>P \<^esub>ltrm(q))" + proof(cases "p0 = \\<^bsub>P\<^esub>") + case True + then show ?thesis + using P1 assms(1) assms(2) ltrm_closed by auto + next + case F: False + then show ?thesis + proof- + have LHS: "degree ((p0 \\<^bsub>P \<^esub>ltrm(q))) < degree p + degree q" + using False F Pp assms(1) assms(2) deg_nzero_nzero + deg_ltrm ltrm_closed p0_def trunc_closed + by (smt add_le_cancel_right deg_mult_ring le_trans not_less) + have RHS: "degree ((ltrm(p))\\<^bsub>P \<^esub>ltrm(q)) = degree p + degree q" + using cring_deg_mult[of "ltrm p" "ltrm q"] assms + by (simp add: ltrm_closed ltrm_cfs deg_ltrm) + then show ?thesis using LHS RHS by auto + qed + qed + have Q2: "ltrm ((p0 \\<^bsub>P\<^esub> ltrm(p))\\<^bsub>P \<^esub>ltrm(q)) = ltrm ((ltrm(p))\\<^bsub>P \<^esub>ltrm(q))" + using Q0 Q1 + by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm + UP_mult_closed assms(1) assms(2) p0_def trunc_closed) + show ?thesis using ltrm_prod_ltrm Q0 Q1 Q2 + by (simp add: assms(1) assms(2)) + qed + then show ?thesis + by (simp add: P2) +qed + +lemma(in UP_domain) ltrm_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "ltrm (p \\<^bsub>P\<^esub> q) = (ltrm p) \\<^bsub>P\<^esub> (ltrm q)" + using cring_ltrm_mult assms + by (smt ltrm_closed ltrm_deg_0 cfs_closed deg_nzero_nzero deg_ltrm local.integral_iff monom_mult monom_zero) + +lemma lcf_deg_0: + assumes "degree p = 0" + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "(p \\<^bsub>P\<^esub> q) = (lcf p)\\<^bsub>P\<^esub>q" + using P_def assms(1) assms(2) assms(3) + by (metis ltrm_deg_0 cfs_closed monom_mult_is_smult) + +text\leading term powers\ + +lemma (in domain) nonzero_pow_nonzero: + assumes "a \ carrier R" + assumes "a \\" + shows "a[^](n::nat) \ \" +proof(induction n) + case 0 + then show ?case + by auto +next + case (Suc n) + fix n::nat + assume IH: "a[^] n \ \" + show "a[^] (Suc n) \ \" + proof- + have "a[^] (Suc n) = a[^] n \ a" + by simp + then show ?thesis using assms IH + using IH assms(1) assms(2) local.integral by auto + qed +qed + +lemma (in UP_cring) cring_monom_degree: + assumes "a \ (carrier R)" + assumes "p = monom P a m" + assumes "a[^]n \ \" + shows "degree (p[^]\<^bsub>P\<^esub> n) = n*m" + by (simp add: assms(1) assms(2) assms(3) monom_pow) + +lemma (in UP_domain) monom_degree: + assumes "a \\" + assumes "a \ (carrier R)" + assumes "p = monom P a m" + shows "degree (p[^]\<^bsub>P\<^esub> n) = n*m" + by (simp add: R.domain_axioms assms(1) assms(2) assms(3) domain.nonzero_pow_nonzero monom_pow) + +lemma(in UP_cring) cring_pow_ltrm: + assumes "p \ carrier P" + assumes "lcf p [^]n \ \" + shows "ltrm (p[^]\<^bsub>P\<^esub>(n::nat)) = (ltrm p)[^]\<^bsub>P\<^esub>n" +proof- + have "lcf p [^]n \ \ \ ltrm (p[^]\<^bsub>P\<^esub>(n::nat)) = (ltrm p)[^]\<^bsub>P\<^esub>n" + proof(induction n) + case 0 + then show ?case + using P.ring_simprules(6) P.nat_pow_0 cfs_one deg_one monom_one by presburger + next + case (Suc n) fix n::nat + assume IH : "(lcf p [^] n \ \ \ ltrm (p [^]\<^bsub>P\<^esub> n) = ltrm p [^]\<^bsub>P\<^esub> n)" + assume A: "lcf p [^] Suc n \ \" + have a: "ltrm (p [^]\<^bsub>P\<^esub> n) = ltrm p [^]\<^bsub>P\<^esub> n" + apply(cases "lcf p [^] n = \") + using A lcf_closed assms(1) apply auto[1] + by(rule IH) + have 0: "lcf (ltrm (p [^]\<^bsub>P\<^esub> n)) = lcf p [^] n" + unfolding a + by (simp add: lcf_monom(1) assms(1) cfs_closed monom_pow) + then have 1: "lcf (ltrm (p [^]\<^bsub>P\<^esub> n)) \ lcf p \ \" + using assms A R.nat_pow_Suc IH by metis + then show "ltrm (p [^]\<^bsub>P\<^esub> Suc n) = ltrm p [^]\<^bsub>P\<^esub> Suc n" + using IH 0 assms(1) cring_ltrm_mult cfs_closed + by (smt A lcf_monom(1) ltrm_closed P.nat_pow_Suc2 P.nat_pow_closed R.nat_pow_Suc2 a) + qed + then show ?thesis + using assms(2) by blast +qed + +lemma(in UP_cring) cring_pow_deg: + assumes "p \ carrier P" + assumes "lcf p [^]n \ \" + shows "degree (p[^]\<^bsub>P\<^esub>(n::nat)) = n*degree p" +proof- + have "degree ( (ltrm p)[^]\<^bsub>P\<^esub>n) = n*degree p" + using assms(1) assms(2) cring_monom_degree lcf_closed lcf_ltrm by auto + then show ?thesis + using assms cring_pow_ltrm + by (metis P.nat_pow_closed P_def UP_ring.deg_ltrm UP_ring_axioms) +qed + +lemma(in UP_cring) cring_pow_deg_bound: + assumes "p \ carrier P" + shows "degree (p[^]\<^bsub>P\<^esub>(n::nat)) \ n*degree p" + apply(induction n) + apply (metis Group.nat_pow_0 deg_one le_zero_eq mult_is_0) + using deg_mult_ring[of _ p] + by (smt P.nat_pow_Suc2 P.nat_pow_closed ab_semigroup_add_class.add_ac(1) assms deg_mult_ring le_iff_add mult_Suc) + +lemma(in UP_cring) deg_smult: + assumes "a \ carrier R" + assumes "f \ carrier (UP R)" + assumes "a \ lcf f \ \" + shows "deg R (a \\<^bsub>UP R\<^esub> f) = deg R f" + using assms P_def cfs_smult deg_eqI deg_smult_decr smult_closed + by (metis deg_gtE le_neq_implies_less) + +lemma(in UP_cring) deg_smult': + assumes "a \ Units R" + assumes "f \ carrier (UP R)" + shows "deg R (a \\<^bsub>UP R\<^esub> f) = deg R f" + apply(cases "deg R f = 0") + apply (metis P_def R.Units_closed assms(1) assms(2) deg_smult_decr le_zero_eq) + apply(rule deg_smult) + using assms apply blast + using assms apply blast +proof + assume A: "deg R f \ 0" "a \ f (deg R f) = \" + have 0: "f (deg R f) = \" + using A assms R.Units_not_right_zero_divisor[of a "f (deg R f)"] UP_car_memE(1) by blast + then show False using assms A + by (metis P_def deg_zero deg_ltrm monom_zero) +qed + +lemma(in UP_domain) pow_sum0: +"\ p q. p \ carrier P \ q \ carrier P \ degree q < degree p \ degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) = (degree p)*n" +proof(induction n) + case 0 + then show ?case + by (metis Group.nat_pow_0 deg_one mult_is_0) +next + case (Suc n) + fix n + assume IH: "\ p q. p \ carrier P \ q \ carrier P \ + degree q < degree p \ degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) = (degree p)*n" + then show "\ p q. p \ carrier P \ q \ carrier P \ + degree q < degree p \ degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>(Suc n)) = (degree p)*(Suc n)" + proof- + fix p q + assume A0: "p \ carrier P" and + A1: "q \ carrier P" and + A2: "degree q < degree p" + show "degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>(Suc n)) = (degree p)*(Suc n)" + proof(cases "q = \\<^bsub>P\<^esub>") + case True + then show ?thesis + by (metis A0 A1 A2 IH P.nat_pow_Suc2 P.nat_pow_closed P.r_zero deg_mult + domain.nonzero_pow_nonzero local.domain_axioms mult_Suc_right nat_neq_iff) + next + case False + then show ?thesis + proof- + have P0: "degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) = (degree p)*n" + using A0 A1 A2 IH by auto + have P1: "(p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>(Suc n) = ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q )" + by simp + then have P2: "(p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>(Suc n) = (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> p) \\<^bsub>P\<^esub> (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> q)" + by (simp add: A0 A1 UP_r_distr) + have P3: "degree (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> p) = (degree p)*n + (degree p)" + using P0 A0 A1 A2 deg_nzero_nzero degree_of_sum_diff_degree local.nonzero_pow_nonzero by auto + have P4: "degree (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> q) = (degree p)*n + (degree q)" + using P0 A0 A1 A2 deg_nzero_nzero degree_of_sum_diff_degree local.nonzero_pow_nonzero False deg_mult + by simp + have P5: "degree (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> p) > degree (((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub> q)" + using P3 P4 A2 by auto + then show ?thesis using P5 P3 P2 + by (simp add: A0 A1 degree_of_sum_diff_degree) + qed + qed + qed +qed + +lemma(in UP_domain) pow_sum: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree q < degree p" + shows "degree ((p \\<^bsub>P\<^esub> q )[^]\<^bsub>P\<^esub>n) = (degree p)*n" + using assms(1) assms(2) assms(3) pow_sum0 by blast + +lemma(in UP_domain) deg_pow0: + "\ p. p \ carrier P \ n \ degree p \ degree (p [^]\<^bsub>P\<^esub> m) = m*(degree p)" +proof(induction n) + case 0 + show "p \ carrier P \ 0 \ degree p \ degree (p [^]\<^bsub>P\<^esub> m) = m*(degree p)" + proof- + assume B0:"p \ carrier P" + assume B1: "0 \ degree p" + then obtain a where a_def: "a \ carrier R \ p = monom P a 0" + using B0 deg_zero_impl_monom by fastforce + show "degree (p [^]\<^bsub>P\<^esub> m) = m*(degree p)" using UP_cring.monom_pow + by (metis P_def R.nat_pow_closed UP_cring_axioms a_def deg_const + mult_0_right mult_zero_left) + qed +next + case (Suc n) + fix n + assume IH: "\p. (p \ carrier P \ n \degree p \ degree (p [^]\<^bsub>P\<^esub> m) = m * (degree p))" + show "p \ carrier P \ Suc n \ degree p \ degree (p [^]\<^bsub>P\<^esub> m) = m * (degree p)" + proof- + assume A0: "p \ carrier P" + assume A1: "Suc n \ degree p" + show "degree (p [^]\<^bsub>P\<^esub> m) = m * (degree p)" + proof(cases "Suc n > degree p") + case True + then show ?thesis using IH A0 by simp + next + case False + then show ?thesis + proof- + obtain q where q_def: "q = trunc p" + by simp + obtain k where k_def: "k = degree q" + by simp + have q_is_poly: "q \ carrier P" + by (simp add: A0 q_def trunc_closed) + have k_bound0: "k n" + using k_bound0 A0 A1 by auto + have P_q:"degree (q [^]\<^bsub>P\<^esub> m) = m * k" + using IH[of "q"] k_bound1 k_def q_is_poly by auto + have P_ltrm: "degree ((ltrm p) [^]\<^bsub>P\<^esub> m) = m*(degree p)" + proof- + have "degree p = degree (ltrm p)" + by (simp add: A0 deg_ltrm) + then show ?thesis using monom_degree + by (metis A0 P.r_zero P_def cfs_closed coeff_simp equal_deg_sum k_bound0 k_def lcoeff_nonzero2 nat_neq_iff q_is_poly) + qed + have "p = q \\<^bsub>P\<^esub> (ltrm p)" + by (simp add: A0 q_def trunc_simps(1)) + then show ?thesis + using P_q pow_sum[of "(ltrm p)" q m] A0 UP_a_comm + deg_ltrm k_bound0 k_def ltrm_closed q_is_poly by auto + qed + qed + qed +qed + +lemma(in UP_domain) deg_pow: + assumes "p \ carrier P" + shows "degree (p [^]\<^bsub>P\<^esub> m) = m*(degree p)" + using deg_pow0 assms by blast + +lemma(in UP_domain) ltrm_pow0: +"\f. f \ carrier P \ ltrm (f [^]\<^bsub>P\<^esub> (n::nat)) = (ltrm f) [^]\<^bsub>P\<^esub> n" +proof(induction n) + case 0 + then show ?case + using ltrm_deg_0 P.nat_pow_0 P.ring_simprules(6) deg_one by presburger +next + case (Suc n) + fix n::nat + assume IH: "\f. f \ carrier P \ ltrm (f [^]\<^bsub>P\<^esub> n) = (ltrm f) [^]\<^bsub>P\<^esub> n" + then show "\f. f \ carrier P \ ltrm (f [^]\<^bsub>P\<^esub> (Suc n)) = (ltrm f) [^]\<^bsub>P\<^esub> (Suc n)" + proof- + fix f + assume A: "f \ carrier P" + show " ltrm (f [^]\<^bsub>P\<^esub> (Suc n)) = (ltrm f) [^]\<^bsub>P\<^esub> (Suc n)" + proof- + have 0: "ltrm (f [^]\<^bsub>P\<^esub> n) = (ltrm f) [^]\<^bsub>P\<^esub> n" + using A IH by blast + have 1: "ltrm (f [^]\<^bsub>P\<^esub> (Suc n)) = ltrm ((f [^]\<^bsub>P\<^esub> n)\\<^bsub>P\<^esub> f)" + by auto then + show ?thesis using ltrm_mult 0 1 + by (simp add: A) + qed + qed +qed + +lemma(in UP_domain) ltrm_pow: + assumes "f \ carrier P" + shows " ltrm (f [^]\<^bsub>P\<^esub> (n::nat)) = (ltrm f) [^]\<^bsub>P\<^esub> n" + using assms ltrm_pow0 by blast + +text\lemma on the leading coefficient\ + +lemma lcf_eq: + assumes "f \ carrier P" + shows "lcf f = lcf (ltrm f)" + using ltrm_deg_0 + by (simp add: ltrm_cfs assms deg_ltrm) + +lemma lcf_eq_deg_eq_imp_ltrm_eq: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree p > 0" + assumes "degree p = degree q" + assumes "lcf p = lcf q" + shows "ltrm p = ltrm q" + using assms(4) assms(5) + by (simp add: leading_term_def) + +lemma ltrm_eq_imp_lcf_eq: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "ltrm p = ltrm q" + shows "lcf p = lcf q" + using assms + by (metis lcf_eq) + +lemma ltrm_eq_imp_deg_drop: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "ltrm p = ltrm q" + assumes "degree p >0" + shows "degree (p \\<^bsub>P\<^esub> q) < degree p" +proof- + have P0: "degree p = degree q" + by (metis assms(1) assms(2) assms(3) deg_ltrm) + then have P1: "degree (p \\<^bsub>P\<^esub> q) \ degree p" + by (metis P.add.inv_solve_right P.minus_closed P.minus_eq assms(1) + assms(2) degree_of_sum_diff_degree neqE order.strict_implies_order order_refl) + have "degree (p \\<^bsub>P\<^esub> q) \ degree p" + proof + assume A: "degree (p \\<^bsub>P\<^esub> q) = degree p" + have Q0: "p \\<^bsub>P\<^esub> q = ((trunc p) \\<^bsub>P\<^esub> (ltrm p)) \\<^bsub>P\<^esub> ((trunc q) \\<^bsub>P\<^esub> (ltrm p))" + using assms(1) assms(2) assms(3) trunc_simps(1) by force + have Q1: "p \\<^bsub>P\<^esub> q = (trunc p) \\<^bsub>P\<^esub> (trunc q)" + proof- + have "p \\<^bsub>P\<^esub> q = ((trunc p) \\<^bsub>P\<^esub> (ltrm p)) \\<^bsub>P\<^esub> (trunc q) \ \<^bsub>P\<^esub> (ltrm p)" + using Q0 + by (simp add: P.minus_add P.minus_eq UP_a_assoc assms(1) assms(2) ltrm_closed trunc_closed) + then show ?thesis + by (metis (no_types, lifting) ltrm_closed P.add.inv_mult_group P.minus_eq + P.r_neg2 UP_a_assoc assms(1) assms(2) assms(3) carrier_is_submodule submoduleE(6) trunc_closed trunc_simps(1)) + qed + have Q2: "degree (trunc p) < degree p" + by (simp add: assms(1) assms(4) trunc_degree) + have Q3: "degree (trunc q) < degree q" + using P0 assms(2) assms(4) trunc_degree by auto + then show False using A Q1 Q2 Q3 by (simp add: P.add.inv_solve_right + P.minus_eq P0 assms(1) assms(2) degree_of_sum_diff_degree trunc_closed) + qed + then show ?thesis + using P1 by auto +qed + +lemma(in UP_cring) cring_lcf_scalar_mult: + assumes "p \ carrier P" + assumes "a \ carrier R" + assumes "a \ (lcf p) \\" + shows "lcf (a \\<^bsub>P\<^esub> p) = a \ (lcf p)" +proof- + have 0: "lcf (a \\<^bsub>P\<^esub> p) = lcf (ltrm (a \\<^bsub>P\<^esub> p))" + using assms lcf_eq smult_closed by blast + have 1: "degree (a \\<^bsub>P\<^esub> p) = degree p" + by (smt lcf_monom(1) P_def R.one_closed R.r_null UP_ring.coeff_smult UP_ring_axioms + assms(1) assms(2) assms(3) coeff_simp cring_deg_mult deg_const monom_closed monom_mult_is_smult smult_one) + then have "lcf (a \\<^bsub>P\<^esub> p) = lcf (a \\<^bsub>P\<^esub> (ltrm p))" + using lcf_eq[of "a \\<^bsub>P\<^esub> p"] smult_closed assms 0 + by (metis cfs_closed cfs_smult monom_mult_smult) + then show ?thesis + unfolding leading_term_def + by (metis P_def R.m_closed UP_cring.lcf_monom UP_cring_axioms assms(1) assms(2) cfs_closed monom_mult_smult) +qed + +lemma(in UP_domain) lcf_scalar_mult: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "lcf (a \\<^bsub>P\<^esub> p) = a \ (lcf p)" +proof- + have "lcf (a \\<^bsub>P\<^esub> p) = lcf (ltrm (a \\<^bsub>P\<^esub> p))" + using lcf_eq UP_smult_closed assms(1) assms(2) by blast + then have "lcf (a \\<^bsub>P\<^esub> p) = lcf (a \\<^bsub>P\<^esub> (ltrm p))" + using ltrm_smult assms(1) assms(2) by metis + then show ?thesis + by (metis (full_types) UP_smult_zero assms(1) assms(2) cfs_smult cfs_zero deg_smult) +qed + +lemma(in UP_cring) cring_lcf_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "(lcf p) \ (lcf q) \\" + shows "lcf (p \\<^bsub>P\<^esub> q) = (lcf p) \ (lcf q)" + using assms cring_ltrm_mult + by (smt lcf_monom(1) P.m_closed R.m_closed cfs_closed monom_mult) + +lemma(in UP_domain) lcf_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "lcf (p \\<^bsub>P\<^esub> q) = (lcf p) \ (lcf q)" + by (smt ltrm_deg_0 R.integral_iff assms(1) assms(2) cfs_closed cring_lcf_mult deg_zero deg_ltrm local.integral_iff monom_zero) + +lemma(in UP_cring) cring_lcf_pow: + assumes "p \ carrier P" + assumes "(lcf p)[^]n \\" + shows "lcf (p[^]\<^bsub>P\<^esub>(n::nat)) = (lcf p)[^]n" + by (smt P.nat_pow_closed R.nat_pow_closed assms(1) assms(2) cring_pow_ltrm lcf_closed lcf_ltrm lcf_monom monom_pow) + +lemma(in UP_domain) lcf_pow: + assumes "p \ carrier P" + shows "lcf (p[^]\<^bsub>P\<^esub>(n::nat)) = (lcf p)[^]n" +proof- + show ?thesis + proof(induction n) + case 0 + then show ?case + by (metis Group.nat_pow_0 P_def R.one_closed UP_cring.lcf_monom UP_cring_axioms monom_one) + next + case (Suc n) + fix n + assume IH: "lcf (p[^]\<^bsub>P\<^esub>(n::nat)) = (lcf p)[^]n" + show "lcf (p[^]\<^bsub>P\<^esub>(Suc n)) = (lcf p)[^](Suc n)" + proof- + have "lcf (p[^]\<^bsub>P\<^esub>(Suc n)) = lcf ((p[^]\<^bsub>P\<^esub>n) \\<^bsub>P\<^esub>p)" + by simp + then have "lcf (p[^]\<^bsub>P\<^esub>(Suc n)) = (lcf p)[^]n \ (lcf p)" + by (simp add: IH assms lcf_mult) + then show ?thesis by auto + qed + qed +qed +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Constant Terms and Constant Coefficients\ + (**************************************************************************************************) + (**************************************************************************************************) + +text\Constant term and coefficient function\ + +definition zcf where +"zcf f = (f 0)" + +abbreviation(in UP_cring)(input) ctrm where +"ctrm f \ monom P (f 0) 0" + +context UP_cring +begin + +lemma ctrm_is_poly: + assumes "p \ carrier P" + shows "ctrm p \ carrier P" + by (simp add: assms cfs_closed) + +lemma ctrm_degree: + assumes "p \ carrier P" + shows "degree (ctrm p) = 0" + by (simp add: assms cfs_closed) + +lemma ctrm_zcf: +assumes "f \ carrier P" +assumes "zcf f = \" +shows "ctrm f = \\<^bsub>P\<^esub>" + by (metis P_def UP_ring.monom_zero UP_ring_axioms zcf_def assms(2)) + +lemma zcf_degree_zero: + assumes "f \ carrier P" + assumes "degree f = 0" + shows "lcf f = zcf f" + by (simp add: zcf_def assms(2)) + +lemma zcf_zero_degree_zero: + assumes "f \ carrier P" + assumes "degree f = 0" + assumes "zcf f = \" + shows "f = \\<^bsub>P\<^esub>" + using zcf_degree_zero[of f] assms ltrm_deg_0[of f] + by simp + +lemma zcf_ctrm: + assumes "p \ carrier P" + shows "zcf (ctrm p) = zcf p" + unfolding zcf_def + using P_def UP_ring.cfs_monom UP_ring_axioms assms cfs_closed by fastforce + +lemma ctrm_trunc: + assumes "p \ carrier P" + assumes "degree p >0" + shows "zcf(trunc p) = zcf p" + by (simp add: zcf_def assms(1) assms(2) trunc_cfs) + +text\Constant coefficient function is a ring homomorphism\ + + +lemma zcf_add: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "zcf(p \\<^bsub>P\<^esub> q) = (zcf p) \ (zcf q)" + by (simp add: zcf_def assms(1) assms(2)) + +lemma coeff_ltrm[simp]: + assumes "p \ carrier P" + assumes "degree p > 0" + shows "zcf(ltrm p) = \" + by (metis ltrm_cfs_above_deg ltrm_cfs zcf_def assms(1) assms(2)) + +lemma zcf_zero[simp]: +"zcf \\<^bsub>P\<^esub> = \" + using zcf_degree_zero by auto + +lemma zcf_one[simp]: +"zcf \\<^bsub>P\<^esub> = \" + by (simp add: zcf_def) + +lemma ctrm_smult: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "ctrm (a \\<^bsub>P\<^esub> f) = a \\<^bsub>P\<^esub>(ctrm f)" + using P_def UP_ring.monom_mult_smult UP_ring_axioms assms(1) assms(2) cfs_smult coeff_simp + by (simp add: UP_ring.monom_mult_smult cfs_closed) + +lemma ctrm_monom[simp]: + assumes "a \ carrier R" + shows "ctrm (monom P a (Suc k)) = \\<^bsub>P\<^esub>" + by (simp add: assms cfs_monom) +end + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Polynomial Induction Rules\ + (**************************************************************************************************) + (**************************************************************************************************) + +context UP_ring +begin + +text\Rule for strong induction on polynomial degree\ + +lemma poly_induct: + assumes "p \ carrier P" + assumes Deg_0: "\p. p \ carrier P \ degree p = 0 \ Q p" + assumes IH: "\p. (\q. q \ carrier P \ degree q < degree p \ Q q) \ p \ carrier P \ degree p > 0 \ Q p" + shows "Q p" +proof- + have "\n. \p. p \ carrier P \ degree p \ n \ Q p" + proof- + fix n + show "\p. p \ carrier P \ degree p \ n \ Q p" + proof(induction n) + case 0 + then show ?case + using Deg_0 by simp + next + case (Suc n) + fix n + assume I: "\p. p \ carrier P \ degree p \ n \ Q p" + show "\p. p \ carrier P \ degree p \ (Suc n) \ Q p" + proof- + fix p + assume A0: " p \ carrier P " + assume A1: "degree p \Suc n" + show "Q p" + proof(cases "degree p < Suc n") + case True + then show ?thesis + using I A0 by auto + next + case False + then have D: "degree p = Suc n" + by (simp add: A1 nat_less_le) + then have "(\q. q \ carrier P \ degree q < degree p \ Q q)" + using I by simp + then show "Q p" + using IH D A0 A1 Deg_0 by blast + qed + qed + qed + qed + then show ?thesis using assms by blast +qed + +text\Variant on induction on degree\ + +lemma poly_induct2: + assumes "p \ carrier P" + assumes Deg_0: "\p. p \ carrier P \ degree p = 0 \ Q p" + assumes IH: "\p. degree p > 0 \ p \ carrier P \ Q (trunc p) \ Q p" + shows "Q p" +proof(rule poly_induct) + show "p \ carrier P" + by (simp add: assms(1)) + show "\p. p \ carrier P \ degree p = 0 \ Q p" + by (simp add: Deg_0) + show "\p. (\q. q \ carrier P \ degree q < degree p \ Q q) \ p \ carrier P \ 0 < degree p \ Q p" + proof- + fix p + assume A0: "(\q. q \ carrier P \ degree q < degree p \ Q q)" + assume A1: " p \ carrier P" + assume A2: "0 < degree p" + show "Q p" + proof- + have "degree (trunc p) < degree p" + by (simp add: A1 A2 trunc_degree) + have "Q (trunc p)" + by (simp add: A0 A1 \degree (trunc p) < degree p\ trunc_closed) + then show ?thesis + by (simp add: A1 A2 IH) + qed + qed +qed + +text\Additive properties which are true for all monomials are true for all polynomials \ + +lemma poly_induct3: + assumes "p \ carrier P" + assumes add: "\p q. q \ carrier P \ p \ carrier P \ Q p \ Q q \ Q (p \\<^bsub>P\<^esub> q)" + assumes monom: "\a n. a \ carrier R \ Q (monom P a n)" + shows "Q p" + apply(rule poly_induct2) + apply (simp add: assms(1)) + apply (metis lcf_closed P_def coeff_simp deg_zero_impl_monom monom) + by (metis lcf_closed ltrm_closed add monom trunc_closed trunc_simps(1)) + +lemma poly_induct4: + assumes "p \ carrier P" + assumes add: "\p q. q \ carrier P \ p \ carrier P \ Q p \ Q q \ Q (p \\<^bsub>P\<^esub> q)" + assumes monom_zero: "\a. a \ carrier R \ Q (monom P a 0)" + assumes monom_Suc: "\a n. a \ carrier R \ Q (monom P a (Suc n))" + shows "Q p" + apply(rule poly_induct3) + using assms(1) apply auto[1] + using add apply blast + using monom_zero monom_Suc + by (metis P_def UP_ring.monom_zero UP_ring_axioms deg_monom deg_monom_le le_0_eq le_SucE zero_induct) + +lemma monic_monom_smult: + assumes "a \ carrier R" + shows "a \\<^bsub>P\<^esub> monom P \ n = monom P a n" + using assms + by (metis R.one_closed R.r_one monom_mult_smult) + +lemma poly_induct5: + assumes "p \ carrier P" + assumes add: "\p q. q \ carrier P \ p \ carrier P \ Q p \ Q q \ Q (p \\<^bsub>P\<^esub> q)" + assumes monic_monom: "\n. Q (monom P \ n)" + assumes smult: "\p a . a \ carrier R \ p \ carrier P \ Q p \ Q (a \\<^bsub>P\<^esub> p)" + shows "Q p" + apply(rule poly_induct3) + apply (simp add: assms(1)) + using add apply blast +proof- + fix a n assume A: "a \ carrier R" show "Q (monom P a n)" + using monic_monom[of n] smult[of a "monom P \ n"] monom_mult_smult[of a \ n] + by (simp add: A) +qed + +lemma poly_induct6: + assumes "p \ carrier P" + assumes monom: "\a n. a \ carrier R \ Q (monom P a 0)" + assumes plus_monom: "\a n p. a \ carrier R \ a \ \ \ p \ carrier P \ degree p < n \ Q p \ + Q(p \\<^bsub>P\<^esub> monom P a n)" + shows "Q p" + apply(rule poly_induct2) + using assms(1) apply auto[1] + apply (metis lcf_closed P_def coeff_simp deg_zero_impl_monom monom) + using plus_monom + by (metis lcf_closed P_def coeff_simp lcoeff_nonzero_deg nat_less_le trunc_closed trunc_degree trunc_simps(1)) + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Mapping a Polynomial to its Associated Ring Function\ +(**************************************************************************************************) +(**************************************************************************************************) + + +text\Turning a polynomial into a function on R:\ +definition to_function where +"to_function S f = (\s \ carrier S. eval S S (\x. x) s f)" + +context UP_cring +begin + +definition to_fun where +"to_fun f \ to_function R f" + +text\Explicit formula for evaluating a polynomial function:\ + +lemma to_fun_eval: + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun f x = eval R R (\x. x) x f" + using assms unfolding to_function_def to_fun_def + by auto + +lemma to_fun_formula: + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun f x = (\i \ {..degree f}. (f i) \ x [^] i)" +proof- + have "f \ carrier (UP R)" + using assms P_def by auto + then have "eval R R (\x. x) x f = (\\<^bsub>R\<^esub>i\{..deg R f}. (\x. x) (coeff (UP R) f i) \\<^bsub>R\<^esub> x [^]\<^bsub>R\<^esub> i)" + apply(simp add:UnivPoly.eval_def) done + then have "to_fun f x = (\\<^bsub>R\<^esub>i\{..deg R f}. (\x. x) (coeff (UP R) f i) \\<^bsub>R\<^esub> x [^]\<^bsub>R\<^esub> i)" + using to_function_def assms unfolding to_fun_def + by (simp add: to_function_def) + then show ?thesis + by(simp add: assms coeff_simp) +qed + +lemma eval_ring_hom: + assumes "a \ carrier R" + shows "eval R R (\x. x) a \ ring_hom P R" +proof- + have "(\x. x) \ ring_hom R R" + apply(rule ring_hom_memI) + apply auto done + then have "UP_pre_univ_prop R R (\x. x)" + using R_cring UP_pre_univ_propI by blast + then show ?thesis + by (simp add: P_def UP_pre_univ_prop.eval_ring_hom assms) +qed + +lemma to_fun_closed: + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun f x \ carrier R" + using assms to_fun_eval[of f x] eval_ring_hom[of x] + ring_hom_closed + by fastforce + +lemma to_fun_plus: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun (f \\<^bsub>P\<^esub> g) x = (to_fun f x) \ (to_fun g x)" + using assms to_fun_eval[of ] eval_ring_hom[of x] + by (simp add: ring_hom_add) + +lemma to_fun_mult: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun (f \\<^bsub>P\<^esub> g) x = (to_fun f x) \ (to_fun g x)" + using assms to_fun_eval[of ] eval_ring_hom[of x] + by (simp add: ring_hom_mult) + +lemma to_fun_ring_hom: + assumes "a \ carrier R" + shows "(\p. to_fun p a) \ ring_hom P R" + apply(rule ring_hom_memI) + apply (simp add: assms to_fun_closed) + apply (simp add: assms to_fun_mult) + apply (simp add: assms to_fun_plus) + using to_fun_eval[of "\\<^bsub>P\<^esub>" a] eval_ring_hom[of a] + ring_hom_closed + by (simp add: assms ring_hom_one) + +lemma ring_hom_uminus: + assumes "ring S" + assumes "f \ (ring_hom S R)" + assumes "a \ carrier S" + shows "f (\\<^bsub>S\<^esub> a) = \ (f a)" +proof- + have "f (a \\<^bsub>S\<^esub> a) = (f a) \ f (\\<^bsub>S\<^esub> a)" + unfolding a_minus_def + by (simp add: assms(1) assms(2) assms(3) ring.ring_simprules(3) ring_hom_add) + then have "(f a) \ f (\\<^bsub>S\<^esub> a) = \ " + by (metis R.ring_axioms a_minus_def assms(1) assms(2) assms(3) + ring.ring_simprules(16) ring_hom_zero) + then show ?thesis + by (metis (no_types, lifting) R.add.m_comm R.minus_equality assms(1) + assms(2) assms(3) ring.ring_simprules(3) ring_hom_closed) +qed + +lemma to_fun_minus: + assumes "f \ carrier P" + assumes "x \ carrier R" + shows "to_fun (\\<^bsub>P\<^esub>f) x = \ (to_fun f x)" + unfolding to_function_def to_fun_def + using eval_ring_hom[of x] assms + by (simp add: UP_ring ring_hom_uminus) + +lemma id_is_hom: +"ring_hom_cring R R (\x. x)" +proof(rule ring_hom_cringI) + show "cring R" + by (simp add: R_cring ) + show "cring R" + by (simp add: R_cring ) + show "(\x. x) \ ring_hom R R" + unfolding ring_hom_def + apply(auto) + done +qed + +lemma UP_pre_univ_prop_fact: +"UP_pre_univ_prop R R (\x. x)" + unfolding UP_pre_univ_prop_def + by (simp add: UP_cring_def R_cring id_is_hom) + +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\to-fun is a Ring Homomorphism from Polynomials to Functions\ + (**************************************************************************************************) + (**************************************************************************************************) + +context UP_cring +begin + +lemma to_fun_is_Fun: + assumes "x \ carrier P" + shows "to_fun x \ carrier (Fun R)" + apply(rule ring_functions.function_ring_car_memI) + unfolding ring_functions_def apply(simp add: R.ring_axioms) + using to_fun_closed assms apply auto[1] + unfolding to_function_def to_fun_def by auto + +lemma to_fun_Fun_mult: + assumes "x \ carrier P" + assumes "y \ carrier P" + shows "to_fun (x \\<^bsub>P\<^esub> y) = to_fun x \\<^bsub>function_ring (carrier R) R\<^esub> to_fun y" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + apply (simp add: R.ring_axioms ring_functions_def) + apply (simp add: assms(1) assms(2) to_fun_is_Fun) + apply (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_mult_closed ring_functions.intro to_fun_is_Fun) + by (simp add: R.ring_axioms assms(1) assms(2) ring_functions.function_mult_eval_car ring_functions.intro to_fun_is_Fun to_fun_mult) + +lemma to_fun_Fun_add: + assumes "x \ carrier P" + assumes "y \ carrier P" + shows "to_fun (x \\<^bsub>P\<^esub> y) = to_fun x \\<^bsub>function_ring (carrier R) R\<^esub> to_fun y" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + apply (simp add: R.ring_axioms ring_functions_def) + apply (simp add: assms(1) assms(2) to_fun_is_Fun) + apply (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_add_closed ring_functions.intro to_fun_is_Fun) + by (simp add: R.ring_axioms assms(1) assms(2) ring_functions.fun_add_eval_car ring_functions.intro to_fun_is_Fun to_fun_plus) + +lemma to_fun_Fun_one: +"to_fun \\<^bsub>P\<^esub> = \\<^bsub>Fun R\<^esub>" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + apply (simp add: R.ring_axioms ring_functions_def) + apply (simp add: to_fun_is_Fun) + apply (simp add: R.ring_axioms ring_functions.function_one_closed ring_functions_def) + using P_def R.ring_axioms UP_cring.eval_ring_hom UP_cring.to_fun_eval UP_cring_axioms UP_one_closed ring_functions.function_one_eval ring_functions.intro ring_hom_one + by fastforce + +lemma to_fun_Fun_zero: +"to_fun \\<^bsub>P\<^esub> = \\<^bsub>Fun R\<^esub>" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + apply (simp add: R.ring_axioms ring_functions_def) + apply (simp add: to_fun_is_Fun) + apply (simp add: R.ring_axioms ring_functions.function_zero_closed ring_functions_def) + using P_def R.ring_axioms UP_cring.eval_ring_hom UP_cring.to_fun_eval UP_cring_axioms UP_zero_closed ring_functions.function_zero_eval ring_functions.intro ring_hom_zero + by (metis UP_ring eval_ring_hom) + +lemma to_fun_function_ring_hom: +"to_fun \ ring_hom P (Fun R)" + apply(rule ring_hom_memI) + using to_fun_is_Fun apply auto[1] + apply (simp add: to_fun_Fun_mult) + apply (simp add: to_fun_Fun_add) + by (simp add: to_fun_Fun_one) + +lemma(in UP_cring) to_fun_one: + assumes "a \ carrier R" + shows "to_fun \\<^bsub>P\<^esub> a = \" + using assms to_fun_Fun_one + by (metis P_def UP_cring.to_fun_eval UP_cring_axioms UP_one_closed eval_ring_hom ring_hom_one) + +lemma(in UP_cring) to_fun_zero: + assumes "a \ carrier R" + shows "to_fun \\<^bsub>P\<^esub> a = \" + by (simp add: assms R.ring_axioms ring_functions.function_zero_eval ring_functions.intro to_fun_Fun_zero) + +lemma(in UP_cring) to_fun_nat_pow: + assumes "h \ carrier (UP R)" + assumes "a \ carrier R" + shows "to_fun (h[^]\<^bsub>UP R\<^esub>(n::nat)) a = (to_fun h a)[^]n" + apply(induction n) + using assms to_fun_one + apply (metis P.nat_pow_0 P_def R.nat_pow_0) + using assms to_fun_mult P.nat_pow_closed P_def by auto + +lemma(in UP_cring) to_fun_finsum: + assumes "finite (Y::'d set)" + assumes "f \ UNIV \ carrier (UP R)" + assumes "t \ carrier R" + shows "to_fun (finsum (UP R) f Y) t = finsum R (\i. (to_fun (f i) t)) Y" +proof(rule finite.induct[of Y]) + show "finite Y" + using assms by blast + show "to_fun (finsum (UP R) f {}) t = (\i\{}. to_fun (f i) t)" + using P.finsum_empty[of f] assms unfolding P_def R.finsum_empty + using P_def to_fun_zero by presburger + show "\A a. finite A \ + to_fun (finsum (UP R) f A) t = (\i\A. to_fun (f i) t) \ to_fun (finsum (UP R) f (insert a A)) t = (\i\insert a A. to_fun (f i) t)" + proof- + fix A :: "'d set" fix a + assume A: "finite A" "to_fun (finsum (UP R) f A) t = (\i\A. to_fun (f i) t)" + show "to_fun (finsum (UP R) f (insert a A)) t = (\i\insert a A. to_fun (f i) t)" + proof(cases "a \ A") + case True + then show ?thesis using A + by (metis insert_absorb) + next + case False + have 0: "finsum (UP R) f (insert a A) = f a \\<^bsub>UP R\<^esub> finsum (UP R) f A" + using A False finsum_insert[of A a f] assms unfolding P_def by blast + have 1: "to_fun (f a \\<^bsub>P\<^esub>finsum (UP R) f A ) t = to_fun (f a) t \ to_fun (finsum (UP R) f A) t" + apply(rule to_fun_plus[of "finsum (UP R) f A" "f a" t]) + using assms(2) finsum_closed[of f A] A unfolding P_def apply blast + using P_def assms apply blast + using assms by blast + have 2: "to_fun (f a \\<^bsub>P\<^esub>finsum (UP R) f A ) t = to_fun (f a) t \ (\i\A. to_fun (f i) t)" + unfolding 1 A by blast + have 3: "(\i\insert a A. to_fun (f i) t) = to_fun (f a) t \ (\i\A. to_fun (f i) t)" + apply(rule R.finsum_insert, rule A, rule False) + using to_fun_closed assms unfolding P_def apply blast + apply(rule to_fun_closed) using assms unfolding P_def apply blast using assms by blast + show ?thesis + unfolding 0 unfolding 3 using 2 unfolding P_def by blast + qed + qed +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Inclusion of a Ring into its Polynomials Ring via Constants\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition to_polynomial where +"to_polynomial R = (\a. monom (UP R) a 0)" + +context UP_cring +begin + +abbreviation(input) to_poly where +"to_poly \ to_polynomial R" + +lemma to_poly_mult_simp: + assumes "b \ carrier R" + assumes "f \ carrier (UP R)" + shows "(to_polynomial R b) \\<^bsub>UP R\<^esub> f = b \\<^bsub>UP R\<^esub> f" + "f \\<^bsub>UP R\<^esub> (to_polynomial R b) = b \\<^bsub>UP R\<^esub> f" + unfolding to_polynomial_def + using assms P_def monom_mult_is_smult apply auto[1] + using UP_cring.UP_m_comm UP_cring_axioms UP_ring.monom_closed + UP_ring.monom_mult_is_smult UP_ring_axioms assms(1) assms(2) + by fastforce + +lemma to_fun_to_poly: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_fun (to_poly a) b = a" + unfolding to_function_def to_fun_def to_polynomial_def + by (simp add: UP_pre_univ_prop.eval_const UP_pre_univ_prop_fact assms(1) assms(2)) + +lemma to_poly_inverse: + assumes "f \ carrier P" + assumes "degree f = 0" + shows "f = to_poly (f 0)" + using P_def assms(1) assms(2) + by (metis ltrm_deg_0 to_polynomial_def) + +lemma to_poly_closed: + assumes "a \ carrier R" + shows "to_poly a \ carrier P" + by (metis P_def assms monom_closed to_polynomial_def) + +lemma degree_to_poly[simp]: + assumes "a \ carrier R" + shows "degree (to_poly a) = 0" + by (metis P_def assms deg_const to_polynomial_def) + +lemma to_poly_is_ring_hom: +"to_poly \ ring_hom R P" + unfolding to_polynomial_def + unfolding P_def + using UP_ring.const_ring_hom[of R] + UP_ring_axioms by simp + +lemma to_poly_add: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_poly (a \ b) = to_poly a \\<^bsub>P\<^esub> to_poly b" + by (simp add: assms(1) assms(2) ring_hom_add to_poly_is_ring_hom) + +lemma to_poly_mult: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_poly (a \ b) = to_poly a \\<^bsub>P\<^esub> to_poly b" + by (simp add: assms(1) assms(2) ring_hom_mult to_poly_is_ring_hom) + +lemma to_poly_minus: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_poly (a \ b) = to_poly a \\<^bsub>P\<^esub> to_poly b" + by (metis P.minus_eq P_def R.add.inv_closed R.ring_axioms UP_ring.monom_add + UP_ring_axioms assms(1) assms(2) monom_a_inv ring.ring_simprules(14) to_polynomial_def) + +lemma to_poly_a_inv: + assumes "a \ carrier R" + shows "to_poly (\ a) = \\<^bsub>P\<^esub> to_poly a" + by (metis P_def assms monom_a_inv to_polynomial_def) + +lemma to_poly_nat_pow: + assumes "a \ carrier R" + shows "(to_poly a) [^]\<^bsub>P\<^esub> (n::nat)= to_poly (a[^]n)" + using assms UP_cring UP_cring_axioms UP_cring_def UnivPoly.ring_hom_cringI ring_hom_cring.hom_pow to_poly_is_ring_hom + by fastforce + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Polynomial Substitution\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition compose where +"compose R f g = eval R (UP R) (to_polynomial R) g f" + +abbreviation(in UP_cring)(input) sub (infixl "of" 70) where +"sub f g \ compose R f g" + +definition rev_compose where +"rev_compose R = eval R (UP R) (to_polynomial R)" + +abbreviation(in UP_cring)(input) rev_sub where +"rev_sub \ rev_compose R" + +context UP_cring +begin + +lemma sub_rev_sub: +"sub f g = rev_sub g f" + unfolding compose_def rev_compose_def + by simp + +lemma(in UP_cring) to_poly_UP_pre_univ_prop: +"UP_pre_univ_prop R P to_poly" +proof + show "to_poly \ ring_hom R P" + by (simp add: to_poly_is_ring_hom) +qed + +lemma rev_sub_is_hom: + assumes "g \ carrier P" + shows "rev_sub g \ ring_hom P P" + unfolding rev_compose_def + using to_poly_UP_pre_univ_prop assms(1) UP_pre_univ_prop.eval_ring_hom[of R P to_poly g] + unfolding P_def apply auto + done + +lemma rev_sub_closed: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "rev_sub q p \ carrier P" + using rev_sub_is_hom[of q] assms ring_hom_closed[of "rev_sub q" P P p] by auto + +lemma sub_closed: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "sub q p \ carrier P" + by (simp add: assms(1) assms(2) rev_sub_closed sub_rev_sub) + +lemma rev_sub_add: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "h \carrier P" + shows "rev_sub g (f \\<^bsub>P\<^esub> h) = (rev_sub g f) \\<^bsub>P\<^esub> (rev_sub g h)" + using rev_sub_is_hom assms ring_hom_add by fastforce + +lemma sub_add: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "h \carrier P" + shows "((f \\<^bsub>P\<^esub> h) of g) = ((f of g) \\<^bsub>P\<^esub> (h of g))" + by (simp add: assms(1) assms(2) assms(3) rev_sub_add sub_rev_sub) + +lemma rev_sub_mult: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "h \carrier P" + shows "rev_sub g (f \\<^bsub>P\<^esub> h) = (rev_sub g f) \\<^bsub>P\<^esub> (rev_sub g h)" + using rev_sub_is_hom assms ring_hom_mult by fastforce + +lemma sub_mult: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "h \carrier P" + shows "((f \\<^bsub>P\<^esub> h) of g) = ((f of g) \\<^bsub>P\<^esub> (h of g))" + by (simp add: assms(1) assms(2) assms(3) rev_sub_mult sub_rev_sub) + +lemma sub_monom: + assumes "g \ carrier (UP R)" + assumes "a \ carrier R" + shows "sub (monom (UP R) a n) g = to_poly a \\<^bsub>UP R\<^esub> (g[^]\<^bsub>UP R\<^esub> (n::nat))" + "sub (monom (UP R) a n) g = a \\<^bsub>UP R\<^esub> (g[^]\<^bsub>UP R\<^esub> (n::nat))" + apply (simp add: UP_cring.to_poly_UP_pre_univ_prop UP_cring_axioms + UP_pre_univ_prop.eval_monom assms(1) assms(2) Cring_Poly.compose_def) + by (metis P_def UP_cring.to_poly_mult_simp(1) UP_cring_axioms UP_pre_univ_prop.eval_monom + UP_ring assms(1) assms(2) Cring_Poly.compose_def monoid.nat_pow_closed ring_def to_poly_UP_pre_univ_prop) + +text\Subbing into a constant does nothing\ + +lemma rev_sub_to_poly: + assumes "g \ carrier P" + assumes "a \ carrier R" + shows "rev_sub g (to_poly a) = to_poly a" + unfolding to_polynomial_def rev_compose_def + using to_poly_UP_pre_univ_prop + unfolding to_polynomial_def + using P_def UP_pre_univ_prop.eval_const assms(1) assms(2) by fastforce + +lemma sub_to_poly: + assumes "g \ carrier P" + assumes "a \ carrier R" + shows "(to_poly a) of g = to_poly a" + by (simp add: assms(1) assms(2) rev_sub_to_poly sub_rev_sub) + +lemma sub_const: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = 0" + shows "f of g = f" + by (metis lcf_closed assms(1) assms(2) assms(3) sub_to_poly to_poly_inverse) + +text\Substitution into a monomial\ + +lemma monom_sub: + assumes "a \ carrier R" + assumes "g \ carrier P" + shows "(monom P a n) of g = a \\<^bsub>P\<^esub> g[^]\<^bsub>P\<^esub> n" + unfolding compose_def + using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g n] to_poly_UP_pre_univ_prop + unfolding P_def + using P.nat_pow_closed P_def to_poly_mult_simp(1) + by (simp add: to_poly_mult_simp(1) UP_cring_axioms) + +lemma(in UP_cring) cring_sub_monom_bound: + assumes "a \ carrier R" + assumes "a \\" + assumes "f = monom P a n" + assumes "g \ carrier P" + shows "degree (f of g) \ n*(degree g)" +proof- + have "f of g = (to_poly a) \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n)" + unfolding compose_def + using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop + unfolding P_def + by blast + then show ?thesis + by (smt P.nat_pow_closed assms(1) assms(4) cring_pow_deg_bound deg_mult_ring + degree_to_poly le_trans plus_nat.add_0 to_poly_closed) +qed + +lemma(in UP_cring) cring_sub_monom: + assumes "a \ carrier R" + assumes "a \\" + assumes "f = monom P a n" + assumes "g \ carrier P" + assumes "a \ (lcf g [^] n) \ \" + shows "degree (f of g) = n*(degree g)" +proof- + have 0: "f of g = (to_poly a) \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n)" + unfolding compose_def + using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop + unfolding P_def + by blast + have 1: "lcf (to_poly a) \ lcf (g [^]\<^bsub>P\<^esub> n) \ \" + using assms + by (smt P.nat_pow_closed P_def R.nat_pow_closed R.r_null cring_pow_ltrm lcf_closed lcf_ltrm lcf_monom monom_pow to_polynomial_def) + then show ?thesis + using 0 1 assms cring_pow_deg[of g n] cring_deg_mult[of "to_poly a" "g[^]\<^bsub>P\<^esub>n"] + by (metis P.nat_pow_closed R.r_null add.right_neutral degree_to_poly to_poly_closed) +qed + +lemma(in UP_domain) sub_monom: + assumes "a \ carrier R" + assumes "a \\" + assumes "f = monom P a n" + assumes "g \ carrier P" + shows "degree (f of g) = n*(degree g)" +proof- + have "f of g = (to_poly a) \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n)" + unfolding compose_def + using assms UP_pre_univ_prop.eval_monom[of R P to_poly a g] to_poly_UP_pre_univ_prop + unfolding P_def + by blast + then show ?thesis using deg_pow deg_mult + by (metis P.nat_pow_closed P_def assms(1) assms(2) + assms(4) deg_smult monom_mult_is_smult to_polynomial_def) +qed + +text\Subbing a constant into a polynomial yields a constant\ +lemma sub_in_const: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree g = 0" + shows "degree (f of g) = 0" +proof- + have "\n. (\p. p \ carrier P \ degree p \ n \ degree (p of g) = 0)" + proof- + fix n + show "\p. p \ carrier P \ degree p \ n \ degree (p of g) = 0" + proof(induction n) + case 0 + then show ?case + by (simp add: assms(1) sub_const) + next + case (Suc n) + fix n + assume IH: "\p. p \ carrier P \ degree p \ n \ degree (p of g) = 0" + show "\p. p \ carrier P \ degree p \ (Suc n) \ degree (p of g) = 0" + proof- + fix p + assume A0: "p \ carrier P" + assume A1: "degree p \ (Suc n)" + show "degree (p of g) = 0" + proof(cases "degree p < Suc n") + case True + then show ?thesis using IH + using A0 by auto + next + case False + then have D: "degree p = Suc n" + by (simp add: A1 nat_less_le) + show ?thesis + proof- + have P0: "degree ((trunc p) of g) = 0" using IH + by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc) + have P1: "degree ((ltrm p) of g) = 0" + proof- + obtain a n where an_def: "ltrm p = monom P a n \ a \ carrier R" + unfolding leading_term_def + using A0 P_def cfs_closed by blast + obtain b where b_def: "g = monom P b 0 \ b \ carrier R" + using assms deg_zero_impl_monom coeff_closed + by blast + have 0: " monom P b 0 [^]\<^bsub>P\<^esub> n = monom P (b[^]n) 0" + apply(induction n) + apply fastforce[1] + proof- fix n::nat assume IH: "monom P b 0 [^]\<^bsub>P\<^esub> n = monom P (b [^] n) 0" + have "monom P b 0 [^]\<^bsub>P\<^esub> Suc n = (monom P (b[^]n) 0) \\<^bsub>P\<^esub> monom P b 0" + using IH by simp + then have "monom P b 0 [^]\<^bsub>P\<^esub> Suc n = (monom P ((b[^]n)\b) 0)" + using b_def + by (simp add: monom_mult_is_smult monom_mult_smult) + then show "monom P b 0 [^]\<^bsub>P\<^esub> Suc n = monom P (b [^] Suc n) 0 " + by simp + qed + + then have 0: "a \\<^bsub>P\<^esub> monom P b 0 [^]\<^bsub>P\<^esub> n = monom P (a \ b[^]n) 0" + by (simp add: an_def b_def monom_mult_smult) + + + then show ?thesis using monom_sub[of a "monom P b 0" n] assms an_def + by (simp add: \\a \ carrier R; monom P b 0 \ carrier P\ \ monom P a n of monom P b 0 = a \\<^bsub>P\<^esub> monom P b 0 [^]\<^bsub>P\<^esub> n\ b_def) + qed + have P2: "p of g = (trunc p of g) \\<^bsub>P\<^esub> ((ltrm p) of g)" + by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed) + then show ?thesis + using P0 P1 P2 deg_add[of "trunc p of g" "ltrm p of g"] + by (metis A0 assms(1) le_0_eq ltrm_closed max_0R sub_closed trunc_closed) + qed + qed + qed + qed + qed + then show ?thesis + using assms(2) by blast +qed + +lemma (in UP_cring) cring_sub_deg_bound: + assumes "g \ carrier P" + assumes "f \ carrier P" + shows "degree (f of g) \ degree f * degree g" +proof- + have "\n. \ p. p \ carrier P \ (degree p) \ n \ degree (p of g) \ degree p * degree g" + proof- + fix n::nat + show "\ p. p \ carrier P \ (degree p) \ n \ degree (p of g) \ degree p * degree g" + proof(induction n) + case 0 + then have B0: "degree p = 0" by auto + then show ?case using sub_const[of g p] + by (simp add: "0.prems"(1) assms(1)) + next + case (Suc n) + fix n + assume IH: "(\p. p \ carrier P \ degree p \ n \ degree (p of g) \ degree p * degree g)" + show " p \ carrier P \ degree p \ Suc n \ degree (p of g) \ degree p * degree g" + proof- + assume A0: "p \ carrier P" + assume A1: "degree p \ Suc n" + show ?thesis + proof(cases "degree p < Suc n") + case True + then show ?thesis using IH + by (simp add: A0) + next + case False + then have D: "degree p = Suc n" + using A1 by auto + have P0: "(p of g) = ((trunc p) of g) \\<^bsub>P\<^esub> ((ltrm p) of g)" + by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed) + have P1: "degree ((trunc p) of g) \ (degree (trunc p))*(degree g)" + using IH by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc) + have P2: "degree ((ltrm p) of g) \ (degree p) * degree g" + using A0 D P_def UP_cring_axioms assms(1) + by (metis False cfs_closed coeff_simp cring_sub_monom_bound deg_zero lcoeff_nonzero2 less_Suc_eq_0_disj) + then show ?thesis + proof(cases "degree g = 0") + case True + then show ?thesis + by (simp add: Suc(2) assms(1) sub_in_const) + next + case F: False + then show ?thesis + proof- + have P3: "degree ((trunc p) of g) \ n*degree g" + using A0 False D P1 P2 IH[of "trunc p"] trunc_degree[of p] + proof - + { assume "degree (trunc p) < degree p" + then have "degree (trunc p) \ n" + using D by auto + then have ?thesis + by (meson P1 le_trans mult_le_cancel2) } + then show ?thesis + by (metis (full_types) A0 D Suc_mult_le_cancel1 nat_mult_le_cancel_disj trunc_degree) + qed + then have P3': "degree ((trunc p) of g) < (degree p)*degree g" + using F D by auto + have P4: "degree (ltrm p of g) \ (degree p)*degree g" + using cring_sub_monom_bound D P2 + by auto + then show ?thesis + using D P0 P1 P3 P4 A0 P3' assms(1) bound_deg_sum less_imp_le_nat + ltrm_closed sub_closed trunc_closed + by metis + qed + qed + qed + qed + qed + qed + then show ?thesis + using assms(2) by blast +qed + +lemma (in UP_cring) cring_sub_deg: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "lcf f \ (lcf g [^] (degree f)) \ \" + shows "degree (f of g) = degree f * degree g" +proof- + have 0: "f of g = (trunc f of g) \\<^bsub>P\<^esub> ((ltrm f) of g)" + by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed) + have 1: "lcf f \ \" + using assms cring.cring_simprules(26) lcf_closed + by auto + have 2: "degree ((ltrm f) of g) = degree f * degree g" + using 0 1 assms cring_sub_monom[of "lcf f" "ltrm f" "degree f" g] lcf_closed lcf_ltrm + by blast + show ?thesis + apply(cases "degree f = 0") + apply (simp add: assms(1) assms(2)) + apply(cases "degree g = 0") + apply (simp add: assms(1) assms(2) sub_in_const) + using 0 1 assms cring_sub_deg_bound[of g "trunc f"] trunc_degree[of f] + using sub_const apply auto[1] + apply(cases "degree g = 0") + using 0 1 assms cring_sub_deg_bound[of g "trunc f"] trunc_degree[of f] + using sub_in_const apply fastforce + unfolding 0 using 1 2 + by (smt "0" ltrm_closed \\f \ carrier P; 0 < deg R f\ \ deg R (Cring_Poly.truncate R f) < deg R f\ + assms(1) assms(2) cring_sub_deg_bound degree_of_sum_diff_degree equal_deg_sum + le_eq_less_or_eq mult_less_cancel2 nat_neq_iff neq0_conv sub_closed trunc_closed) +qed + +lemma (in UP_domain) sub_deg0: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "g \ \\<^bsub>P\<^esub>" + assumes "f \ \\<^bsub>P\<^esub>" + shows "degree (f of g) = degree f * degree g" +proof- + have "\n. \ p. p \ carrier P \ (degree p) \ n \ degree (p of g) = degree p * degree g" + proof- + fix n::nat + show "\ p. p \ carrier P \ (degree p) \ n \ degree (p of g) = degree p * degree g" + proof(induction n) + case 0 + then have B0: "degree p = 0" by auto + then show ?case using sub_const[of g p] + by (simp add: "0.prems"(1) assms(1)) + next + case (Suc n) + fix n + assume IH: "(\p. p \ carrier P \ degree p \ n \ degree (p of g) = degree p * degree g)" + show " p \ carrier P \ degree p \ Suc n \ degree (p of g) = degree p * degree g" + proof- + assume A0: "p \ carrier P" + assume A1: "degree p \ Suc n" + show ?thesis + proof(cases "degree p < Suc n") + case True + then show ?thesis using IH + by (simp add: A0) + next + case False + then have D: "degree p = Suc n" + using A1 by auto + have P0: "(p of g) = ((trunc p) of g) \\<^bsub>P\<^esub> ((ltrm p) of g)" + by (metis A0 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed) + have P1: "degree ((trunc p) of g) = (degree (trunc p))*(degree g)" + using IH by (metis A0 D less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc) + have P2: "degree ((ltrm p) of g) = (degree p) * degree g" + using A0 D P_def UP_domain.sub_monom UP_cring_axioms assms(1) + by (metis False UP_domain_axioms UP_ring.coeff_simp UP_ring.lcoeff_nonzero2 UP_ring_axioms cfs_closed deg_nzero_nzero less_Suc_eq_0_disj) + + then show ?thesis + proof(cases "degree g = 0") + case True + then show ?thesis + by (simp add: Suc(2) assms(1) sub_in_const) + next + case False + then show ?thesis + proof- + have P3: "degree ((trunc p) of g) < degree ((ltrm p) of g)" + using False D P1 P2 + by (metis (no_types, lifting) A0 mult.commute mult_right_cancel + nat_less_le nat_mult_le_cancel_disj trunc_degree zero_less_Suc) + then show ?thesis + by (simp add: A0 ltrm_closed P0 P2 assms(1) equal_deg_sum sub_closed trunc_closed) + qed + qed + qed + qed + qed + qed + then show ?thesis + using assms(2) by blast +qed + +lemma(in UP_domain) sub_deg: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "g \ \\<^bsub>P\<^esub>" + shows "degree (f of g) = degree f * degree g" +proof(cases "f = \\<^bsub>P\<^esub>") + case True + then show ?thesis + using assms(1) sub_const by auto +next + case False + then show ?thesis + by (simp add: assms(1) assms(2) assms(3) sub_deg0) +qed + +lemma(in UP_cring) cring_ltrm_sub: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree g > 0" + assumes "lcf f \ (lcf g [^] (degree f)) \ \" + shows "ltrm (f of g) = ltrm ((ltrm f) of g)" +proof- + have P0: "degree (f of g) = degree ((ltrm f) of g)" + using assms(1) assms(2) assms(4) cring_sub_deg lcf_eq ltrm_closed deg_ltrm + by auto + have P1: "f of g = ((trunc f) of g) \\<^bsub>P\<^esub>((ltrm f) of g)" + by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed) + then show ?thesis + proof(cases "degree f = 0") + case True + then show ?thesis + using ltrm_deg_0 assms(2) by auto + next + case False + have P2: "degree (f of g) = degree f * degree g" + by (simp add: assms(1) assms(2) assms(4) cring_sub_deg) + then have P3: "degree ((trunc f) of g) < degree ((ltrm f) of g)" + using False P0 P1 P_def UP_cring.sub_closed trunc_closed UP_cring_axioms + UP_ring.degree_of_sum_diff_degree UP_ring.ltrm_closed UP_ring_axioms assms(1) + assms(2) assms(4) cring_sub_deg_bound le_antisym less_imp_le_nat less_nat_zero_code + mult_right_le_imp_le nat_neq_iff trunc_degree + by (smt assms(3)) + then show ?thesis using P0 P1 P2 + by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm assms(1) assms(2) sub_closed trunc_closed) + qed +qed + +lemma(in UP_domain) ltrm_sub: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree g > 0" + shows "ltrm (f of g) = ltrm ((ltrm f) of g)" +proof- + have P0: "degree (f of g) = degree ((ltrm f) of g)" + using sub_deg + by (metis ltrm_closed assms(1) assms(2) assms(3) deg_zero deg_ltrm nat_neq_iff) + have P1: "f of g = ((trunc f) of g) \\<^bsub>P\<^esub>((ltrm f) of g)" + by (metis assms(1) assms(2) ltrm_closed rev_sub_add sub_rev_sub trunc_simps(1) trunc_closed) + then show ?thesis + proof(cases "degree f = 0") + case True + then show ?thesis + using ltrm_deg_0 assms(2) by auto + next + case False + then have P2: "degree ((trunc f) of g) < degree ((ltrm f) of g)" + using sub_deg + by (metis (no_types, lifting) ltrm_closed assms(1) assms(2) assms(3) deg_zero + deg_ltrm mult_less_cancel2 neq0_conv trunc_closed trunc_degree) + then show ?thesis + using P0 P1 P2 + by (metis (no_types, lifting) ltrm_closed ltrm_of_sum_diff_degree P.add.m_comm assms(1) assms(2) sub_closed trunc_closed) + qed +qed + +lemma(in UP_cring) cring_lcf_of_sub_in_ltrm: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + assumes "(lcf f) \ ((lcf g)[^]n) \\" + shows "lcf ((ltrm f) of g) = (lcf f) \ ((lcf g)[^]n)" + by (metis (no_types, lifting) P.nat_pow_closed P_def R.r_null UP_cring.monom_sub UP_cring_axioms + assms(1) assms(2) assms(3) assms(5) cfs_closed cring_lcf_pow cring_lcf_scalar_mult) + +lemma(in UP_domain) lcf_of_sub_in_ltrm: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + shows "lcf ((ltrm f) of g) = (lcf f) \ ((lcf g)[^]n)" +proof(cases "degree f = 0") + case True + then show ?thesis + using ltrm_deg_0 assms(1) assms(2) assms(3) cfs_closed + by (simp add: sub_const) +next + case False + then show ?thesis + proof- + have P0: "(ltrm f) of g = (to_poly (lcf f)) \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n)" + unfolding compose_def + using assms UP_pre_univ_prop.eval_monom[of R P to_poly "(lcf f)" g n] to_poly_UP_pre_univ_prop + unfolding P_def + using P_def cfs_closed by blast + have P1: "(ltrm f) of g = (lcf f) \\<^bsub>P\<^esub>(g[^]\<^bsub>P\<^esub>n)" + using P0 P.nat_pow_closed + by (simp add: assms(1) assms(2) assms(3) cfs_closed monom_sub) + have P2: "ltrm ((ltrm f) of g) = (ltrm (to_poly (lcf f))) \\<^bsub>P\<^esub> (ltrm (g[^]\<^bsub>P\<^esub>n))" + using P0 ltrm_mult P.nat_pow_closed P_def assms(1) assms(2) + to_poly_closed + by (simp add: cfs_closed) + have P3: "ltrm ((ltrm f) of g) = (to_poly (lcf f)) \\<^bsub>P\<^esub> (ltrm (g[^]\<^bsub>P\<^esub>n))" + using P2 ltrm_deg_0 assms(2) to_poly_closed + by (simp add: cfs_closed) + have P4: "ltrm ((ltrm f) of g) = (lcf f) \\<^bsub>P\<^esub> ((ltrm g)[^]\<^bsub>P\<^esub>n)" + using P.nat_pow_closed P1 P_def assms(1) assms(2) ltrm_pow0 ltrm_smult + by (simp add: cfs_closed) + have P5: "lcf ((ltrm f) of g) = (lcf f) \ (lcf ((ltrm g)[^]\<^bsub>P\<^esub>n))" + using lcf_scalar_mult P4 by (metis P.nat_pow_closed P1 cfs_closed + UP_smult_closed assms(1) assms(2) assms(3) lcf_eq ltrm_closed sub_rev_sub) + show ?thesis + using P5 ltrm_pow lcf_pow assms(1) lcf_eq ltrm_closed by presburger + qed +qed + +lemma(in UP_cring) cring_ltrm_of_sub_in_ltrm: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + assumes "(lcf f) \ ((lcf g)[^]n) \\" + shows "ltrm ((ltrm f) of g) = (lcf f) \\<^bsub>P\<^esub> ((ltrm g)[^]\<^bsub>P\<^esub>n)" + by (smt lcf_eq ltrm_closed R.nat_pow_closed R.r_null assms(1) assms(2) assms(3) + assms(4) assms(5) cfs_closed cring_lcf_of_sub_in_ltrm cring_lcf_pow cring_pow_ltrm + cring_pow_deg cring_sub_deg deg_zero deg_ltrm monom_mult_smult neq0_conv) + +lemma(in UP_domain) ltrm_of_sub_in_ltrm: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + shows "ltrm ((ltrm f) of g) = (lcf f) \\<^bsub>P\<^esub> ((ltrm g)[^]\<^bsub>P\<^esub>n)" + by (smt Group.nat_pow_0 lcf_of_sub_in_ltrm lcf_pow lcf_scalar_mult ltrm_closed + ltrm_pow0 ltrm_smult P.nat_pow_closed P_def UP_ring.monom_one UP_ring_axioms assms(1) + assms(2) assms(3) assms(4) cfs_closed coeff_simp deg_const deg_nzero_nzero deg_pow + deg_smult deg_ltrm lcoeff_nonzero2 nat_less_le sub_deg) + +text\formula for the leading term of a composition \ + +lemma(in UP_domain) cring_ltrm_of_sub: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + assumes "(lcf f) \ ((lcf g)[^]n) \\" + shows "ltrm (f of g) = (lcf f) \\<^bsub>P\<^esub> ((ltrm g)[^]\<^bsub>P\<^esub>n)" + using ltrm_of_sub_in_ltrm ltrm_sub assms(1) assms(2) assms(3) assms(4) by presburger + +lemma(in UP_domain) ltrm_of_sub: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "degree f = n" + assumes "degree g > 0" + shows "ltrm (f of g) = (lcf f) \\<^bsub>P\<^esub> ((ltrm g)[^]\<^bsub>P\<^esub>n)" + using ltrm_of_sub_in_ltrm ltrm_sub assms(1) assms(2) assms(3) assms(4) by presburger + +text\subtitution is associative\ + +lemma sub_assoc_monom: + assumes "f \ carrier P" + assumes "q \ carrier P" + assumes "r \ carrier P" + shows "(ltrm f) of (q of r) = ((ltrm f) of q) of r" +proof- + obtain n where n_def: "n = degree f" + by simp + obtain a where a_def: "a \ carrier R \ (ltrm f) = monom P a n" + using assms(1) cfs_closed n_def by blast + have LHS: "(ltrm f) of (q of r) = a \\<^bsub>P\<^esub> (q of r)[^]\<^bsub>P\<^esub> n" + by (metis P.nat_pow_closed P_def UP_pre_univ_prop.eval_monom a_def assms(2) + assms(3) compose_def monom_mult_is_smult sub_closed to_poly_UP_pre_univ_prop to_polynomial_def) + have RHS0: "((ltrm f) of q) of r = (a \\<^bsub>P\<^esub> q[^]\<^bsub>P\<^esub> n)of r" + by (metis P.nat_pow_closed P_def UP_pre_univ_prop.eval_monom a_def + assms(2) compose_def monom_mult_is_smult to_poly_UP_pre_univ_prop to_polynomial_def) + have RHS1: "((ltrm f) of q) of r = ((to_poly a) \\<^bsub>P\<^esub> q[^]\<^bsub>P\<^esub> n)of r" + using RHS0 by (metis P.nat_pow_closed P_def a_def + assms(2) monom_mult_is_smult to_polynomial_def) + have RHS2: "((ltrm f) of q) of r = ((to_poly a) of r) \\<^bsub>P\<^esub> (q[^]\<^bsub>P\<^esub> n of r)" + using RHS1 a_def assms(2) assms(3) sub_mult to_poly_closed by auto + have RHS3: "((ltrm f) of q) of r = (to_poly a) \\<^bsub>P\<^esub> (q[^]\<^bsub>P\<^esub> n of r)" + using RHS2 a_def assms(3) sub_to_poly by auto + have RHS4: "((ltrm f) of q) of r = a \\<^bsub>P\<^esub> ((q[^]\<^bsub>P\<^esub> n)of r)" + using RHS3 + by (metis P.nat_pow_closed P_def a_def assms(2) assms(3) + monom_mult_is_smult sub_closed to_polynomial_def) + have "(q of r)[^]\<^bsub>P\<^esub> n = ((q[^]\<^bsub>P\<^esub> n)of r)" + apply(induction n) + apply (metis Group.nat_pow_0 P.ring_simprules(6) assms(3) deg_one sub_const) + by (simp add: assms(2) assms(3) sub_mult) + then show ?thesis using RHS4 LHS by simp +qed + +lemma sub_assoc: + assumes "f \ carrier P" + assumes "q \ carrier P" + assumes "r \ carrier P" + shows "f of (q of r) = (f of q) of r" +proof- + have "\ n. \ p. p \ carrier P \ degree p \ n \ p of (q of r) = (p of q) of r" + proof- + fix n + show "\ p. p \ carrier P \ degree p \ n \ p of (q of r) = (p of q) of r" + proof(induction n) + case 0 + then have deg_p: "degree p = 0" + by blast + then have B0: "p of (q of r) = p" + using sub_const[of "q of r" p] assms "0.prems"(1) sub_closed by blast + have B1: "(p of q) of r = p" + proof- + have p0: "p of q = p" + using deg_p 0 assms(2) + by (simp add: P_def UP_cring.sub_const UP_cring_axioms) + show ?thesis + unfolding p0 using deg_p 0 assms(3) + by (simp add: P_def UP_cring.sub_const UP_cring_axioms) + qed + then show "p of (q of r) = (p of q) of r" using B0 B1 by auto + next + case (Suc n) + fix n + assume IH: "\ p. p \ carrier P \ degree p \ n \ p of (q of r) = (p of q) of r" + then show "\ p. p \ carrier P \ degree p \ Suc n \ p of (q of r) = (p of q) of r" + proof- + fix p + assume A0: " p \ carrier P " + assume A1: "degree p \ Suc n" + show "p of (q of r) = (p of q) of r" + proof(cases "degree p < Suc n") + case True + then show ?thesis using A0 A1 IH by auto + next + case False + then have "degree p = Suc n" + using A1 by auto + have I0: "p of (q of r) = ((trunc p) \\<^bsub>P\<^esub> (ltrm p)) of (q of r)" + using A0 trunc_simps(1) by auto + have I1: "p of (q of r) = ((trunc p) of (q of r)) \\<^bsub>P\<^esub> ((ltrm p) of (q of r))" + using I0 sub_add + by (simp add: A0 assms(2) assms(3) ltrm_closed rev_sub_closed sub_rev_sub trunc_closed) + have I2: "p of (q of r) = (((trunc p) of q) of r) \\<^bsub>P\<^esub> (((ltrm p) of q) of r)" + using IH[of "trunc p"] sub_assoc_monom[of p q r] + by (metis A0 I1 \degree p = Suc n\ assms(2) assms(3) + less_Suc_eq_le trunc_degree trunc_closed zero_less_Suc) + have I3: "p of (q of r) = (((trunc p) of q) \\<^bsub>P\<^esub> ((ltrm p) of q)) of r" + using sub_add trunc_simps(1) assms + by (simp add: A0 I2 ltrm_closed sub_closed trunc_closed) + have I4: "p of (q of r) = (((trunc p)\\<^bsub>P\<^esub>(ltrm p)) of q) of r" + using sub_add trunc_simps(1) assms + by (simp add: trunc_simps(1) A0 I3 ltrm_closed trunc_closed) + then show ?thesis + using A0 trunc_simps(1) by auto + qed + qed + qed + qed + then show ?thesis + using assms(1) by blast +qed + +lemma sub_smult: + assumes "f \ carrier P" + assumes "q \ carrier P" + assumes "a \ carrier R" + shows "(a\\<^bsub>P\<^esub>f ) of q = a\\<^bsub>P\<^esub>(f of q)" +proof- + have "(a\\<^bsub>P\<^esub>f ) of q = ((to_poly a) \\<^bsub>P\<^esub>f) of q" + using assms by (metis P_def monom_mult_is_smult to_polynomial_def) + then have "(a\\<^bsub>P\<^esub>f ) of q = ((to_poly a) of q) \\<^bsub>P\<^esub>(f of q)" + by (simp add: assms(1) assms(2) assms(3) sub_mult to_poly_closed) + then have "(a\\<^bsub>P\<^esub>f ) of q = (to_poly a) \\<^bsub>P\<^esub>(f of q)" + by (simp add: assms(2) assms(3) sub_to_poly) + then show ?thesis + by (metis P_def assms(1) assms(2) assms(3) + monom_mult_is_smult sub_closed to_polynomial_def) +qed + +lemma to_fun_sub_monom: + assumes "is_UP_monom f" + assumes "g \ carrier P" + assumes "a \ carrier R" + shows "to_fun (f of g) a = to_fun f (to_fun g a)" +proof- + obtain b n where b_def: "b \ carrier R \ f = monom P b n" + using assms unfolding is_UP_monom_def + using P_def cfs_closed by blast + then have P0: "f of g = b \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n)" + using b_def assms(2) monom_sub by blast + have P1: "UP_pre_univ_prop R R (\x. x)" + by (simp add: UP_pre_univ_prop_fact) + then have P2: "to_fun f (to_fun g a) = b \((to_fun g a)[^]n)" + using P1 to_fun_eval[of f "to_fun g a"] P_def UP_pre_univ_prop.eval_monom assms(1) + assms(2) assms(3) b_def is_UP_monomE(1) to_fun_closed + by force + have P3: "to_fun (monom P b n of g) a = b \((to_fun g a)[^]n)" + proof- + have 0: "to_fun (monom P b n of g) a = eval R R (\x. x) a (b \\<^bsub>P\<^esub> (g[^]\<^bsub>P\<^esub>n) )" + + using UP_pre_univ_prop.eval_monom[of R "(UP R)" to_poly b g n] + P_def assms(2) b_def to_poly_UP_pre_univ_prop to_fun_eval P0 + by (metis assms(3) monom_closed sub_closed) + have 1: "to_fun (monom P b n of g) a = (eval R R (\x. x) a (to_poly b)) \ ( eval R R (\x. x) a ( g [^]\<^bsub>UP R\<^esub> n ))" + using 0 eval_ring_hom + by (metis P.nat_pow_closed P0 P_def assms(2) assms(3) b_def monom_mult_is_smult to_fun_eval to_fun_mult to_poly_closed to_polynomial_def) + have 2: "to_fun (monom P b n of g) a = b \ ( eval R R (\x. x) a ( g [^]\<^bsub>UP R\<^esub> n ))" + using 1 assms(3) b_def to_fun_eval to_fun_to_poly to_poly_closed by auto + then show ?thesis + unfolding to_function_def to_fun_def + using eval_ring_hom P_def UP_pre_univ_prop.ring_homD UP_pre_univ_prop_fact + assms(2) assms(3) ring_hom_cring.hom_pow by fastforce + qed + then show ?thesis + using b_def P2 by auto +qed + +lemma to_fun_sub: + assumes "g \ carrier P" + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "to_fun (f of g) a = (to_fun f) (to_fun g a)" +proof(rule poly_induct2[of f]) + show "f \ carrier P" + using assms by auto + show "\p. p \ carrier P \ degree p = 0 \ to_fun (p of g) a = to_fun p (to_fun g a)" + proof- + fix p + assume A0: "p \ carrier P" + assume A1: "degree p = 0" + then have P0: "degree (p of g) = 0" + by (simp add: A0 assms(1) sub_const) + then obtain b where b_def: "p of g = to_poly b \ b \ carrier R" + using A0 A1 cfs_closed assms(1) to_poly_inverse + by (meson sub_closed) + then have "to_fun (p of g) a = b" + by (simp add: assms(3) to_fun_to_poly) + have "p of g = p" + using A0 A1 P_def sub_const UP_cring_axioms assms(1) by blast + then have P1: "p = to_poly b" + using b_def by auto + have "to_fun g a \ carrier R" + using assms + by (simp add: to_fun_closed) + then show "to_fun (p of g) a = to_fun p (to_fun g a)" + using P1 \to_fun (p of g) a = b\ b_def + by (simp add: to_fun_to_poly) + qed + show "\p. 0 < degree p \ p \ carrier P \ + to_fun (trunc p of g) a = to_fun (trunc p) (to_fun g a) \ + to_fun (p of g) a = to_fun p (to_fun g a)" + proof- + fix p + assume A0: "0 < degree p" + assume A1: " p \ carrier P" + assume A2: "to_fun (trunc p of g) a = to_fun (trunc p) (to_fun g a)" + show "to_fun (p of g) a = to_fun p (to_fun g a)" + proof- + have "p of g = (trunc p) of g \\<^bsub>P\<^esub> (ltrm p) of g" + by (metis A1 assms(1) ltrm_closed sub_add trunc_simps(1) trunc_closed) + then have "to_fun (p of g) a = to_fun ((trunc p) of g) a \ (to_fun ((ltrm p) of g) a)" + by (simp add: A1 assms(1) assms(3) to_fun_plus ltrm_closed sub_closed trunc_closed) + then have 0: "to_fun (p of g) a = to_fun (trunc p) (to_fun g a) \ (to_fun ((ltrm p) of g) a)" + by (simp add: A2) + have "(to_fun ((ltrm p) of g) a) = to_fun (ltrm p) (to_fun g a)" + using to_fun_sub_monom + by (simp add: A1 assms(1) assms(3) ltrm_is_UP_monom) + then have "to_fun (p of g) a = to_fun (trunc p) (to_fun g a) \ to_fun (ltrm p) (to_fun g a)" + using 0 by auto + then show ?thesis + by (metis A1 assms(1) assms(3) to_fun_closed to_fun_plus ltrm_closed trunc_simps(1) trunc_closed) + qed + qed +qed +end + + +text\More material on constant terms and constant coefficients\ + +context UP_cring +begin + +lemma to_fun_ctrm: + assumes "f \ carrier P" + assumes "b \ carrier R" + shows "to_fun (ctrm f) b = (f 0)" + using assms + by (metis ctrm_degree ctrm_is_poly lcf_monom(2) P_def cfs_closed to_fun_to_poly to_poly_inverse) + +lemma to_fun_smult: + assumes "f \ carrier P" + assumes "b \ carrier R" + assumes "c \ carrier R" + shows "to_fun (c \\<^bsub>P\<^esub> f) b = c \(to_fun f b)" +proof- + have "(c \\<^bsub>P\<^esub> f) = (to_poly c) \\<^bsub>P\<^esub> f" + by (metis P_def assms(1) assms(3) monom_mult_is_smult to_polynomial_def) + then have "to_fun (c \\<^bsub>P\<^esub> f) b = to_fun (to_poly c) b \ to_fun f b" + by (simp add: assms(1) assms(2) assms(3) to_fun_mult to_poly_closed) + then show ?thesis + by (simp add: assms(2) assms(3) to_fun_to_poly) +qed + +lemma to_fun_monom: + assumes "c \ carrier R" + assumes "x \ carrier R" + shows "to_fun (monom P c n) x = c \ x [^] n" + by (smt P_def R.m_comm R.nat_pow_closed UP_cring.to_poly_nat_pow UP_cring_axioms assms(1) + assms(2) monom_is_UP_monom(1) sub_monom(1) to_fun_smult to_fun_sub_monom to_fun_to_poly + to_poly_closed to_poly_mult_simp(2)) + +lemma zcf_monom: + assumes "a \ carrier R" + shows "zcf (monom P a n) = to_fun (monom P a n) \" + using to_fun_monom unfolding zcf_def + by (simp add: R.nat_pow_zero assms cfs_monom) + +lemma zcf_to_fun: + assumes "p \ carrier P" + shows "zcf p = to_fun p \" + apply(rule poly_induct3[of p]) + apply (simp add: assms) + using R.zero_closed zcf_add to_fun_plus apply presburger + using zcf_monom by blast + +lemma zcf_to_poly[simp]: + assumes "a \ carrier R" + shows "zcf (to_poly a) = a" + by (metis assms cfs_closed degree_to_poly to_fun_to_poly to_poly_inverse to_poly_closed zcf_def) + +lemma zcf_ltrm_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree p > 0" + shows "zcf((ltrm p) \\<^bsub>P\<^esub> q) = \" + using zcf_to_fun[of "ltrm p \\<^bsub>P\<^esub> q" ] + by (metis ltrm_closed P.l_null P.m_closed R.zero_closed UP_zero_closed zcf_to_fun + zcf_zero assms(1) assms(2) assms(3) coeff_ltrm to_fun_mult) + +lemma zcf_mult: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "zcf(p \\<^bsub>P\<^esub> q) = (zcf p) \ (zcf q)" + using zcf_to_fun[of " p \\<^bsub>P\<^esub> q" ] zcf_to_fun[of "p" ] zcf_to_fun[of "q" ] to_fun_mult[of q p \] + by (simp add: assms(1) assms(2)) + +lemma zcf_is_ring_hom: +"zcf\ ring_hom P R" + apply(rule ring_hom_memI) + using zcf_mult zcf_add + apply (simp add: P_def UP_ring.cfs_closed UP_ring_axioms zcf_def) + apply (simp add: zcf_mult) + using zcf_add apply auto[1] + by simp + +lemma ctrm_is_ring_hom: +"ctrm \ ring_hom P P" + apply(rule ring_hom_memI) + apply (simp add: ctrm_is_poly) + apply (metis zcf_def zcf_mult cfs_closed monom_mult zero_eq_add_iff_both_eq_0) + using cfs_add[of _ _ 0] + apply (simp add: cfs_closed) + by auto + + +(**************************************************************************************************) +(**************************************************************************************************) +section\Describing the Image of (UP R) in the Ring of Functions from R to R\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma to_fun_diff: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "a \ carrier R" + shows "to_fun (p \\<^bsub>P\<^esub> q) a = to_fun p a \ to_fun q a" + using to_fun_plus[of "\\<^bsub>P\<^esub> q" p a] + by (simp add: P.minus_eq R.minus_eq assms(1) assms(2) assms(3) to_fun_minus) + +lemma to_fun_const: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_fun (monom P a 0) b = a" + by (metis lcf_monom(2) P_def UP_cring.to_fun_ctrm UP_cring_axioms assms(1) assms(2) deg_const monom_closed) + +lemma to_fun_monic_monom: + assumes "b \ carrier R" + shows "to_fun (monom P \ n) b = b[^]n" + by (simp add: assms to_fun_monom) + + +text\Constant polynomials map to constant polynomials\ + +lemma const_to_constant: + assumes "a \ carrier R" + shows "to_fun (monom P a 0) = constant_function (carrier R) a" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + unfolding ring_functions_def apply(simp add: R.ring_axioms) + apply (simp add: assms to_fun_is_Fun) + using assms ring_functions.constant_function_closed[of R a "carrier R"] + unfolding ring_functions_def apply (simp add: R.ring_axioms) + using assms to_fun_const[of a ] unfolding constant_function_def + by auto + +text\Monomial polynomials map to monomial functions\ + +lemma monom_to_monomial: + assumes "a \ carrier R" + shows "to_fun (monom P a n) = monomial_function R a n" + apply(rule ring_functions.function_ring_car_eqI[of R _ "carrier R"]) + unfolding ring_functions_def apply(simp add: R.ring_axioms) + apply (simp add: assms to_fun_is_Fun) + using assms U_function_ring.monomial_functions[of R a n] R.ring_axioms + unfolding U_function_ring_def + apply auto[1] + unfolding monomial_function_def + using assms to_fun_monom[of a _ n] + by auto +end + + +(**************************************************************************************************) +(**************************************************************************************************) +section\Taylor Expansions\ +(**************************************************************************************************) +(**************************************************************************************************) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Monic Linear Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +text\The polynomial representing the variable X\ + +definition X_poly where +"X_poly R = monom (UP R) \\<^bsub>R\<^esub> 1" + +context UP_cring +begin + +abbreviation(input) X where +"X \ X_poly R" + +lemma X_closed: +"X \ carrier P" + unfolding X_poly_def + using P_def monom_closed by blast + +lemma degree_X[simp]: + assumes "\ \\" + shows"degree X = 1" + unfolding X_poly_def + using assms P_def deg_monom[of \ 1] + by blast + +lemma X_not_zero: + assumes "\ \\" + shows"X \ \\<^bsub>P\<^esub>" + using degree_X assms by force + +lemma sub_X[simp]: + assumes "p \ carrier P" + shows "X of p = p" + unfolding X_poly_def + using P_def UP_pre_univ_prop.eval_monom1 assms compose_def to_poly_UP_pre_univ_prop + by metis + +lemma sub_monom_deg_one: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "monom P a 1 of p = a \\<^bsub>P\<^esub> p" + using assms sub_smult[of X p a] unfolding X_poly_def + by (metis P_def R.one_closed R.r_one X_closed X_poly_def monom_mult_smult sub_X) + +lemma monom_rep_X_pow: + assumes "a \ carrier R" + shows "monom P a n = a\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)" +proof- + have "monom P a n = a\\<^bsub>P\<^esub>monom P \ n" + by (metis R.one_closed R.r_one assms monom_mult_smult) + then show ?thesis + unfolding X_poly_def + using monom_pow + by (simp add: P_def) +qed + +lemma X_sub[simp]: + assumes "p \ carrier P" + shows "p of X = p" + apply(rule poly_induct3) + apply (simp add: assms) + using X_closed sub_add apply presburger + using sub_monom[of X] P_def monom_rep_X_pow X_closed by auto + +text\representation of monomials as scalar multiples of powers of X\ + +lemma ltrm_rep_X_pow: + assumes "p \ carrier P" + shows "ltrm p = (lcf p)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(degree p))" +proof- + have "ltrm p = monom P (lcf p) (degree p)" + using assms unfolding leading_term_def by (simp add: P_def) + then show ?thesis + using monom_rep_X_pow P_def assms + by (simp add: cfs_closed) +qed + +lemma to_fun_monom': + assumes "c \ carrier R" + assumes "c \\" + assumes "x \ carrier R" + shows "to_fun (c \\<^bsub>P\<^esub> X[^]\<^bsub>P\<^esub>(n::nat)) x = c \ x [^] n" + using P_def to_fun_monom monom_rep_X_pow UP_cring_axioms assms(1) assms(2) assms(3) by fastforce + +lemma to_fun_X_pow: + assumes "x \ carrier R" + shows "to_fun (X[^]\<^bsub>P\<^esub>(n::nat)) x = x [^] n" + using to_fun_monom'[of \ x n] assms + by (metis P.nat_pow_closed R.l_one R.nat_pow_closed R.one_closed R.r_null R.r_one + UP_one_closed X_closed to_fun_to_poly ring_hom_one smult_l_null smult_one to_poly_is_ring_hom) +end + +text\Monic linear polynomials\ + +definition X_poly_plus where +"X_poly_plus R a = (X_poly R) \\<^bsub>(UP R)\<^esub> to_polynomial R a" + +definition X_poly_minus where +"X_poly_minus R a = (X_poly R) \\<^bsub>(UP R)\<^esub> to_polynomial R a" + +context UP_cring +begin + +abbreviation(input) X_plus where +"X_plus \ X_poly_plus R" + +abbreviation(input) X_minus where +"X_minus \ X_poly_minus R" + +lemma X_plus_closed: + assumes "a \ carrier R" + shows "(X_plus a) \ carrier P" + unfolding X_poly_plus_def using X_closed to_poly_closed + using P_def UP_a_closed assms by auto + +lemma X_minus_closed: + assumes "a \ carrier R" + shows "(X_minus a) \ carrier P" + unfolding X_poly_minus_def using X_closed to_poly_closed + by (simp add: P_def UP_cring.UP_cring UP_cring_axioms assms cring.cring_simprules(4)) + +lemma X_minus_plus: + assumes "a \ carrier R" + shows "(X_minus a) = X_plus (\a)" + using P_def UP_ring.UP_ring UP_ring_axioms + by (simp add: X_poly_minus_def X_poly_plus_def a_minus_def assms to_poly_a_inv) + +lemma degree_of_X_plus: + assumes "a \ carrier R" + assumes "\ \\" + shows "degree (X_plus a) = 1" +proof- + have 0:"degree (X_plus a) \ 1" + using deg_add degree_X P_def unfolding X_poly_plus_def + using UP_cring.to_poly_closed UP_cring_axioms X_closed assms(1) assms(2) by fastforce + have 1:"degree (X_plus a) > 0" + by (metis One_nat_def P_def R.one_closed R.r_zero X_poly_def + X_closed X_poly_plus_def X_plus_closed assms coeff_add coeff_monom deg_aboveD + gr0I lessI n_not_Suc_n to_polynomial_def to_poly_closed) + then show ?thesis + using "0" by linarith +qed + +lemma degree_of_X_minus: + assumes "a \ carrier R" + assumes "\ \\" + shows "degree (X_minus a) = 1" + using degree_of_X_plus[of "\a"] X_minus_plus[simp] assms by auto + +lemma ltrm_of_X: +shows"ltrm X = X" + unfolding leading_term_def + by (metis P_def R.one_closed X_poly_def is_UP_monom_def is_UP_monomI leading_term_def) + +lemma ltrm_of_X_plus: + assumes "a \ carrier R" + assumes "\ \\" + shows "ltrm (X_plus a) = X" + unfolding X_poly_plus_def + using X_closed assms ltrm_of_sum_diff_degree[of X "to_poly a"] + degree_to_poly[of a] to_poly_closed[of a] degree_X ltrm_of_X + by (simp add: P_def) + +lemma ltrm_of_X_minus: + assumes "a \ carrier R" + assumes "\ \\" + shows "ltrm (X_minus a) = X" + using X_minus_plus[of a] assms + by (simp add: ltrm_of_X_plus) + +lemma lcf_of_X_minus: + assumes "a \ carrier R" + assumes "\ \\" + shows "lcf (X_minus a) = \" + using ltrm_of_X_minus unfolding X_poly_def + using P_def UP_cring.X_minus_closed UP_cring.lcf_eq UP_cring_axioms assms(1) assms(2) lcf_monom + by (metis R.one_closed) + +lemma lcf_of_X_plus: + assumes "a \ carrier R" + assumes "\ \\" + shows "lcf (X_plus a) = \" + using ltrm_of_X_plus unfolding X_poly_def + by (metis lcf_of_X_minus P_def UP_cring.lcf_eq UP_cring.X_plus_closed UP_cring_axioms X_minus_closed assms(1) assms(2) degree_of_X_minus) + +lemma to_fun_X[simp]: + assumes "a \ carrier R" + shows "to_fun X a = a" + using X_closed assms to_fun_sub_monom ltrm_is_UP_monom ltrm_of_X to_poly_closed + by (metis sub_X to_fun_to_poly) + +lemma to_fun_X_plus[simp]: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_fun (X_plus a) b = b \ a" + unfolding X_poly_plus_def + using assms to_fun_X[of b] to_fun_plus[of "to_poly a" X b] to_fun_to_poly[of a b] + using P_def X_closed to_poly_closed by auto + +lemma to_fun_X_minus[simp]: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_fun (X_minus a) b = b \ a" + using to_fun_X_plus[of "\ a" b] X_minus_plus[of a] assms + by (simp add: R.minus_eq) + +lemma cfs_X_plus: + assumes "a \ carrier R" + shows "X_plus a n = (if n = 0 then a else (if n = 1 then \ else \))" + using assms cfs_add monom_closed UP_ring_axioms cfs_monom + unfolding X_poly_plus_def to_polynomial_def X_poly_def P_def + by auto + +lemma cfs_X_minus: + assumes "a \ carrier R" + shows "X_minus a n = (if n = 0 then \ a else (if n = 1 then \ else \))" + using cfs_X_plus[of "\ a"] assms + unfolding X_poly_plus_def X_poly_minus_def + by (simp add: P_def a_minus_def to_poly_a_inv) + +text\Linear substituions\ + +lemma X_plus_sub_deg: + assumes "a \ carrier R" + assumes "f \ carrier P" + shows "degree (f of (X_plus a)) = degree f" + apply(cases "\ = \") + apply (metis P_def UP_one_closed X_plus_closed X_poly_def sub_X assms(1) assms(2) deg_one monom_one monom_zero sub_const) + using cring_sub_deg[of "X_plus a" f] assms X_plus_closed[of a] lcf_of_X_plus[of a] + ltrm_of_X_plus degree_of_X_plus[of a] P_def + by (metis lcf_eq R.nat_pow_one R.r_one UP_cring.cring_sub_deg UP_cring_axioms X_closed X_sub + cfs_closed coeff_simp deg_nzero_nzero degree_X lcoeff_nonzero2 sub_const) + +lemma X_minus_sub_deg: + assumes "a \ carrier R" + assumes "f \ carrier P" + shows "degree (f of (X_minus a)) = degree f" + using X_plus_sub_deg[of "\a"] assms X_minus_plus[of a] + by simp + +lemma plus_minus_sub: + assumes " a \ carrier R" + shows "X_plus a of X_minus a = X" + unfolding X_poly_plus_def +proof- + have "(X \\<^bsub>P\<^esub> to_poly a) of X_minus a = (X of X_minus a) \\<^bsub>P\<^esub> (to_poly a) of X_minus a" + using sub_add + by (simp add: X_closed X_minus_closed assms to_poly_closed) + then have "(X \\<^bsub>P\<^esub> to_poly a) of X_minus a = (X_minus a) \\<^bsub>P\<^esub> (to_poly a)" + by (simp add: X_minus_closed assms sub_to_poly) + then show "(X \\<^bsub>UP R\<^esub> to_poly a) of X_minus a = X" + unfolding to_polynomial_def X_poly_minus_def + by (metis P.add.inv_solve_right P.minus_eq P_def + X_closed X_poly_minus_def X_minus_closed assms monom_closed to_polynomial_def) +qed + +lemma minus_plus_sub: + assumes " a \ carrier R" + shows "X_minus a of X_plus a = X" + using plus_minus_sub[of "\a"] + unfolding X_poly_minus_def + unfolding X_poly_plus_def + using assms apply simp + by (metis P_def R.add.inv_closed R.minus_minus a_minus_def to_poly_a_inv) + +lemma ltrm_times_X: + assumes "p \ carrier P" + shows "ltrm (X \\<^bsub>P\<^esub> p) = X \\<^bsub>P\<^esub> (ltrm p)" + using assms ltrm_of_X cring_ltrm_mult[of X p] + by (metis ltrm_deg_0 P.r_null R.l_one R.one_closed UP_cring.lcf_monom(1) + UP_cring_axioms X_closed X_poly_def cfs_closed deg_zero deg_ltrm monom_zero) + +lemma times_X_not_zero: + assumes "p \ carrier P" + assumes "p \ \\<^bsub>P\<^esub>" + shows "(X \\<^bsub>P\<^esub> p) \ \\<^bsub>P\<^esub>" + by (metis (no_types, hide_lams) lcf_monom(1) lcf_of_X_minus ltrm_of_X_minus P.inv_unique + P.r_null R.l_one R.one_closed UP_zero_closed X_closed zcf_def + zcf_zero_degree_zero assms(1) assms(2) cfs_closed cfs_zero cring_lcf_mult + deg_monom deg_nzero_nzero deg_ltrm degree_X degree_of_X_minus + monom_one monom_zero) + +lemma degree_times_X: + assumes "p \ carrier P" + assumes "p \ \\<^bsub>P\<^esub>" + shows "degree (X \\<^bsub>P\<^esub> p) = degree p + 1" + using cring_deg_mult[of X p] assms times_X_not_zero[of p] + by (metis (no_types, lifting) P.r_null P.r_one P_def R.l_one R.one_closed + UP_cring.lcf_monom(1) UP_cring_axioms UP_zero_closed X_closed X_poly_def cfs_closed + deg_zero deg_ltrm degree_X monom_one monom_zero to_poly_inverse) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Basic Facts About Taylor Expansions\ +(**************************************************************************************************) +(**************************************************************************************************) +definition taylor_expansion where +"taylor_expansion R a p = compose R p (X_poly_plus R a)" + +definition(in UP_cring) taylor where +"taylor \ taylor_expansion R" + +context UP_cring +begin + +lemma taylor_expansion_ring_hom: + assumes "c \ carrier R" + shows "taylor_expansion R c \ ring_hom P P" + unfolding taylor_expansion_def + using rev_sub_is_hom[of "X_plus c"] + unfolding rev_compose_def compose_def + using X_plus_closed assms by auto + +notation taylor ("T\<^bsub>_\<^esub>") + +lemma(in UP_cring) taylor_closed: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "T\<^bsub>a\<^esub> f \ carrier P" + unfolding taylor_def + by (simp add: X_plus_closed assms(1) assms(2) sub_closed taylor_expansion_def) + +lemma taylor_deg: + assumes "a \ carrier R" + assumes "p \ carrier P" + shows "degree (T\<^bsub>a\<^esub> p) = degree p" + unfolding taylor_def taylor_expansion_def + using X_plus_sub_deg[of a p] assms + by (simp add: taylor_expansion_def) + +lemma taylor_id: + assumes "a \ carrier R" + assumes "p \ carrier P" + shows "p = (T\<^bsub>a\<^esub> p) of (X_minus a)" + unfolding taylor_expansion_def taylor_def + using assms sub_assoc[of p "X_plus a" "X_minus a"] X_plus_closed[of a] X_minus_closed[of a] + by (metis X_sub plus_minus_sub taylor_expansion_def) + +lemma taylor_eval: + assumes "a \ carrier R" + assumes "f \ carrier P" + assumes "b \ carrier R" + shows "to_fun (T\<^bsub>a\<^esub> f) b = to_fun f (b \ a)" + unfolding taylor_expansion_def taylor_def + using to_fun_sub[of "(X_plus a)" f b] to_fun_X_plus[of a b] + assms X_plus_closed[of a] by auto + +lemma taylor_eval': + assumes "a \ carrier R" + assumes "f \ carrier P" + assumes "b \ carrier R" + shows "to_fun f (b) = to_fun (T\<^bsub>a\<^esub> f) (b \ a) " + unfolding taylor_expansion_def taylor_def + using to_fun_sub[of "(X_minus a)" "T\<^bsub>a\<^esub> f" b] to_fun_X_minus[of b a] + assms X_minus_closed[of a] + by (metis taylor_closed taylor_def taylor_id taylor_expansion_def to_fun_X_minus) + +lemma(in UP_cring) degree_monom: + assumes "a \ carrier R" + shows "degree (a \\<^bsub>UP R\<^esub> (X_poly R)[^]\<^bsub>UP R\<^esub>n) = (if a = \ then 0 else n)" + apply(cases "a = \") + apply (metis (full_types) P.nat_pow_closed P_def R.one_closed UP_smult_zero X_poly_def deg_zero monom_closed) + using P_def UP_cring.monom_rep_X_pow UP_cring_axioms assms deg_monom by fastforce + +lemma(in UP_cring) poly_comp_finsum: + assumes "\i::nat. i \ n \ g i \ carrier P" + assumes "q \ carrier P" + assumes "p = (\\<^bsub>P\<^esub> i \ {..n}. g i)" + shows "p of q = (\\<^bsub>P\<^esub> i \ {..n}. (g i) of q)" +proof- + have 0: "p of q = rev_sub q p" + unfolding compose_def rev_compose_def by blast + have 1: "p of q = finsum P (rev_compose R q \ g) {..n}" + unfolding 0 unfolding assms + apply(rule ring_hom_finsum[of "rev_compose R q" P "{..n}" g ]) + using assms(2) rev_sub_is_hom apply blast + apply (simp add: UP_ring) + apply simp + by (simp add: assms(1)) + show ?thesis unfolding 1 + unfolding comp_apply rev_compose_def compose_def + by auto +qed + +lemma(in UP_cring) poly_comp_expansion: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "degree p \ n" + shows "p of q = (\\<^bsub>P\<^esub> i \ {..n}. (p i) \\<^bsub>P\<^esub> q[^]\<^bsub>P\<^esub>i)" +proof- + obtain g where g_def: "g = (\i. monom P (p i) i)" + by blast + have 0: "\i. (g i) of q = (p i) \\<^bsub>P\<^esub> q[^]\<^bsub>P\<^esub>i" + proof- fix i show "g i of q = p i \\<^bsub>P\<^esub> q [^]\<^bsub>P\<^esub> i" + using assms g_def P_def coeff_simp monom_sub + by (simp add: cfs_closed) + qed + have 1: "(\i. i \ n \ g i \ carrier P)" + using g_def assms + by (simp add: cfs_closed) + have "(\\<^bsub>P\<^esub>i\{..n}. monom P (p i) i) = p" + using assms up_repr_le[of p n] coeff_simp[of p] unfolding P_def + by auto + then have "p = (\\<^bsub>P\<^esub> i \ {..n}. g i)" + using g_def by auto + then have "p of q = (\\<^bsub>P\<^esub>i\{..n}. g i of q)" + using 0 1 poly_comp_finsum[of n g q p] + using assms(2) + by blast + then show ?thesis + by(simp add: 0) +qed + +lemma(in UP_cring) taylor_sum: + assumes "p \ carrier P" + assumes "degree p \ n" + assumes "a \ carrier R" + shows "p = (\\<^bsub>P\<^esub> i \ {..n}. T\<^bsub>a\<^esub> p i \\<^bsub>P\<^esub> (X_minus a)[^]\<^bsub>P\<^esub>i)" +proof- + have 0: "(T\<^bsub>a\<^esub> p) of X_minus a = p" + using P_def taylor_id assms(1) assms(3) + by fastforce + have 1: "degree (T\<^bsub>a\<^esub> p) \ n" + using assms + by (simp add: taylor_deg) + have 2: "T\<^bsub>a\<^esub> p of X_minus a = (\\<^bsub>P\<^esub>i\{..n}. T\<^bsub>a\<^esub> p i \\<^bsub>P\<^esub> X_minus a [^]\<^bsub>P\<^esub> i)" + using 1 X_minus_closed[of a] poly_comp_expansion[of "T\<^bsub>a\<^esub> p" "X_minus a" n] + assms taylor_closed + by blast + then show ?thesis + using 0 + by simp +qed + +text\The $i^{th}$ term in the taylor expansion\ +definition taylor_term where +"taylor_term c p i = (taylor_expansion R c p i) \\<^bsub>UP R\<^esub> (UP_cring.X_minus R c) [^]\<^bsub>UP R\<^esub>i" + +lemma (in UP_cring) taylor_term_closed: +assumes "p \ carrier P" +assumes "a \ carrier R" +shows "taylor_term a p i \ carrier (UP R)" + unfolding taylor_term_def + using P.nat_pow_closed P_def taylor_closed taylor_def X_minus_closed assms(1) assms(2) smult_closed + by (simp add: cfs_closed) + +lemma(in UP_cring) taylor_term_sum: + assumes "p \ carrier P" + assumes "degree p \ n" + assumes "a \ carrier R" + shows "p = (\\<^bsub>P\<^esub> i \ {..n}. taylor_term a p i)" + unfolding taylor_term_def taylor_def + using assms taylor_sum[of p n a] P_def + using taylor_def by auto + +lemma (in UP_cring) taylor_expansion_add: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "c \ carrier R" + shows "taylor_expansion R c (p \\<^bsub>UP R\<^esub> q) = (taylor_expansion R c p) \\<^bsub>UP R\<^esub> (taylor_expansion R c q)" + unfolding taylor_expansion_def + using assms X_plus_closed[of c] P_def sub_add + by blast + +lemma (in UP_cring) taylor_term_add: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "a \ carrier R" + shows "taylor_term a (p \\<^bsub>UP R\<^esub>q) i = taylor_term a p i \\<^bsub>UP R\<^esub> taylor_term a q i" + using assms taylor_expansion_add[of p q a] + unfolding taylor_term_def + using P.nat_pow_closed P_def taylor_closed X_minus_closed cfs_add smult_l_distr + by (simp add: taylor_def cfs_closed) + +lemma (in UP_cring) to_fun_taylor_term: +assumes "p \ carrier P" +assumes "a \ carrier R" +assumes "c \ carrier R" +shows "to_fun (taylor_term c p i) a = (T\<^bsub>c\<^esub> p i) \ (a \ c)[^]i" + using assms to_fun_smult[of "X_minus c [^]\<^bsub>UP R\<^esub> i" a "taylor_expansion R c p i"] + to_fun_X_minus[of c a] to_fun_nat_pow[of "X_minus c" a i] + unfolding taylor_term_def + using P.nat_pow_closed P_def taylor_closed taylor_def X_minus_closed + by (simp add: cfs_closed) + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\Defining the (Scalar-Valued) Derivative of a Polynomial Using the Taylor Expansion\ +(**************************************************************************************************) +(**************************************************************************************************) +definition derivative where +"derivative R f a = (taylor_expansion R a f) 1" + +context UP_cring +begin + +abbreviation(in UP_cring) deriv where +"deriv \ derivative R" + +lemma(in UP_cring) deriv_closed: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "(derivative R f a) \ carrier R" + unfolding derivative_def + using taylor_closed taylor_def assms(1) assms(2) cfs_closed by auto + +lemma(in UP_cring) deriv_add: + assumes "f \ carrier P" + assumes "g \ carrier P" + assumes "a \ carrier R" + shows "deriv (f \\<^bsub>P\<^esub> g) a = deriv f a \ deriv g a" + unfolding derivative_def taylor_expansion_def using assms + by (simp add: X_plus_closed sub_add sub_closed) + +end +(**************************************************************************************************) +(**************************************************************************************************) +section\The Polynomial-Valued Derivative Operator\ +(**************************************************************************************************) +(**************************************************************************************************) + +context UP_cring +begin + + (**********************************************************************) + (**********************************************************************) + subsection\Operator Which Shifts Coefficients\ + (**********************************************************************) + (**********************************************************************) +lemma cfs_times_X: + assumes "g \ carrier P" + shows "(X \\<^bsub>P\<^esub> g) (Suc n) = g n" + apply(rule poly_induct3[of g]) + apply (simp add: assms) + apply (metis (no_types, lifting) P.m_closed P.r_distr X_closed cfs_add) + by (metis (no_types, lifting) P_def R.l_one R.one_closed R.r_null Suc_eq_plus1 X_poly_def + cfs_monom coeff_monom_mult coeff_simp monom_closed monom_mult) + +lemma times_X_pow_coeff: + assumes "g \ carrier P" + shows "(monom P \ k \\<^bsub>P\<^esub> g) (n + k) = g n" + using coeff_monom_mult P.m_closed P_def assms coeff_simp monom_closed + by (simp add: cfs_closed) + +lemma zcf_eq_zero_unique: + assumes "f \ carrier P" + assumes "g \ carrier P \ (f = X \\<^bsub>P\<^esub> g)" + shows "\ h. h \ carrier P \ (f = X \\<^bsub>P\<^esub> h) \ h = g" +proof- + fix h + assume A: "h \ carrier P \ (f = X \\<^bsub>P\<^esub> h)" + then have 0: " X \\<^bsub>P\<^esub> g = X \\<^bsub>P\<^esub> h" + using assms(2) by auto + show "h = g" + using 0 A assms + by (metis P_def coeff_simp cfs_times_X up_eqI) +qed + +lemma f_minus_ctrm: + assumes "f \ carrier P" + shows "zcf(f \\<^bsub>P\<^esub> ctrm f) = \" + using assms + by (smt ctrm_is_poly P.add.inv_closed P.minus_closed P_def R.r_neg R.zero_closed zcf_to_fun + to_fun_minus to_fun_plus UP_cring_axioms zcf_ctrm zcf_def a_minus_def cfs_closed) + +definition poly_shift where +"poly_shift f n = f (Suc n)" + +lemma poly_shift_closed: + assumes "f \ carrier P" + shows "poly_shift f \ carrier P" + apply(rule UP_car_memI[of "deg R f"]) + unfolding poly_shift_def +proof - + fix n :: nat + assume "deg R f < n" + then have "deg R f < Suc n" + using Suc_lessD by blast + then have "f (Suc n) = \\<^bsub>P\<^esub> (Suc n)" + by (metis P.l_zero UP_zero_closed assms coeff_of_sum_diff_degree0) + then show "f (Suc n) = \" + by simp +next + show " \n. f (Suc n) \ carrier R" + by(rule cfs_closed, rule assms) +qed + +lemma poly_shift_eq_0: + assumes "f \ carrier P" + shows "f n = (ctrm f \\<^bsub>P\<^esub> X \\<^bsub>P\<^esub> poly_shift f) n" + apply(cases "n = 0") + apply (smt ctrm_degree ctrm_is_poly ltrm_of_X One_nat_def P.r_null P.r_zero P_def UP_cring.lcf_monom(1) UP_cring_axioms UP_mult_closed UP_r_one UP_zero_closed X_closed zcf_ltrm_mult zcf_def zcf_zero assms cfs_add cfs_closed deg_zero degree_X lessI monom_one poly_shift_closed to_poly_inverse) +proof- assume A: "n \ 0" + then obtain k where k_def: " n = Suc k" + by (meson lessI less_Suc_eq_0_disj) + show ?thesis + using cfs_times_X[of "poly_shift f" k] poly_shift_def[of f k] poly_shift_closed assms + cfs_add[of "ctrm f" "X \\<^bsub>P\<^esub> poly_shift f" n] unfolding k_def + by (simp add: X_closed cfs_closed cfs_monom) +qed + +lemma poly_shift_eq: + assumes "f \ carrier P" + shows "f = (ctrm f \\<^bsub>P\<^esub> X \\<^bsub>P\<^esub> poly_shift f)" +by(rule ext, rule poly_shift_eq_0, rule assms) + +lemma poly_shift_id: + assumes "f \ carrier P" + shows "f \\<^bsub>P\<^esub> ctrm f = X \\<^bsub>P\<^esub> poly_shift f" + using assms poly_shift_eq[of f] poly_shift_closed unfolding a_minus_def + by (metis ctrm_is_poly P.add.inv_solve_left P.m_closed UP_a_comm UP_a_inv_closed X_closed) + +lemma poly_shift_degree_zero: + assumes "p \ carrier P" + assumes "degree p = 0" + shows "poly_shift p = \\<^bsub>P\<^esub>" + by (metis ltrm_deg_0 P.r_neg P.r_null UP_ring UP_zero_closed X_closed zcf_eq_zero_unique + abelian_group.minus_eq assms(1) assms(2) poly_shift_closed poly_shift_id ring_def) + +lemma poly_shift_degree: + assumes "p \ carrier P" + assumes "degree p > 0" + shows "degree (poly_shift p) = degree p - 1 " + using poly_shift_id[of p] + by (metis ctrm_degree ctrm_is_poly P.r_null X_closed add_diff_cancel_right' assms(1) assms(2) + deg_zero degree_of_difference_diff_degree degree_times_X nat_less_le poly_shift_closed) + +lemma poly_shift_monom: + assumes "a \ carrier R" + shows "poly_shift (monom P a (Suc k)) = (monom P a k)" +proof- + have "(monom P a (Suc k)) = ctrm (monom P a (Suc k)) \\<^bsub>P\<^esub> X \\<^bsub>P\<^esub>poly_shift (monom P a (Suc k))" + using poly_shift_eq[of "monom P a (Suc k)"] assms monom_closed + by blast + then have "(monom P a (Suc k)) = \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> X \\<^bsub>P\<^esub>poly_shift (monom P a (Suc k))" + using assms by simp + then have "(monom P a (Suc k)) = X \\<^bsub>P\<^esub>poly_shift (monom P a (Suc k))" + using X_closed assms poly_shift_closed by auto + then have "X \\<^bsub>P\<^esub>(monom P a k) = X \\<^bsub>P\<^esub>poly_shift (monom P a (Suc k))" + by (metis P_def R.l_one R.one_closed X_poly_def assms monom_mult plus_1_eq_Suc) + then show ?thesis + using X_closed X_not_zero assms + by (meson UP_mult_closed zcf_eq_zero_unique monom_closed poly_shift_closed) +qed + +lemma(in UP_cring) poly_shift_add: + assumes "f \ carrier P" + assumes "g \ carrier P" + shows "poly_shift (f \\<^bsub>P\<^esub> g) = (poly_shift f) \\<^bsub>P\<^esub> (poly_shift g)" + apply(rule ext) + using cfs_add[of "poly_shift f" "poly_shift g"] poly_shift_closed poly_shift_def + by (simp add: poly_shift_def assms(1) assms(2)) + +lemma(in UP_cring) poly_shift_s_mult: + assumes "f \ carrier P" + assumes "s \ carrier R" + shows "poly_shift (s \\<^bsub>P\<^esub>f) = s \\<^bsub>P\<^esub> (poly_shift f)" +proof- + have "(s \\<^bsub>P\<^esub>f) = (ctrm (s \\<^bsub>P\<^esub>f)) \\<^bsub>P\<^esub>(X \\<^bsub>P\<^esub> poly_shift (s \\<^bsub>P\<^esub>f))" + using poly_shift_eq[of "(s \\<^bsub>P\<^esub>f)"] assms(1) assms(2) + by blast + then have 0: "(s \\<^bsub>P\<^esub>f) = (s \\<^bsub>P\<^esub>(ctrm f)) \\<^bsub>P\<^esub>(X \\<^bsub>P\<^esub> poly_shift (s \\<^bsub>P\<^esub>f))" + using ctrm_smult assms(1) assms(2) by auto + have 1: "(s \\<^bsub>P\<^esub>f) = s \\<^bsub>P\<^esub> ((ctrm f) \\<^bsub>P\<^esub> (X \\<^bsub>P\<^esub> (poly_shift f)))" + using assms(1) poly_shift_eq by auto + have 2: "(s \\<^bsub>P\<^esub>f) = (s \\<^bsub>P\<^esub>(ctrm f)) \\<^bsub>P\<^esub> (s \\<^bsub>P\<^esub>(X \\<^bsub>P\<^esub> (poly_shift f)))" + by (simp add: "1" X_closed assms(1) assms(2) ctrm_is_poly poly_shift_closed smult_r_distr) + have 3: "(s \\<^bsub>P\<^esub>f) = (s \\<^bsub>P\<^esub>(ctrm f)) \\<^bsub>P\<^esub> (X \\<^bsub>P\<^esub> (s \\<^bsub>P\<^esub>(poly_shift f)))" + using "2" UP_m_comm X_closed assms(1) assms(2) smult_assoc2 + by (simp add: poly_shift_closed) + have 4: "(X \\<^bsub>P\<^esub> poly_shift (s \\<^bsub>P\<^esub>f)) = (X \\<^bsub>P\<^esub> (s \\<^bsub>P\<^esub>(poly_shift f)))" + using 3 0 X_closed assms(1) assms(2) ctrm_is_poly poly_shift_closed + by auto + then show ?thesis + using X_closed X_not_zero assms(1) assms(2) + by (metis UP_mult_closed UP_smult_closed zcf_eq_zero_unique poly_shift_closed) +qed + +lemma zcf_poly_shift: + assumes "f \ carrier P" + shows "zcf (poly_shift f) = f 1" + apply(rule poly_induct3) + apply (simp add: assms) + using poly_shift_add zcf_add cfs_add poly_shift_closed apply metis + unfolding zcf_def using poly_shift_monom poly_shift_degree_zero + by (simp add: poly_shift_def) + +fun poly_shift_iter ("shift") where +Base:"poly_shift_iter 0 f = f"| +Step:"poly_shift_iter (Suc n) f = poly_shift (poly_shift_iter n f)" + +lemma shift_closed: + assumes "f \ carrier P" + shows "shift n f \ carrier P" + apply(induction n) + using assms poly_shift_closed by auto + + (**********************************************************************) + (**********************************************************************) + subsection\Operator Which Multiplies Coefficients by Their Degree\ + (**********************************************************************) + (**********************************************************************) + +definition n_mult where +"n_mult f = (\n. [n]\\<^bsub>R\<^esub>(f n))" + +lemma(in UP_cring) n_mult_closed: + assumes "f \ carrier P" + shows "n_mult f \ carrier P" + apply(rule UP_car_memI[of "deg R f"]) + unfolding n_mult_def + apply (metis P.l_zero R.add.nat_pow_one UP_zero_closed assms cfs_zero coeff_of_sum_diff_degree0) + using assms cfs_closed by auto + +text\Facts about the shift function\ + +lemma shift_one: +"shift (Suc 0) = poly_shift" + by auto + +lemma shift_factor0: + assumes "f \ carrier P" + shows "degree f \ (Suc k) \ degree (f \\<^bsub>P\<^esub> ((shift (Suc k) f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc k)))) < (Suc k)" +proof(induction k) + case 0 + have 0: " f \\<^bsub>P\<^esub> (ctrm f) = (shift (Suc 0) f)\\<^bsub>P\<^esub>X" + by (metis UP_m_comm X_closed assms poly_shift_id shift_closed shift_one) + then have " f \\<^bsub>P\<^esub>(shift (Suc 0) f)\\<^bsub>P\<^esub>X = (ctrm f) " + proof- + have " f \\<^bsub>P\<^esub> (ctrm f) \\<^bsub>P\<^esub> (shift (Suc 0) f)\\<^bsub>P\<^esub>X= (shift (Suc 0) f)\\<^bsub>P\<^esub>X \\<^bsub>P\<^esub> (shift (Suc 0) f)\\<^bsub>P\<^esub>X" + using 0 by simp + then have " f \\<^bsub>P\<^esub> (ctrm f) \\<^bsub>P\<^esub> (shift (Suc 0) f)\\<^bsub>P\<^esub>X = \\<^bsub>P\<^esub>" + using UP_cring.UP_cring[of R] assms + by (metis "0" P.ring_simprules(4) P_def UP_ring.UP_ring UP_ring_axioms + a_minus_def abelian_group.r_neg ctrm_is_poly ring_def) + then have " f \\<^bsub>P\<^esub> ((ctrm f) \\<^bsub>P\<^esub> (shift (Suc 0) f)\\<^bsub>P\<^esub>X) = \\<^bsub>P\<^esub>" + using assms P.ring_simprules + by (metis "0" poly_shift_id poly_shift_eq) + then have " f \\<^bsub>P\<^esub> ((shift (Suc 0) f)\\<^bsub>P\<^esub>X \\<^bsub>P\<^esub> (ctrm f) ) = \\<^bsub>P\<^esub>" + using P.m_closed UP_a_comm X_closed assms ctrm_is_poly shift_closed + by presburger + then have "f \\<^bsub>P\<^esub> ((shift (Suc 0) f)\\<^bsub>P\<^esub>X) \\<^bsub>P\<^esub> (ctrm f)= \\<^bsub>P\<^esub>" + using P.add.m_assoc P.ring_simprules(14) P.ring_simprules(19) assms "0" + P.add.inv_closed P.r_neg P.r_zero ctrm_is_poly + by smt + then show ?thesis + by (metis "0" P.add.m_comm P.m_closed P.ring_simprules(14) P.ring_simprules(18) + P.ring_simprules(3) X_closed assms ctrm_is_poly poly_shift_id poly_shift_eq + shift_closed) + qed + then have " f \\<^bsub>P\<^esub>(shift (Suc 0) f)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc 0)) = (ctrm f) " + proof- + have "X = X[^]\<^bsub>P\<^esub>(Suc 0)" + by (simp add: X_closed) + then show ?thesis + using 0 \f \\<^bsub>P\<^esub> shift (Suc 0) f \\<^bsub>P\<^esub> X = ctrm f\ + by auto + qed + then have " degree (f \\<^bsub>P\<^esub>(shift (Suc 0) f)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc 0))) < 1" + using ctrm_degree[of f] assms by simp + then show ?case + by blast +next + case (Suc n) + fix k + assume IH: "degree f \ (Suc k) \ degree (f \\<^bsub>P\<^esub> ((shift (Suc k) f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc k)))) < (Suc k)" + show "degree f \ (Suc (Suc k)) \ degree (f \\<^bsub>P\<^esub> ((shift (Suc (Suc k)) f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc (Suc k))))) < (Suc (Suc k))" + proof- + obtain n where n_def: "n = Suc k" + by simp + have IH': "degree f \ n \ degree (f \\<^bsub>P\<^esub> ((shift n f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))) < n" + using n_def IH by auto + have P: "degree f \ (Suc n) \ degree (f \\<^bsub>P\<^esub> ((shift (Suc n) f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc n)))) < (Suc n)" + proof- + obtain g where g_def: "g = (f \\<^bsub>P\<^esub> ((shift n f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)))" + by simp + obtain s where s_def: "s = shift n f" + by simp + obtain s' where s'_def: "s' = shift (Suc n) f" + by simp + have P: "g \ carrier P" "s \ carrier P" "s' \ carrier P" "(X[^]\<^bsub>P\<^esub>n) \ carrier P" + using s_def s'_def g_def assms shift_closed[of f n] + apply (simp add: X_closed) + apply (simp add: \f \ carrier P \ shift n f \ carrier P\ assms s_def) + using P_def UP_cring.shift_closed UP_cring_axioms assms s'_def apply blast + using X_closed by blast + have g_def': "g = (f \\<^bsub>P\<^esub> (s \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)))" + using g_def s_def by auto + assume "degree f \ (Suc n)" + then have " degree (f \\<^bsub>P\<^esub> (s \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))) < n" + using IH' Suc_leD s_def by blast + then have d_g: "degree g < n" using g_def' by auto + have P0: "f \\<^bsub>P\<^esub> (s' \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc n))) = ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g" + proof- + have "s = (ctrm s) \\<^bsub>P\<^esub> (X \\<^bsub>P\<^esub> s')" + using s_def s'_def P_def poly_shift_eq UP_cring_axioms assms shift_closed + by (simp add: UP_cring.poly_shift_eq) + then have 0: "g = f \\<^bsub>P\<^esub> ((ctrm s) \\<^bsub>P\<^esub> (X \\<^bsub>P\<^esub> s')) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)" + using g_def' by auto + then have "g = f \\<^bsub>P\<^esub> ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> ((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))" + using P cring X_closed P.l_distr P.ring_simprules(19) UP_a_assoc a_minus_def assms + by (simp add: a_minus_def ctrm_is_poly) + then have "g \\<^bsub>P\<^esub> ((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) = f \\<^bsub>P\<^esub> ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))" + using P cring X_closed P.l_distr P.ring_simprules UP_a_assoc a_minus_def assms + by (simp add: P.r_neg2 ctrm_is_poly) + then have " ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) = f \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)))" + using P cring X_closed P.ring_simprules UP_a_assoc a_minus_def assms + by (simp add: P.ring_simprules(17) ctrm_is_poly) + then have " ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) = f \\<^bsub>P\<^esub> (((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g)" + by (simp add: P(1) P(3) UP_a_comm X_closed) + then have "((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) = f \\<^bsub>P\<^esub> ((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g" + using P(1) P(3) P.ring_simprules(19) UP_a_assoc a_minus_def assms + by (simp add: a_minus_def X_closed) + then have "((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g= f \\<^bsub>P\<^esub> ((X \\<^bsub>P\<^esub> s') \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))" + by (metis P(1) P(3) P(4) P.add.inv_solve_right P.m_closed P.ring_simprules(14) + P.ring_simprules(4) P_def UP_cring.X_closed UP_cring_axioms assms) + then have "((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g= f \\<^bsub>P\<^esub> ((s' \\<^bsub>P\<^esub> X) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n))" + by (simp add: P(3) UP_m_comm X_closed) + then have "((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g= f \\<^bsub>P\<^esub> (s' \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc n)))" + using P(3) P.nat_pow_Suc2 UP_m_assoc X_closed by auto + then show ?thesis + by auto + qed + have P1: "degree (((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \\<^bsub>P\<^esub> g) \ n" + proof- + have Q0: "degree ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \ n" + proof(cases "ctrm s = \\<^bsub>P\<^esub>") + case True + then show ?thesis + by (simp add: P(4)) + next + case False + then have F0: "degree ((ctrm s)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)) \ degree (ctrm s) + degree (X[^]\<^bsub>P\<^esub>n) " + by (meson ctrm_is_poly P(2) P(4) deg_mult_ring) + have F1: "\\\\ degree (X[^]\<^bsub>P\<^esub>n) = n" + unfolding X_poly_def + using P_def cring_monom_degree by auto + show ?thesis + by (metis (no_types, hide_lams) F0 F1 ltrm_deg_0 P(2) P.r_null P_def R.l_null R.l_one + R.nat_pow_closed R.zero_closed X_poly_def assms cfs_closed + add_0 deg_const deg_zero deg_ltrm + monom_pow monom_zero zero_le) + qed + then show ?thesis + using d_g + by (simp add: P(1) P(2) P(4) bound_deg_sum ctrm_is_poly) + qed + then show ?thesis + using s'_def P0 by auto + qed + assume "degree f \ (Suc (Suc k)) " + then show "degree (f \\<^bsub>P\<^esub> ((shift (Suc (Suc k)) f) \\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc (Suc k))))) < (Suc (Suc k))" + using P by(simp add: n_def) + qed +qed + +lemma(in UP_cring) shift_degree0: + assumes "f \ carrier P" + shows "degree f >n \ Suc (degree (shift (Suc n) f)) = degree (shift n f)" +proof(induction n) + case 0 + assume B: "0< degree f" + have 0: "degree (shift 0 f) = degree f" + by simp + have 1: "degree f = degree (f \\<^bsub>P\<^esub> (ctrm f))" + using assms(1) B ctrm_degree degree_of_difference_diff_degree + by (simp add: ctrm_is_poly) + have "(f \\<^bsub>P\<^esub> (ctrm f)) = X \\<^bsub>P\<^esub>(shift 1 f)" + using P_def poly_shift_id UP_cring_axioms assms(1) by auto + then have "degree (f \\<^bsub>P\<^esub> (ctrm f)) = 1 + (degree (shift 1 f))" + by (metis "1" B P.r_null X_closed add.commute assms deg_nzero_nzero degree_times_X not_gr_zero shift_closed) + then have "degree (shift 0 f) = 1 + (degree (shift 1 f))" + using 0 1 by auto + then show ?case + by simp +next + case (Suc n) + fix n + assume IH: "(n < degree f \ Suc (degree (shift (Suc n) f)) = degree (shift n f))" + show "Suc n < degree f \ Suc (degree (shift (Suc (Suc n)) f)) = degree (shift (Suc n) f)" + proof- + assume A: " Suc n < degree f" + then have 0: "(shift (Suc n) f) = ctrm ((shift (Suc n) f)) \\<^bsub>P\<^esub> (shift (Suc (Suc n)) f)\\<^bsub>P\<^esub>X" + by (metis UP_m_comm X_closed assms local.Step poly_shift_eq shift_closed) + have N: "(shift (Suc (Suc n)) f) \ \\<^bsub>P\<^esub>" + proof + assume C: "shift (Suc (Suc n)) f = \\<^bsub>P\<^esub>" + obtain g where g_def: "g = f \\<^bsub>P\<^esub> (shift (Suc (Suc n)) f)\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>(Suc (Suc n)))" + by simp + have C0: "degree g < degree f" + using g_def assms A + by (meson Suc_leI Suc_less_SucD Suc_mono less_trans_Suc shift_factor0) + have C1: "g = f" + using C + by (simp add: P.minus_eq X_closed assms g_def) + then show False + using C0 by auto + qed + have 1: "degree (shift (Suc n) f) = degree ((shift (Suc n) f) \\<^bsub>P\<^esub> ctrm ((shift (Suc n) f)))" + proof(cases "degree (shift (Suc n) f) = 0") + case True + then show ?thesis + using N assms poly_shift_degree_zero poly_shift_closed shift_closed by auto + next + case False + then have "degree (shift (Suc n) f) > degree (ctrm ((shift (Suc n) f)))" + proof - + have "shift (Suc n) f \ carrier P" + using assms shift_closed by blast + then show ?thesis + using False ctrm_degree by auto + qed + then show ?thesis + proof - + show ?thesis + using \degree (ctrm (shift (Suc n) f)) < degree (shift (Suc n) f)\ + assms ctrm_is_poly degree_of_difference_diff_degree shift_closed by presburger + qed + qed + have 2: "(shift (Suc n) f) \\<^bsub>P\<^esub> ctrm ((shift (Suc n) f)) = (shift (Suc (Suc n)) f)\\<^bsub>P\<^esub>X" + using 0 + by (metis Cring_Poly.INTEG.Step P.m_comm X_closed assms poly_shift_id shift_closed) + have 3: "degree ((shift (Suc n) f) \\<^bsub>P\<^esub> ctrm ((shift (Suc n) f))) = degree (shift (Suc (Suc n)) f) + 1" + using 2 N X_closed X_not_zero assms degree_X shift_closed + by (metis UP_m_comm degree_times_X) + then show ?thesis using 1 + by linarith + qed +qed + +lemma(in UP_cring) shift_degree: + assumes "f \ carrier P" + shows "degree f \ n \ degree (shift n f) + n = degree f" +proof(induction n) + case 0 + then show ?case + by auto +next + case (Suc n) + fix n + assume IH: "(n \ degree f \ degree (shift n f) + n = degree f)" + show "Suc n \ degree f \ degree (shift (Suc n) f) + Suc n = degree f" + proof- + assume A: "Suc n \ degree f " + have 0: "degree (shift n f) + n = degree f" + using IH A by auto + have 1: "degree (shift n f) = Suc (degree (shift (Suc n) f))" + using A assms shift_degree0 by auto + show "degree (shift (Suc n) f) + Suc n = degree f" + using 0 1 by simp + qed +qed + +lemma(in UP_cring) shift_degree': + assumes "f \ carrier P" + shows "degree (shift (degree f) f) = 0" + using shift_degree assms + by fastforce + +lemma(in UP_cring) shift_above_degree: + assumes "f \ carrier P" + assumes "k > degree f" + shows "(shift k f) = \\<^bsub>P\<^esub>" +proof- + have "\n. shift ((degree f)+ (Suc n)) f = \\<^bsub>P\<^esub>" + proof- + fix n + show "shift ((degree f)+ (Suc n)) f = \\<^bsub>P\<^esub>" + proof(induction n) + case 0 + have B0:"shift (degree f) f = ctrm(shift (degree f) f) \\<^bsub>P\<^esub> (shift (degree f + Suc 0) f)\\<^bsub>P\<^esub>X" + proof - + have f1: "\f n. f \ carrier P \ shift n f \ carrier P" + by (meson shift_closed) + then have "shift (degree f + Suc 0) f \ carrier P" + using assms(1) by blast + then show ?thesis + using f1 by (simp add: P.m_comm X_closed assms(1) poly_shift_eq) + qed + have B1:"shift (degree f) f = ctrm(shift (degree f) f)" + proof - + have "shift (degree f) f \ carrier P" + using assms(1) shift_closed by blast + then show ?thesis + using ltrm_deg_0 assms(1) shift_degree' by auto + qed + have B2: "(shift (degree f + Suc 0) f)\\<^bsub>P\<^esub>X = \\<^bsub>P\<^esub>" + using B0 B1 X_closed assms(1) + proof - + have "\f n. f \ carrier P \ shift n f \ carrier P" + using shift_closed by blast + then show ?thesis + by (metis (no_types) B0 B1 P.add.l_cancel_one UP_mult_closed X_closed assms(1)) + qed + then show ?case + by (metis P.r_null UP_m_comm UP_zero_closed X_closed assms(1) zcf_eq_zero_unique shift_closed) + next + case (Suc n) + fix n + assume "shift (degree f + Suc n) f = \\<^bsub>P\<^esub>" + then show "shift (degree f + Suc (Suc n)) f = \\<^bsub>P\<^esub>" + by (simp add: poly_shift_degree_zero) + qed + qed + then show ?thesis + using assms(2) less_iff_Suc_add by auto +qed + +lemma(in UP_domain) shift_cfs0: + assumes "f \ carrier P" + shows "zcf(shift 1 f) = f 1" + using assms + by (simp add: zcf_poly_shift) + +lemma(in UP_cring) X_mult_cf: + assumes "p \ carrier P" + shows "(p \\<^bsub>P\<^esub> X) (k+1) = p k" + unfolding X_poly_def + using assms + by (metis UP_m_comm X_closed X_poly_def add.commute plus_1_eq_Suc cfs_times_X) + +lemma(in UP_cring) X_pow_cf: + assumes "p \ carrier P" + shows "(p \\<^bsub>P\<^esub> X[^]\<^bsub>P\<^esub>(n::nat)) (n + k) = p k" +proof- + have P: "\f. f \ carrier P \ (f \\<^bsub>P\<^esub> X[^]\<^bsub>P\<^esub>(n::nat)) (n + k) = f k" + proof(induction n) + show "\f. f \ carrier P \ (f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> (0::nat)) (0 + k) = f k" + proof- + fix f + assume B0: "f \ carrier P" + show "(f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> (0::nat)) (0 + k) = f k" + by (simp add: B0) + qed + fix n + fix f + assume IH: "(\f. f \ carrier P \ (f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> n) (n + k) = f k)" + assume A0: " f \ carrier P" + show "(f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> Suc n) (Suc n + k) = f k" + proof- + have 0: "(f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> n)(n + k) = f k" + using A0 IH by simp + have 1: "((f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> n)\\<^bsub>P\<^esub>X) (Suc n + k) = (f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> n)(n + k)" + using X_mult_cf A0 P.m_closed P.nat_pow_closed + Suc_eq_plus1 X_closed add_Suc by presburger + have 2: "(f \\<^bsub>P\<^esub> (X [^]\<^bsub>P\<^esub> n \\<^bsub>P\<^esub>X)) (Suc n + k) = (f \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> n)(n + k)" + using 1 + by (simp add: A0 UP_m_assoc X_closed) + then show ?thesis + by (simp add: "0") + qed + qed + show ?thesis using assms P[of p] by auto +qed + +lemma poly_shift_cfs: + assumes "f \ carrier P" + shows "poly_shift f n = f (Suc n)" +proof- + have "(f \\<^bsub>P\<^esub> ctrm f) (Suc n) = (X \\<^bsub>P\<^esub> (poly_shift f)) (Suc n)" + using assms poly_shift_id by auto + then show ?thesis unfolding X_poly_def using poly_shift_closed assms + by (metis (no_types, lifting) ctrm_degree ctrm_is_poly + P.add.m_comm P.minus_closed coeff_of_sum_diff_degree0 poly_shift_id poly_shift_eq cfs_times_X zero_less_Suc) +qed + +lemma(in UP_cring) shift_cfs: + assumes "p \ carrier P" + shows "(shift k p) n = p (k + n)" + apply(induction k arbitrary: n) + by (auto simp: assms poly_shift_cfs shift_closed) + + (**********************************************************************) + (**********************************************************************) + subsection\The Derivative Operator\ + (**********************************************************************) + (**********************************************************************) +definition pderiv where +"pderiv p = poly_shift (n_mult p)" + +lemma pderiv_closed: + assumes "p \ carrier P" + shows "pderiv p \ carrier P" + unfolding pderiv_def + using assms n_mult_closed[of p] poly_shift_closed[of "n_mult p"] + by blast + +text\Function which obtains the first n+1 terms of f, in ascending order of degree:\ + +definition trms_of_deg_leq where +"trms_of_deg_leq n f \ f \\<^bsub>(UP R)\<^esub> ((shift (Suc n) f) \\<^bsub>UP R\<^esub> monom P \ (Suc n))" + +lemma trms_of_deg_leq_closed: + assumes "f \ carrier P" + shows "trms_of_deg_leq n f \ carrier P" + unfolding trms_of_deg_leq_def using assms + by (metis P.m_closed P.minus_closed P_def R.one_closed monom_closed shift_closed) + +lemma trms_of_deg_leq_id: + assumes "f \ carrier P" + shows "f \\<^bsub>P\<^esub> (trms_of_deg_leq k f) = shift (Suc k) f \\<^bsub>P\<^esub> monom P \ (Suc k)" + unfolding trms_of_deg_leq_def + using assms + by (smt P.add.inv_closed P.l_zero P.m_closed P.minus_add P.minus_minus P.r_neg + P_def R.one_closed UP_a_assoc a_minus_def monom_closed shift_closed) + +lemma trms_of_deg_leq_id': + assumes "f \ carrier P" + shows "f = (trms_of_deg_leq k f) \\<^bsub>P\<^esub> shift (Suc k) f \\<^bsub>P\<^esub> monom P \ (Suc k)" + using trms_of_deg_leq_id assms trms_of_deg_leq_closed[of f] + by (smt P.add.inv_closed P.l_zero P.m_closed P.minus_add P.minus_minus P.r_neg R.one_closed UP_a_assoc a_minus_def monom_closed shift_closed) + +lemma deg_leqI: + assumes "p \ carrier P" + assumes "\n. n > k \ p n = \" + shows "degree p \ k" + by (metis assms(1) assms(2) deg_zero deg_ltrm le0 le_less_linear monom_zero) + +lemma deg_leE: + assumes "p \ carrier P" + assumes "degree p < k" + shows "p k = \" + using assms coeff_of_sum_diff_degree0 P_def coeff_simp deg_aboveD + by auto + +lemma trms_of_deg_leq_deg: + assumes "f \ carrier P" + shows "degree (trms_of_deg_leq k f) \ k" +proof- + have "\n. (trms_of_deg_leq k f) (Suc k + n) = \" + proof- + fix n + have 0: "(shift (Suc k) f \\<^bsub>UP R\<^esub> monom P \ (Suc k)) (Suc k + n) = shift (Suc k) f n" + using assms shift_closed cfs_monom_mult_l + by (metis P.m_comm P_def R.one_closed add.commute monom_closed times_X_pow_coeff) + then show "trms_of_deg_leq k f (Suc k + n) = \" + unfolding trms_of_deg_leq_def + using shift_cfs[of f "Suc k" n] + cfs_minus[of f "shift (Suc k) f \\<^bsub>UP R\<^esub> monom P \ (Suc k)" "Suc k + n"] + by (metis P.m_closed P.r_neg P_def R.one_closed a_minus_def assms + cfs_minus cfs_zero monom_closed shift_closed) + qed + then show ?thesis using deg_leqI + by (metis (no_types, lifting) assms le_iff_add less_Suc_eq_0_disj less_Suc_eq_le trms_of_deg_leq_closed) +qed + +lemma trms_of_deg_leq_zero_is_ctrm: + assumes "f \ carrier P" + assumes "degree f > 0" + shows "trms_of_deg_leq 0 f = ctrm f" +proof- + have "f = ctrm f \\<^bsub>P\<^esub> (X \\<^bsub>P\<^esub> (shift (Suc 0) f))" + using assms poly_shift_eq + by simp + then have "f = ctrm f \\<^bsub>P\<^esub> (X [^]\<^bsub>UP R\<^esub> Suc 0 \\<^bsub>P\<^esub> (shift (Suc 0) f))" + using P.nat_pow_eone P_def X_closed by auto + then show ?thesis + unfolding trms_of_deg_leq_def + by (metis (no_types, lifting) ctrm_is_poly One_nat_def P.add.right_cancel P.m_closed + P.minus_closed P.nat_pow_eone P_def UP_m_comm X_closed X_poly_def assms(1) shift_closed + trms_of_deg_leq_def trms_of_deg_leq_id') +qed + +lemma cfs_monom_mult: + assumes "p \ carrier P" + assumes "a \ carrier R" + assumes "k < n" + shows "(p \\<^bsub>P\<^esub> (monom P a n)) k = \" + apply(rule poly_induct3[of p]) + apply (simp add: assms(1)) + apply (metis (no_types, lifting) P.l_distr P.m_closed R.r_zero R.zero_closed assms(2) cfs_add monom_closed) + using assms monom_mult[of _ a _ n] + by (metis R.m_closed R.m_comm add.commute cfs_monom not_add_less1) + +lemma(in UP_cring) cfs_monom_mult_2: + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "m < n" + shows "((monom P a n) \\<^bsub>P\<^esub> f) m = \" + using cfs_monom_mult + by (simp add: P.m_comm assms(1) assms(2) assms(3)) + +lemma trms_of_deg_leq_cfs: + assumes "f \ carrier P" + shows "trms_of_deg_leq n f k = (if k \ n then (f k) else \)" + unfolding trms_of_deg_leq_def + apply(cases "k \ n") + using cfs_minus[of f "shift (Suc n) f \\<^bsub>UP R\<^esub> monom P \ (Suc n)"] + cfs_monom_mult[of _ \ k "Suc n"] + apply (metis (no_types, lifting) P.m_closed P.minus_closed P_def R.one_closed R.r_zero assms + cfs_add cfs_closed le_refl monom_closed nat_less_le nat_neq_iff not_less_eq_eq shift_closed + trms_of_deg_leq_def trms_of_deg_leq_id') + using trms_of_deg_leq_deg[of f n] deg_leE + unfolding trms_of_deg_leq_def + using assms trms_of_deg_leq_closed trms_of_deg_leq_def by auto + +lemma trms_of_deg_leq_iter: + assumes "f \ carrier P" + shows "trms_of_deg_leq (Suc k) f = (trms_of_deg_leq k f) \\<^bsub>P\<^esub> monom P (f (Suc k)) (Suc k)" +proof fix x + show "trms_of_deg_leq (Suc k) f x = (trms_of_deg_leq k f \\<^bsub>P\<^esub> monom P (f (Suc k)) (Suc k)) x" + apply(cases "x \ k") + using trms_of_deg_leq_cfs trms_of_deg_leq_closed cfs_closed[of f "Suc k"] + cfs_add[of "trms_of_deg_leq k f" "monom P (f (Suc k)) (Suc k)" x] + apply (simp add: assms) + using deg_leE assms cfs_closed cfs_monom apply auto[1] + by (simp add: assms cfs_closed cfs_monom trms_of_deg_leq_cfs trms_of_deg_leq_closed) +qed + +lemma trms_of_deg_leq_0: + assumes "f \ carrier P" + shows "trms_of_deg_leq 0 f = ctrm f" + by (metis One_nat_def P.r_null P_def UP_m_comm UP_zero_closed X_closed X_poly_def assms not_gr_zero + poly_shift_degree_zero shift_one trms_of_deg_leq_def trms_of_deg_leq_zero_is_ctrm trunc_simps(2) trunc_zero) + +lemma trms_of_deg_leq_degree_f: + assumes "f \ carrier P" + shows "trms_of_deg_leq (degree f) f = f" +proof fix x + show "trms_of_deg_leq (deg R f) f x = f x" + using assms trms_of_deg_leq_cfs deg_leE[of f x] + by simp +qed + +definition(in UP_cring) lin_part where +"lin_part f = trms_of_deg_leq 1 f" + +lemma(in UP_cring) lin_part_id: + assumes "f \ carrier P" + shows "lin_part f = (ctrm f) \\<^bsub>P\<^esub> monom P (f 1) 1" + unfolding lin_part_def + by (simp add: assms trms_of_deg_leq_0 trms_of_deg_leq_iter) + +lemma(in UP_cring) lin_part_eq: + assumes "f \ carrier P" + shows "f = lin_part f \\<^bsub>P\<^esub> (shift 2 f) \\<^bsub>P\<^esub> monom P \ 2" + unfolding lin_part_def + by (metis Suc_1 assms trms_of_deg_leq_id') + +text\Constant term of a substitution:\ + +lemma zcf_eval: + assumes "f \ carrier P" + shows "zcf f = to_fun f \" + using assms zcf_to_fun by blast + +lemma ctrm_of_sub: + assumes "f \ carrier P" + assumes "g \ carrier P" + shows "zcf(f of g) = to_fun f (zcf g)" + apply(rule poly_induct3[of f]) + apply (simp add: assms(1)) + using P_def UP_cring.to_fun_closed UP_cring_axioms zcf_add zcf_to_fun assms(2) to_fun_plus sub_add sub_closed apply fastforce + using R.zero_closed zcf_to_fun assms(2) to_fun_sub monom_closed sub_closed by presburger + +text\Evaluation of linear part:\ + +lemma to_fun_lin_part: + assumes "f \ carrier P" + assumes "b \ carrier R" + shows "to_fun (lin_part f) b = (f 0) \ (f 1) \ b" + using assms lin_part_id[of f] to_fun_ctrm to_fun_monom monom_closed + by (simp add: cfs_closed to_fun_plus) + +text\Constant term of taylor expansion:\ + +lemma taylor_zcf: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "zcf(T\<^bsub>a\<^esub> f) = to_fun f a" + unfolding taylor_expansion_def + using ctrm_of_sub assms P_def zcf_eval X_plus_closed taylor_closed taylor_eval by auto + +lemma(in UP_cring) taylor_eq_1: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "(T\<^bsub>a\<^esub> f) \\<^bsub>P\<^esub> (trms_of_deg_leq 1 (T\<^bsub>a\<^esub> f)) = (shift (2::nat) (T\<^bsub>a\<^esub> f))\\<^bsub>P\<^esub> (X[^]\<^bsub>P\<^esub>(2::nat))" + by (metis P.nat_pow_eone P.nat_pow_mult P_def Suc_1 taylor_closed X_closed X_poly_def assms(1) + assms(2) monom_one_Suc2 one_add_one trms_of_deg_leq_id) + +lemma(in UP_cring) taylor_deg_1: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "f of (X_plus a) = (lin_part (T\<^bsub>a\<^esub> f)) \\<^bsub>P\<^esub> (shift (2::nat) (T\<^bsub>a\<^esub> f))\\<^bsub>P\<^esub> (X[^]\<^bsub>P\<^esub>(2::nat))" + using taylor_eq_1[of f a] + unfolding taylor_expansion_def lin_part_def + using One_nat_def X_plus_closed assms(1) + assms(2) trms_of_deg_leq_id' numeral_2_eq_2 sub_closed + by (metis P.nat_pow_Suc2 P.nat_pow_eone P_def taylor_def X_closed X_poly_def monom_one_Suc taylor_expansion_def) + +lemma(in UP_cring) taylor_deg_1_eval: + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c = to_fun (shift (2::nat) (T\<^bsub>a\<^esub> f)) b" + assumes "fa = to_fun f a" + assumes "f'a = deriv f a" + shows "to_fun f (b \ a) = fa \ (f'a \ b) \ (c \ b[^](2::nat))" + using assms taylor_deg_1 unfolding derivative_def +proof- + have 0: "to_fun f (b \ a) = to_fun (f of (X_plus a)) b" + using to_fun_sub assms X_plus_closed by auto + have 1: "to_fun (lin_part (T\<^bsub>a\<^esub> f)) b = fa \ (f'a \ b) " + using assms to_fun_lin_part[of "(T\<^bsub>a\<^esub> f)" b] + by (metis P_def taylor_def UP_cring.taylor_zcf UP_cring.taylor_closed UP_cring_axioms zcf_def derivative_def) + have 2: "(T\<^bsub>a\<^esub> f) = (lin_part (T\<^bsub>a\<^esub> f)) \\<^bsub>P\<^esub> ((shift 2 (T\<^bsub>a\<^esub> f))\\<^bsub>P\<^esub>X[^]\<^bsub>P\<^esub>(2::nat))" + using lin_part_eq[of "(T\<^bsub>a\<^esub>f)"] assms(1) assms(2) taylor_closed + by (metis taylor_def taylor_deg_1 taylor_expansion_def) + then have "to_fun (T\<^bsub>a\<^esub>f) b = fa \ (f'a \ b) \ to_fun ((shift 2 (T\<^bsub>a\<^esub> f))\\<^bsub>P\<^esub>X[^]\<^bsub>P\<^esub>(2::nat)) b" + using 1 2 + by (metis P.nat_pow_closed taylor_closed UP_mult_closed X_closed assms(1) assms(2) assms(3) + to_fun_plus lin_part_def shift_closed trms_of_deg_leq_closed) + then have "to_fun (T\<^bsub>a\<^esub>f) b = fa \ (f'a \ b) \ c \ to_fun (X[^]\<^bsub>P\<^esub>(2::nat)) b" + by (simp add: taylor_closed X_closed assms(1) assms(2) assms(3) assms(4) to_fun_mult shift_closed) + then have 3: "to_fun f (b \ a)= fa \ (f'a \ b) \ c \ to_fun (X[^]\<^bsub>P\<^esub>(2::nat)) b" + using taylor_eval assms(1) assms(2) assms(3) by auto + have "to_fun (X[^]\<^bsub>P\<^esub>(2::nat)) b = b[^](2::nat)" + by (metis P.nat_pow_Suc2 P.nat_pow_eone R.nat_pow_Suc2 + R.nat_pow_eone Suc_1 to_fun_X + X_closed assms(3) to_fun_mult) + then show ?thesis + using 3 by auto +qed + +lemma(in UP_cring) taylor_deg_1_eval': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c = to_fun (shift (2::nat) (T\<^bsub>a\<^esub> f)) b" + assumes "fa = to_fun f a" + assumes "f'a = deriv f a" + shows "to_fun f (a \ b) = fa \ (f'a \ b) \ (c \ b[^](2::nat))" + using R.add.m_comm taylor_deg_1_eval assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) + by auto + +lemma(in UP_cring) taylor_deg_1_eval'': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c = to_fun (shift (2::nat) (T\<^bsub>a\<^esub> f)) (\b)" + shows "to_fun f (a \ b) = (to_fun f a) \ (deriv f a \ b) \ (c \ b[^](2::nat))" +proof- + have "\b \ carrier R" + using assms + by blast + then have 0: "to_fun f (a \ b) = (to_fun f a)\ (deriv f a \ (\b)) \ (c \ (\b)[^](2::nat))" + unfolding a_minus_def + using taylor_deg_1_eval'[of f a "\b" c "(to_fun f a)" "deriv f a"] assms + by auto + have 1: "\ (deriv f a \ b) = (deriv f a \ (\b))" + using assms + by (simp add: R.r_minus deriv_closed) + have 2: "(c \ b[^](2::nat)) = (c \ (\b)[^](2::nat))" + using assms + by (metis R.add.inv_closed R.add.inv_solve_right R.l_zero R.nat_pow_Suc2 + R.nat_pow_eone R.zero_closed Suc_1 UP_ring_axioms UP_ring_def + ring.ring_simprules(26) ring.ring_simprules(27)) + show ?thesis + using 0 1 2 + unfolding a_minus_def + by simp +qed + +lemma(in UP_cring) taylor_deg_1_expansion: + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c = to_fun (shift (2::nat) (T\<^bsub>a\<^esub> f)) (b \ a)" + assumes "fa = to_fun f a" + assumes "f'a = deriv f a" + shows "to_fun f (b) = fa \ f'a \ (b \ a) \ (c \ (b \ a)[^](2::nat))" +proof- + obtain b' where b'_def: "b'= b \ a " + by simp + then have b'_def': "b = b' \ a" + using assms + by (metis R.add.inv_solve_right R.minus_closed R.minus_eq) + have "to_fun f (b' \ a) = fa \ (f'a \ b') \ (c \ b'[^](2::nat))" + using assms taylor_deg_1_eval[of f a b' c fa f'a] b'_def + by blast + then have "to_fun f (b) = fa \ (f'a \ b') \ (c \ b'[^](2::nat))" + using b'_def' + by auto + then show "to_fun f (b) = fa \ f'a \ (b \ a) \ c \ (b \ a) [^] (2::nat)" + using b'_def + by auto +qed + +lemma(in UP_cring) Taylor_deg_1_expansion': + assumes "f \ carrier (UP R)" + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "\c \ carrier R. to_fun f (b) = (to_fun f a) \ (deriv f a) \ (b \ a) \ (c \ (b \ a)[^](2::nat))" + using taylor_deg_1_expansion[of f a b] assms unfolding P_def + by (metis P_def R.minus_closed taylor_closed shift_closed to_fun_closed) + + +text\Basic Properties of deriv and pderiv:\ + +lemma n_mult_degree_bound: + assumes "f \ carrier P" + shows "degree (n_mult f) \ degree f" + apply(rule deg_leqI) + apply (simp add: assms n_mult_closed) + by (simp add: assms deg_leE n_mult_def) + +lemma pderiv_deg_0[simp]: + assumes "f \ carrier P" + assumes "degree f = 0" + shows "pderiv f = \\<^bsub>P\<^esub>" +proof- + have "degree (n_mult f) = 0" + using P_def n_mult_degree_bound assms(1) assms(2) by fastforce + then show ?thesis + unfolding pderiv_def + by (simp add: assms(1) n_mult_closed poly_shift_degree_zero) +qed + +lemma deriv_deg_0: + assumes "f \ carrier P" + assumes "degree f = 0" + assumes "a \ carrier R" + shows "deriv f a = \" + unfolding derivative_def taylor_expansion_def + using X_plus_closed assms(1) assms(2) assms(3) deg_leE sub_const by force + +lemma poly_shift_monom': + assumes "a \ carrier R" + shows "poly_shift (a \\<^bsub>P\<^esub> (X[^]\<^bsub>P\<^esub>(Suc n))) = a\\<^bsub>P\<^esub>(X[^]\<^bsub>P\<^esub>n)" + using assms monom_rep_X_pow poly_shift_monom by auto + +lemma monom_coeff: + assumes "a \ carrier R" + shows "(a \\<^bsub>P\<^esub> X [^]\<^bsub>P\<^esub> (n::nat)) k = (if (k = n) then a else \)" + using assms cfs_monom monom_rep_X_pow by auto + +lemma cfs_n_mult: + assumes "p \ carrier P" + shows "n_mult p n = [n]\(p n)" + by (simp add: n_mult_def) + +lemma cfs_add_nat_pow: + assumes "p \ carrier P" + shows "([(n::nat)]\\<^bsub>P\<^esub>p) k = [n]\(p k)" + apply(induction n) by (auto simp: assms) + +lemma cfs_add_int_pow: + assumes "p \ carrier P" + shows "([(n::int)]\\<^bsub>P\<^esub>p) k = [n]\(p k)" + apply(induction n) + by(auto simp: add_pow_int_ge assms cfs_add_nat_pow add_pow_int_lt) + +lemma add_nat_pow_monom: + assumes "a \ carrier R" + shows "[(n::nat)]\\<^bsub>P\<^esub>monom P a k = monom P ([n]\a) k" + apply(rule ext) + by (simp add: assms cfs_add_nat_pow cfs_monom) + +lemma add_int_pow_monom: + assumes "a \ carrier R" + shows "[(n::int)]\\<^bsub>P\<^esub>monom P a k = monom P ([n]\a) k" + apply(rule ext) + by (simp add: assms cfs_add_int_pow cfs_monom) + +lemma n_mult_monom: + assumes "a \ carrier R" + shows "n_mult (monom P a (Suc n)) = monom P ([Suc n]\a) (Suc n)" + apply(rule ext) + unfolding n_mult_def + using assms cfs_monom by auto + +lemma pderiv_monom: + assumes "a \ carrier R" + shows "pderiv (monom P a n) = monom P ([n]\a) (n-1)" + apply(cases "n = 0") + apply (simp add: assms) + unfolding pderiv_def + using assms Suc_diff_1[of n] n_mult_monom[of a "n-1"] poly_shift_monom[of "[Suc (n-1)]\a" "Suc (n-1)"] + by (metis R.add.nat_pow_closed neq0_conv poly_shift_monom) + +lemma pderiv_monom': + assumes "a \ carrier R" + shows "pderiv (a \\<^bsub>P\<^esub> X[^]\<^bsub>P\<^esub>(n::nat)) = ([n]\a)\\<^bsub>P\<^esub> X[^]\<^bsub>P\<^esub>(n-1)" + using assms pderiv_monom[of a n ] + by (simp add: P_def UP_cring.monom_rep_X_pow UP_cring_axioms) + +lemma n_mult_add: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "n_mult (p \\<^bsub>P\<^esub> q) = n_mult p \\<^bsub>P\<^esub> n_mult q" +proof(rule ext) fix x show "n_mult (p \\<^bsub>P\<^esub> q) x = (n_mult p \\<^bsub>P\<^esub> n_mult q) x" + using assms R.add.nat_pow_distrib[of "p x" "q x" x] cfs_add[of p q x] + cfs_add[of "n_mult p" "n_mult q" x] n_mult_closed + unfolding n_mult_def + by (simp add: cfs_closed) +qed + +lemma pderiv_add: + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "pderiv (p \\<^bsub>P\<^esub> q) = pderiv p \\<^bsub>P\<^esub> pderiv q" + unfolding pderiv_def + using assms poly_shift_add n_mult_add + by (simp add: n_mult_closed) + +lemma zcf_monom_sub: + assumes "p \ carrier P" + shows "zcf ((monom P \ (Suc n)) of p) = zcf p [^] (Suc n)" + apply(induction n) + using One_nat_def P.nat_pow_eone R.nat_pow_eone R.one_closed R.zero_closed zcf_to_fun assms to_fun_closed monom_sub smult_one apply presburger + using P_def UP_cring.ctrm_of_sub UP_cring_axioms zcf_to_fun assms to_fun_closed to_fun_monom monom_closed + by fastforce + +lemma zcf_monom_sub': + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "zcf ((monom P a (Suc n)) of p) = a \ zcf p [^] (Suc n)" + using zcf_monom_sub assms P_def R.zero_closed UP_cring.ctrm_of_sub UP_cring.to_fun_monom UP_cring_axioms + zcf_to_fun to_fun_closed monom_closed by fastforce + +lemma deriv_monom: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "deriv (monom P a n) b = ([n]\a)\(b[^](n-1))" +proof(induction n) + case 0 + have 0: "b [^] ((0::nat) - 1) \ carrier R" + using assms + by simp + then show ?case unfolding derivative_def using assms + by (smt One_nat_def P_def R.add.nat_pow_0 R.nat_pow_Suc2 R.nat_pow_eone R.zero_closed + taylor_def taylor_deg UP_cring.taylor_closed UP_cring.zcf_monom UP_cring.shift_one + UP_cring_axioms zcf_degree_zero zcf_zero_degree_zero degree_monom monom_closed + monom_rep_X_pow plus_1_eq_Suc poly_shift_degree_zero shift_cfs to_fun_monom to_fun_zero zero_diff) +next + case (Suc n) + show ?case + proof(cases "n = 0") + case True + have T0: "[Suc n] \ a \ b [^] (Suc n - 1) = a" + by (simp add: True assms(1)) + have T1: "(X_poly R \\<^bsub>UP R\<^esub> to_polynomial R b) [^]\<^bsub>UP R\<^esub> Suc n = X_poly R \\<^bsub>UP R\<^esub> to_polynomial R b " + using P.nat_pow_eone P_def True UP_a_closed X_closed assms(2) to_poly_closed by auto + then show ?thesis + unfolding derivative_def taylor_expansion_def + using T0 T1 True sub_monom(2)[of "X_plus b" a "Suc n"] cfs_add assms + unfolding P_def X_poly_plus_def to_polynomial_def X_poly_def + by (smt Group.nat_pow_0 lcf_eq lcf_monom(2) ltrm_of_X_plus One_nat_def P_def R.one_closed + R.r_one R.r_zero UP_cring.zcf_monom UP_cring.degree_of_X_plus + UP_cring.poly_shift_degree_zero UP_cring_axioms X_closed X_plus_closed X_poly_def + X_poly_plus_def zcf_zero_degree_zero cfs_monom_mult_l degree_to_poly to_fun_X_pow + plus_1_eq_Suc poly_shift_cfs poly_shift_monom to_poly_closed to_poly_mult_simp(2) + to_poly_nat_pow to_polynomial_def) + next + case False + have "deriv (monom P a (Suc n)) b = ((monom P a (Suc n)) of (X_plus b)) 1" + unfolding derivative_def taylor_expansion_def + by auto + then have "deriv (monom P a (Suc n)) b = (((monom P a n) of (X_plus b)) \\<^bsub>P\<^esub> (X_plus b)) 1" + using monom_mult[of a \ n 1] sub_mult[of "X_plus b" "monom P a n" "monom P \ 1" ] X_plus_closed[of b] assms + by (metis lcf_monom(1) P.l_one P.nat_pow_eone P_def R.one_closed R.r_one Suc_eq_plus1 + deg_one monom_closed monom_one sub_monom(1) to_poly_inverse) + then have "deriv (monom P a (Suc n)) b = (((monom P a n) of (X_plus b)) \\<^bsub>P\<^esub> (monom P \ 1) \\<^bsub>P\<^esub> + (((monom P a n) of (X_plus b)) \\<^bsub>P\<^esub> to_poly b)) 1" + unfolding X_poly_plus_def + by (metis P.r_distr P_def X_closed X_plus_closed X_poly_def X_poly_plus_def assms(1) assms(2) monom_closed sub_closed to_poly_closed) + then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0 \ b \ ((monom P a n) of (X_plus b)) 1" + unfolding X_poly_plus_def + by (smt One_nat_def P.m_closed P_def UP_m_comm X_closed X_plus_closed X_poly_def X_poly_plus_def + assms(1) assms(2) cfs_add cfs_monom_mult_l monom_closed plus_1_eq_Suc sub_closed cfs_times_X to_polynomial_def) + then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0 \ b \ (deriv (monom P a n) b)" + by (simp add: derivative_def taylor_expansion_def) + then have "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0 \ b \ ( ([n]\a)\(b[^](n-1)))" + by (simp add: Suc) + then have 0: "deriv (monom P a (Suc n)) b = ((monom P a n) of (X_plus b)) 0 \ ([n]\a)\(b[^]n)" + using assms R.m_comm[of b] R.nat_pow_mult[of b "n-1" 1] False + by (metis (no_types, lifting) R.add.nat_pow_closed R.m_lcomm R.nat_pow_closed R.nat_pow_eone add.commute add_eq_if plus_1_eq_Suc) + have 1: "((monom P a n) of (X_plus b)) 0 = a \ b[^]n" + unfolding X_poly_plus_def using zcf_monom_sub' + by (smt ctrm_of_sub One_nat_def P_def R.l_zero R.one_closed UP_cring.zcf_to_poly + UP_cring.f_minus_ctrm UP_cring_axioms X_plus_closed X_poly_def X_poly_plus_def zcf_add + zcf_def assms(1) assms(2) to_fun_monom monom_closed monom_one_Suc2 poly_shift_id poly_shift_monom to_poly_closed) + show ?thesis + using 0 1 R.add.nat_pow_Suc2 R.add.nat_pow_closed R.l_distr R.nat_pow_closed assms(1) assms(2) diff_Suc_1 by presburger + qed +qed + +lemma deriv_smult: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "g \ carrier P" + shows "deriv (a \\<^bsub>P\<^esub> g) b = a \ (deriv g b)" + unfolding derivative_def taylor_expansion_def + using assms sub_smult X_plus_closed cfs_smult + by (simp add: sub_closed) + +lemma deriv_const: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "deriv (monom P a 0) b = \" + unfolding derivative_def + using assms taylor_closed taylor_def taylor_deg deg_leE by auto + +lemma deriv_monom_deg_one: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "deriv (monom P a 1) b = a" + unfolding derivative_def taylor_expansion_def + using assms cfs_X_plus[of b 1] sub_monom_deg_one X_plus_closed[of b] + by simp + +lemma monom_Suc: + assumes "a \ carrier R" + shows "monom P a (Suc n) = monom P \ 1 \\<^bsub>P\<^esub> monom P a n" + "monom P a (Suc n) = monom P a n \\<^bsub>P\<^esub> monom P \ 1" + apply (metis R.l_one R.one_closed Suc_eq_plus1_left assms monom_mult) + by (metis R.one_closed R.r_one Suc_eq_plus1 assms monom_mult) + + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\The Product Rule\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma(in UP_cring) times_x_product_rule: + assumes "f \ carrier P" + shows "pderiv (f \\<^bsub>P\<^esub> up_ring.monom P \ 1) = f \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1" +proof(rule poly_induct3[of f]) + show "f \ carrier P" + using assms by blast + show "\p q. q \ carrier P \ + p \ carrier P \ + pderiv (p \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> pderiv p \\<^bsub>P\<^esub> up_ring.monom P \ 1 \ + pderiv (q \\<^bsub>P\<^esub> up_ring.monom P \ 1) = q \\<^bsub>P\<^esub> pderiv q \\<^bsub>P\<^esub> up_ring.monom P \ 1 \ + pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> pderiv (p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1" + proof- fix p q assume A: "q \ carrier P" + "p \ carrier P" + "pderiv (p \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> pderiv p \\<^bsub>P\<^esub> up_ring.monom P \ 1" + "pderiv (q \\<^bsub>P\<^esub> up_ring.monom P \ 1) = q \\<^bsub>P\<^esub> pderiv q \\<^bsub>P\<^esub> up_ring.monom P \ 1" + have 0: "(p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1 = (p \\<^bsub>P\<^esub> up_ring.monom P \ 1) \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + using A assms by (meson R.one_closed UP_l_distr is_UP_monomE(1) is_UP_monomI) + have 1: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = pderiv (p \\<^bsub>P\<^esub> up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (q \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + unfolding 0 apply(rule pderiv_add) + using A is_UP_monomE(1) monom_is_UP_monom(1) apply blast + using A is_UP_monomE(1) monom_is_UP_monom(1) by blast + have 2: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> pderiv p \\<^bsub>P\<^esub> up_ring.monom P \ 1 \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> pderiv q \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + unfolding 1 A by blast + have 3: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> (pderiv p \\<^bsub>P\<^esub> up_ring.monom P \ 1 \\<^bsub>P\<^esub> pderiv q \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + unfolding 2 + using A P.add.m_lcomm R.one_closed UP_a_assoc UP_a_closed UP_mult_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed by presburger + have 4: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> ((pderiv p \\<^bsub>P\<^esub> pderiv q) \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + unfolding 3 using A P.l_distr R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed by presburger + show 5: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1) = p \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> pderiv (p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> up_ring.monom P \ 1" + unfolding 4 using pderiv_add A by presburger + qed + show "\a n. a \ carrier R \ + pderiv (up_ring.monom P a n \\<^bsub>P\<^esub> up_ring.monom P \ 1) = up_ring.monom P a n \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> up_ring.monom P \ 1" + proof- fix a n assume A: "a \ carrier R" + have 0: "up_ring.monom P a n \\<^bsub>P\<^esub> up_ring.monom P \ 1 = up_ring.monom P a (Suc n)" + using A monom_Suc(2) by presburger + have 1: "pderiv (up_ring.monom P a n \\<^bsub>P\<^esub> up_ring.monom P \ 1) = [(Suc n)] \\<^bsub>P\<^esub> (up_ring.monom P a n)" + unfolding 0 using A add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom + by (simp add: P_def) + have 2: "pderiv (up_ring.monom P a n \\<^bsub>P\<^esub> up_ring.monom P \ 1) = (up_ring.monom P a n) \\<^bsub>P\<^esub> [n] \\<^bsub>P\<^esub> (up_ring.monom P a n)" + unfolding 1 using A P.add.nat_pow_Suc2 is_UP_monomE(1) monom_is_UP_monom(1) by blast + have 3: "pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> up_ring.monom P \ 1 = [n] \\<^bsub>P\<^esub> (up_ring.monom P a n)" + apply(cases "n = 0") + using A add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom pderiv_deg_0 apply auto[1] + using monom_Suc(2)[of a "n-1"] A add_nat_pow_monom n_mult_monom pderiv_def poly_shift_monom + by (metis R.add.nat_pow_closed Suc_eq_plus1 add_eq_if monom_Suc(2) pderiv_monom) + show "pderiv (up_ring.monom P a n \\<^bsub>P\<^esub> up_ring.monom P \ 1) = up_ring.monom P a n \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> up_ring.monom P \ 1" + unfolding 2 3 by blast + qed +qed + +lemma(in UP_cring) deg_one_eval: + assumes "g \ carrier (UP R)" + assumes "deg R g = 1" + shows "\t. t \ carrier R \ to_fun g t = g 0 \ (g 1)\t" +proof- + obtain h where h_def: "h = ltrm g" + by blast + have 0: "deg R (g \\<^bsub>UP R\<^esub> h) = 0" + using assms unfolding h_def + by (metis ltrm_closed ltrm_eq_imp_deg_drop ltrm_monom P_def UP_car_memE(1) less_one) + have 1: "g \\<^bsub>UP R\<^esub> h = to_poly (g 0)" + proof(rule ext) fix x show "(g \\<^bsub>UP R\<^esub> h) x = to_polynomial R (g 0) x" + proof(cases "x = 0") + case True + have T0: "h 0 = \" + unfolding h_def using assms UP_car_memE(1) cfs_monom by presburger + have T1: "(g \\<^bsub>UP R\<^esub> h) 0 = g 0 \ h 0" + using ltrm_closed P_def assms(1) cfs_minus h_def by blast + then show ?thesis using T0 assms + by (smt "0" ltrm_closed ltrm_deg_0 P.minus_closed P_def UP_car_memE(1) UP_zero_closed zcf_def zcf_zero deg_zero degree_to_poly h_def to_poly_closed to_poly_inverse to_poly_minus trunc_simps(2) trunc_zero) + next + case False + then have "x > 0" + by presburger + then show ?thesis + by (metis "0" ltrm_closed P.minus_closed P_def UP_car_memE(1) UP_cring.degree_to_poly UP_cring_axioms assms(1) deg_leE h_def to_poly_closed) + qed + qed + have 2: "g = (g \\<^bsub>UP R\<^esub> h) \\<^bsub>UP R\<^esub> h" + unfolding h_def using assms + by (metis "1" P_def h_def lin_part_def lin_part_id to_polynomial_def trms_of_deg_leq_degree_f) + fix t assume A: "t \ carrier R" + have 3: " to_fun g t = to_fun (g \\<^bsub>UP R\<^esub> h) t \ to_fun h t" + using 2 + by (metis "1" A P_def UP_car_memE(1) assms(1) h_def monom_closed to_fun_plus to_polynomial_def) + then show "to_fun g t = g 0 \ g 1 \ t " + unfolding 1 h_def + using A P_def UP_cring.lin_part_def UP_cring_axioms assms(1) assms(2) to_fun_lin_part trms_of_deg_leq_degree_f by fastforce +qed + +lemma nmult_smult: + assumes "a \ carrier R" + assumes "f \ carrier P" + shows "n_mult (a \\<^bsub>P\<^esub> f) = a \\<^bsub>P\<^esub> (n_mult f)" + apply(rule poly_induct4[of f]) + apply (simp add: assms(2)) + using assms(1) n_mult_add n_mult_closed smult_closed smult_r_distr apply presburger + using assms apply(intro ext, metis (no_types, lifting) ctrm_smult ltrm_deg_0 P_def R.add.nat_pow_0 UP_cring.ctrm_degree UP_cring.n_mult_closed UP_cring.n_mult_def UP_cring_axioms UP_smult_closed UP_zero_closed zcf_degree_zero zcf_zero deg_const deg_zero le_0_eq monom_closed n_mult_degree_bound smult_r_null) + using monom_mult_smult n_mult_monom assms + by (smt lcf_monom(1) P_def R.add.nat_pow_closed R.add_pow_rdistr R.zero_closed UP_cring.to_poly_mult_simp(1) UP_cring_axioms UP_smult_closed cfs_closed cring_lcf_mult monom_closed to_polynomial_def) + +lemma pderiv_smult: + assumes "a \ carrier R" + assumes "f \ carrier P" + shows "pderiv (a \\<^bsub>P\<^esub> f) = a \\<^bsub>P\<^esub> (pderiv f)" + unfolding pderiv_def + using assms + by (simp add: n_mult_closed nmult_smult poly_shift_s_mult) + +lemma(in UP_cring) pderiv_minus: + assumes "a \ carrier P" + assumes "b \ carrier P" + shows "pderiv (a \\<^bsub>P\<^esub> b) = pderiv a \\<^bsub>P\<^esub> pderiv b" +proof- + have "\\<^bsub>P\<^esub> b = (\\)\\<^bsub>P\<^esub>b" + using R.one_closed UP_smult_one assms(2) smult_l_minus by presburger + thus ?thesis unfolding a_minus_def using pderiv_add assms pderiv_smult + by (metis P.add.inv_closed R.add.inv_closed R.one_closed UP_smult_one pderiv_closed smult_l_minus) +qed + +lemma(in UP_cring) pderiv_const: + assumes "b \ carrier R" + shows "pderiv (up_ring.monom P b 0) = \\<^bsub>P\<^esub>" + using assms pderiv_monom[of b 0] deg_const is_UP_monomE(1) monom_is_UP_monom(1) pderiv_deg_0 + by blast + +lemma(in UP_cring) pderiv_minus_const: + assumes "a \ carrier P" + assumes "b \ carrier R" + shows "pderiv (a \\<^bsub>P\<^esub> up_ring.monom P b 0) = pderiv a" + using pderiv_minus[of a "up_ring.monom P b 0" ] assms pderiv_const[of b] + by (smt P.l_zero P.minus_closed P_def UP_cring.pderiv_const UP_cring.pderiv_minus UP_cring.poly_shift_eq UP_cring_axioms cfs_closed monom_closed pderiv_add pderiv_closed poly_shift_id) + +lemma(in UP_cring) monom_product_rule: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a n) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a n" +proof- + have "\f. f \ carrier P \ pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a n) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a n" + proof(induction n) + case 0 + show ?case + proof fix f show "f \ carrier P \ pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a 0) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a 0) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a 0 " + proof assume A: "f \ carrier P" + have 0: "f \\<^bsub>P\<^esub> up_ring.monom P a 0 = a \\<^bsub>P\<^esub>f" + using assms A UP_m_comm is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult by presburger + have 1: "f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a 0) = \\<^bsub>P\<^esub>" + using A assms P.r_null pderiv_const by presburger + have 2: "pderiv f \\<^bsub>P\<^esub> up_ring.monom P a 0 = a \\<^bsub>P\<^esub> pderiv f" + using assms A UP_m_comm is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult pderiv_closed by presburger + show "pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a 0) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a 0) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a 0" + unfolding 0 1 2 using A UP_l_zero UP_smult_closed assms(2) pderiv_closed pderiv_smult by presburger + qed + qed + next + case (Suc n) + show "\f. f \ carrier P \ + pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a (Suc n)) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)" + proof fix f + show "f \ carrier P \ + pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a (Suc n)) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)" + proof + assume A: "f \ carrier P" + show " pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a (Suc n)) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)" + proof(cases "n = 0") + case True + have 0: "(f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = a \\<^bsub>P\<^esub> f \\<^bsub>P\<^esub> up_ring.monom P \ 1" + proof - + have "\n. up_ring.monom P a n \ carrier P" + using assms(2) is_UP_monomE(1) monom_is_UP_monom(1) by presburger + then show ?thesis + by (metis A P.m_assoc P.m_comm R.one_closed True assms(2) is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_mult_is_smult) + qed + have 1: "f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a (Suc n)) = a \\<^bsub>P\<^esub> f" + using assms True + by (metis A One_nat_def P.m_comm R.add.nat_pow_eone diff_Suc_1 is_UP_monomE(1) is_UP_monomI monom_mult_is_smult pderiv_monom) + have 2: "pderiv f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n) = a \\<^bsub>P\<^esub> (pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + using A assms unfolding True + by (metis P.m_lcomm R.one_closed UP_mult_closed is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_mult_is_smult pderiv_closed) + have 3: "a \\<^bsub>P\<^esub> f \\<^bsub>P\<^esub> a \\<^bsub>P\<^esub> (pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1) = a \\<^bsub>P\<^esub> (f \\<^bsub>P\<^esub>(pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1))" + using assms A P.m_closed R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed smult_r_distr by presburger + show ?thesis + unfolding 0 1 2 3 + using A times_x_product_rule P.m_closed R.one_closed UP_smult_assoc2 assms(2) is_UP_monomE(1) monom_is_UP_monom(1) pderiv_smult by presburger + next + case False + have IH: "pderiv ((f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n) = (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n" + using Suc A P.m_closed R.one_closed is_UP_monomE(1) is_UP_monomI by presburger + have 0: "f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n) = (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n" + using A R.one_closed UP_m_assoc assms(2) is_UP_monomE(1) monom_Suc(1) monom_is_UP_monom(1) by presburger + have 1: "(f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n = + (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> (f \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n " + using A times_x_product_rule by presburger + have 2: "(f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) =(f \\<^bsub>P\<^esub>up_ring.monom P ([n]\a) n)" + proof- + have 20: "up_ring.monom P ([n] \ a) (n) = up_ring.monom P \ 1 \\<^bsub>P\<^esub> up_ring.monom P ([n] \ a) (n - 1)" + using A assms False monom_mult[of \ "[n]\a" 1 "n-1"] + by (metis R.add.nat_pow_closed R.l_one R.one_closed Suc_eq_plus1 add.commute add_eq_if ) + show ?thesis unfolding 20 using assms A False pderiv_monom[of a n] + using P.m_assoc R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) by simp + qed + have 3: "(f \\<^bsub>P\<^esub>up_ring.monom P ([n]\a) n) = [n]\\<^bsub>P\<^esub>(f \\<^bsub>P\<^esub>up_ring.monom P a n)" + using A assms by (metis P.add_pow_rdistr add_nat_pow_monom is_UP_monomE(1) monom_is_UP_monom(1)) + have 4: "pderiv (f \\<^bsub>P\<^esub> up_ring.monom P \ 1) = (f \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1)" + using times_x_product_rule A by blast + have 5: " (f \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n = + (f \\<^bsub>P\<^esub> up_ring.monom P a n ) \\<^bsub>P\<^esub> (pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1 \\<^bsub>P\<^esub> up_ring.monom P a n )" + using A assms by (meson P.l_distr P.m_closed R.one_closed is_UP_monomE(1) is_UP_monomI pderiv_closed) + have 6: " (f \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n = + (f \\<^bsub>P\<^esub> up_ring.monom P a n ) \\<^bsub>P\<^esub> (pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1 \\<^bsub>P\<^esub> up_ring.monom P a n )" + using A assms False 5 by blast + have 7: "(f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> pderiv (up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv (f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n = + [(Suc n)] \\<^bsub>P\<^esub> (f \\<^bsub>P\<^esub> up_ring.monom P a n) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P \ 1 \\<^bsub>P\<^esub> up_ring.monom P a n" + unfolding 2 3 5 6 using assms A P.a_assoc + by (smt "1" "2" "3" "6" P.add.nat_pow_Suc P.m_closed R.one_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_closed) + have 8: "pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = pderiv ((f \\<^bsub>P\<^esub>up_ring.monom P \ 1) \\<^bsub>P\<^esub> up_ring.monom P a n)" + using A assms 0 by presburger + show " pderiv (f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)) = f \\<^bsub>P\<^esub> pderiv (up_ring.monom P a (Suc n)) \\<^bsub>P\<^esub> pderiv f \\<^bsub>P\<^esub> up_ring.monom P a (Suc n)" + unfolding 8 IH 0 1 2 3 4 5 6 + by (smt "2" "4" "6" "7" A P.add_pow_rdistr R.one_closed UP_m_assoc add_nat_pow_monom assms(2) diff_Suc_1 is_UP_monomE(1) is_UP_monomI monom_Suc(1) pderiv_closed pderiv_monom) + qed + qed + qed + qed + thus ?thesis using assms by blast +qed + +lemma(in UP_cring) product_rule: + assumes "f \ carrier (UP R)" + assumes "g \ carrier (UP R)" + shows "pderiv (f \\<^bsub>UP R\<^esub>g) = (pderiv f \\<^bsub>UP R\<^esub> g) \\<^bsub>UP R\<^esub> (f \\<^bsub>UP R\<^esub> pderiv g)" +proof(rule poly_induct3[of f]) + show "f \ carrier P" + using assms unfolding P_def by blast + show "\p q. q \ carrier P \ + p \ carrier P \ + pderiv (p \\<^bsub>UP R\<^esub> g) = pderiv p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> pderiv g \ + pderiv (q \\<^bsub>UP R\<^esub> g) = pderiv q \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> pderiv g \ + pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = pderiv (p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> (p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> pderiv g" + proof- fix p q + assume A: "q \ carrier P" "p \ carrier P" + "pderiv (p \\<^bsub>UP R\<^esub> g) = pderiv p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> pderiv g" + "pderiv (q \\<^bsub>UP R\<^esub> g) = pderiv q \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> pderiv g" + have 0: "(p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g = p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> g" + using A assms unfolding P_def using P_def UP_l_distr by blast + have 1: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = pderiv (p \\<^bsub>UP R\<^esub> g) \\<^bsub>UP R\<^esub> pderiv (q \\<^bsub>UP R\<^esub> g)" + unfolding 0 using pderiv_add[of "p \\<^bsub>P\<^esub> g" "q \\<^bsub>P\<^esub> g"] unfolding P_def + using A(1) A(2) P_def UP_mult_closed assms(2) by blast + have 2: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = pderiv p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> pderiv g \\<^bsub>UP R\<^esub> (pderiv q \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> pderiv g)" + unfolding 1 A by blast + have 3: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = pderiv p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> pderiv q \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> pderiv g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> pderiv g" + using A assms + by (smt "2" P.add.m_lcomm P.m_closed P_def UP_a_assoc pderiv_closed) + have 4: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = (pderiv p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> pderiv q \\<^bsub>UP R\<^esub> g) \\<^bsub>UP R\<^esub> (p \\<^bsub>UP R\<^esub> pderiv g \\<^bsub>UP R\<^esub> q \\<^bsub>UP R\<^esub> pderiv g)" + unfolding 3 using A assms P_def UP_a_assoc UP_a_closed UP_mult_closed pderiv_closed by auto + have 5: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = ((pderiv p \\<^bsub>UP R\<^esub> pderiv q) \\<^bsub>UP R\<^esub> g) \\<^bsub>UP R\<^esub> ((p \\<^bsub>UP R\<^esub> q) \\<^bsub>UP R\<^esub> pderiv g)" + unfolding 4 using A assms by (metis P.l_distr P_def pderiv_closed) + have 6: "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = ((pderiv (p \\<^bsub>P\<^esub> q)) \\<^bsub>UP R\<^esub> g) \\<^bsub>UP R\<^esub> ((p \\<^bsub>UP R\<^esub> q) \\<^bsub>UP R\<^esub> pderiv g)" + unfolding 5 using A assms + by (metis P_def pderiv_add) + show "pderiv ((p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g) = pderiv (p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> (p \\<^bsub>P\<^esub> q) \\<^bsub>UP R\<^esub> pderiv g" + unfolding 6 using A assms P_def by blast + qed + show "\a n. a \ carrier R \ + pderiv (up_ring.monom P a n \\<^bsub>UP R\<^esub> g) = pderiv (up_ring.monom P a n) \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> up_ring.monom P a n \\<^bsub>UP R\<^esub> pderiv g" + using P_def UP_m_comm assms(2) is_UP_monomE(1) monom_is_UP_monom(1) monom_product_rule pderiv_closed by presburger +qed + +(**************************************************************************************************) +(**************************************************************************************************) +subsection\The Chain Rule\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma(in UP_cring) chain_rule: + assumes "f \ carrier P" + assumes "g \ carrier P" + shows "pderiv (compose R f g) = compose R (pderiv f) g \\<^bsub>UP R\<^esub> pderiv g" +proof(rule poly_induct3[of f]) + show "f \ carrier P" + using assms by blast + show "\p q. q \ carrier P \ + p \ carrier P \ + pderiv (Cring_Poly.compose R p g) = Cring_Poly.compose R (pderiv p) g \\<^bsub>UP R\<^esub> pderiv g \ + pderiv (Cring_Poly.compose R q g) = Cring_Poly.compose R (pderiv q) g \\<^bsub>UP R\<^esub> pderiv g \ + pderiv (Cring_Poly.compose R (p \\<^bsub>P\<^esub> q) g) = Cring_Poly.compose R (pderiv (p \\<^bsub>P\<^esub> q)) g \\<^bsub>UP R\<^esub> pderiv g" + using pderiv_add sub_add + by (smt P_def UP_a_closed UP_m_comm UP_r_distr assms(2) pderiv_closed sub_closed) + show "\a n. a \ carrier R \ + pderiv (compose R (up_ring.monom P a n) g) = compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>UP R\<^esub> pderiv g" + proof- + fix a n assume A: "a \ carrier R" + show "pderiv (compose R (up_ring.monom P a n) g) = compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>UP R\<^esub> pderiv g" + proof(induction n) + case 0 + have 00: "(compose R (up_ring.monom P a 0) g) = (up_ring.monom P a 0)" + using A P_def assms(2) deg_const is_UP_monom_def monom_is_UP_monom(1) sub_const by presburger + have 01: "pderiv (up_ring.monom P a 0) = \\<^bsub>P\<^esub>" + using A pderiv_const by blast + show ?case unfolding 00 01 + by (metis P.l_null P_def UP_zero_closed assms(2) deg_zero pderiv_closed sub_const) + next + case (Suc n) + show "pderiv (Cring_Poly.compose R (up_ring.monom P a (Suc n)) g) = Cring_Poly.compose R (pderiv (up_ring.monom P a (Suc n))) g \\<^bsub>UP R\<^esub> pderiv g" + proof(cases "n = 0") + case True + have 0: "compose R (up_ring.monom P a (Suc n)) g = a \\<^bsub>P\<^esub> g" + using A assms sub_monom_deg_one[of g a] unfolding True using One_nat_def + by presburger + have 1: "(pderiv (up_ring.monom P a (Suc n))) = up_ring.monom P a 0" + unfolding True + proof - + have "pderiv (up_ring.monom P a 0) = \\<^bsub>P\<^esub>" + using A pderiv_const by blast + then show "pderiv (up_ring.monom P a (Suc 0)) = up_ring.monom P a 0" + using A lcf_monom(1) P_def X_closed deg_const deg_nzero_nzero is_UP_monomE(1) monom_Suc(2) monom_is_UP_monom(1) monom_rep_X_pow pderiv_monom poly_shift_degree_zero poly_shift_eq sub_monom(2) sub_monom_deg_one to_poly_inverse to_poly_mult_simp(2) + by (metis (no_types, lifting) P.l_null P.r_zero X_poly_def times_x_product_rule) + qed + then show ?thesis unfolding 0 1 + using A P_def assms(2) deg_const is_UP_monomE(1) monom_is_UP_monom(1) monom_mult_is_smult pderiv_closed pderiv_smult sub_const + by presburger + next + case False + have 0: "compose R (up_ring.monom P a (Suc n)) g = (compose R (up_ring.monom P a n) g) \\<^bsub>P\<^esub> (compose R (up_ring.monom P \ 1) g)" + using assms A by (metis R.one_closed monom_Suc(2) monom_closed sub_mult) + have 1: "compose R (up_ring.monom P a (Suc n)) g = (compose R (up_ring.monom P a n) g) \\<^bsub>P\<^esub> g" + unfolding 0 using A assms + by (metis P_def R.one_closed UP_cring.lcf_monom(1) UP_cring.to_poly_inverse UP_cring_axioms UP_l_one UP_one_closed deg_one monom_one sub_monom_deg_one to_poly_mult_simp(1)) + have 2: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = + ((pderiv (compose R (up_ring.monom P a n) g)) \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> ((compose R (up_ring.monom P a n) g) \\<^bsub>P\<^esub> pderiv g)" + unfolding 1 unfolding P_def apply(rule product_rule) + using A assms unfolding P_def using P_def is_UP_monomE(1) is_UP_monomI rev_sub_closed sub_rev_sub apply presburger + using assms unfolding P_def by blast + have 3: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = + (compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>UP R\<^esub> pderiv g \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> ((compose R (up_ring.monom P a n) g) \\<^bsub>P\<^esub> pderiv g)" + unfolding 2 Suc by blast + have 4: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = + ((compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>P\<^esub> g) \\<^bsub>UP R\<^esub> pderiv g) \\<^bsub>P\<^esub> ((compose R (up_ring.monom P a n) g) \\<^bsub>P\<^esub> pderiv g)" + unfolding 3 using A assms m_assoc m_comm + by (smt P_def monom_closed monom_rep_X_pow pderiv_closed sub_closed) + have 5: "pderiv (compose R (up_ring.monom P a (Suc n)) g ) = + ((compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> (compose R (up_ring.monom P a n) g)) \\<^bsub>P\<^esub> pderiv g" + unfolding 4 using A assms + by (metis P.l_distr P.m_closed P_def UP_cring.pderiv_closed UP_cring_axioms monom_closed sub_closed) + have 6: "compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>P\<^esub> g = [n]\\<^bsub>P\<^esub>compose R ((up_ring.monom P a n)) g" + proof- + have 60: "(pderiv (up_ring.monom P a n)) = (up_ring.monom P ([n]\a) (n-1))" + using A assms pderiv_monom by blast + have 61: "compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>P\<^esub> g = compose R ((up_ring.monom P ([n]\a) (n-1))) g \\<^bsub>P\<^esub> (compose R (up_ring.monom P \ 1) g)" + unfolding 60 using A assms sub_monom_deg_one[of g \ ] R.one_closed smult_one by presburger + have 62: "compose R (pderiv (up_ring.monom P a n)) g \\<^bsub>P\<^esub> g = compose R (up_ring.monom P ([n]\a) n) g" + unfolding 61 using False A assms sub_mult[of g "up_ring.monom P ([n] \ a) (n - 1)" "up_ring.monom P \ 1" ] monom_mult[of "[n]\a" \ "n-1" 1] + by (metis Nat.add_0_right R.add.nat_pow_closed R.one_closed R.r_one Suc_eq_plus1 add_eq_if monom_closed) + have 63: "\k::nat. Cring_Poly.compose R (up_ring.monom P ([k] \ a) n) g = [k] \\<^bsub>P\<^esub>Cring_Poly.compose R (up_ring.monom P a n) g" + proof- fix k::nat show "Cring_Poly.compose R (up_ring.monom P ([k] \ a) n) g = [k] \\<^bsub>P\<^esub>Cring_Poly.compose R (up_ring.monom P a n) g" + apply(induction k) + using UP_zero_closed assms(2) deg_zero monom_zero sub_const + apply (metis A P.add.nat_pow_0 add_nat_pow_monom) + proof- + fix k::nat + assume a: "Cring_Poly.compose R (monom P ([k] \ a) n) g = + [k] \\<^bsub>P\<^esub> Cring_Poly.compose R (monom P a n) g" + have 0: "(monom P ([Suc k] \ a) n) = [Suc k] \ a \\<^bsub>P\<^esub>(monom P \ n)" + by (simp add: A monic_monom_smult) + have 1: "(monom P ([Suc k] \ a) n) = [k] \ a \\<^bsub>P\<^esub>(monom P \ n) \\<^bsub>P\<^esub>a \\<^bsub>P\<^esub>(monom P \ n) " + unfolding 0 + by (simp add: A UP_smult_l_distr) + show "Cring_Poly.compose R (monom P ([Suc k] \ a) n) g = + [Suc k] \\<^bsub>P\<^esub> (Cring_Poly.compose R (monom P a n) g) " + unfolding 1 + by (simp add: A a assms(2) monic_monom_smult sub_add) + qed + qed + have 64: "Cring_Poly.compose R (up_ring.monom P ([n] \ a) n) g = [n] \\<^bsub>P\<^esub>Cring_Poly.compose R (up_ring.monom P a n) g" + using 63 by blast + show ?thesis unfolding 62 64 by blast + qed + have 63: "\k::nat. Cring_Poly.compose R (up_ring.monom P ([k] \ a) n) g = [k] \\<^bsub>P\<^esub>Cring_Poly.compose R (up_ring.monom P a n) g" + proof- fix k::nat show "Cring_Poly.compose R (up_ring.monom P ([k] \ a) n) g = [k] \\<^bsub>P\<^esub>Cring_Poly.compose R (up_ring.monom P a n) g" + apply(induction k) + using UP_zero_closed assms(2) deg_zero monom_zero sub_const + apply (metis A P.add.nat_pow_0 add_nat_pow_monom) + using A P.add.nat_pow_Suc add_nat_pow_monom assms(2) is_UP_monomE(1) monom_is_UP_monom(1) rev_sub_add sub_rev_sub + by (metis P.add.nat_pow_closed) + qed + have 7: "([n] \\<^bsub>P\<^esub> Cring_Poly.compose R (up_ring.monom P a n) g \\<^bsub>P\<^esub> Cring_Poly.compose R (up_ring.monom P a n) g) = + [Suc n] \\<^bsub>P\<^esub> (Cring_Poly.compose R (up_ring.monom P a n) g)" + using A assms P.add.nat_pow_Suc by presburger + have 8: "[Suc n] \\<^bsub>P\<^esub> Cring_Poly.compose R (up_ring.monom P a n) g \\<^bsub>P\<^esub> pderiv g = Cring_Poly.compose R (up_ring.monom P ([Suc n] \ a) n) g \\<^bsub>P\<^esub> pderiv g" + unfolding 63[of "Suc n"] by blast + show ?thesis unfolding 5 6 7 8 using A assms pderiv_monom[of "a" "Suc n"] + using P_def diff_Suc_1 by metis + qed + qed + qed +qed + +lemma deriv_prod_rule_times_monom: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "q \ carrier P" + shows "deriv ((monom P a n) \\<^bsub>P\<^esub> q) b = (deriv (monom P a n) b) \ (to_fun q b) \ (to_fun (monom P a n) b) \ deriv q b" +proof(rule poly_induct3[of q]) + show "q \ carrier P" + using assms by simp + show " \p q. q \ carrier P \ + p \ carrier P \ + deriv (monom P a n \\<^bsub>P\<^esub> p) b = deriv (monom P a n) b \ to_fun p b \ to_fun (monom P a n) b \ deriv p b \ + deriv (monom P a n \\<^bsub>P\<^esub> q) b = deriv (monom P a n) b \ to_fun q b \ to_fun (monom P a n) b \ deriv q b \ + deriv (monom P a n \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)) b = deriv (monom P a n) b \ to_fun (p \\<^bsub>P\<^esub> q) b \ to_fun (monom P a n) b \ deriv (p \\<^bsub>P\<^esub> q) b" + proof- fix p q assume A: "q \ carrier P" " p \ carrier P" + "deriv (monom P a n \\<^bsub>P\<^esub> p) b = deriv (monom P a n) b \ to_fun p b \ to_fun (monom P a n) b \ deriv p b" + "deriv (monom P a n \\<^bsub>P\<^esub> q) b = deriv (monom P a n) b \ to_fun q b \ to_fun (monom P a n) b \ deriv q b" + have "deriv (monom P a n \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)) b = deriv (monom P a n) b \ to_fun p b \ to_fun (monom P a n) b \ deriv p b + \deriv (monom P a n) b \ to_fun q b \ to_fun (monom P a n) b \ deriv q b" + using A assms + by (simp add: P.r_distr R.add.m_assoc deriv_add deriv_closed to_fun_closed) + hence "deriv (monom P a n \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)) b = deriv (monom P a n) b \ to_fun p b \deriv (monom P a n) b \ to_fun q b + \ to_fun (monom P a n) b \ deriv p b \ to_fun (monom P a n) b \ deriv q b" + using A(1) A(2) R.add.m_assoc R.add.m_comm assms(1) assms(2) deriv_closed to_fun_closed by auto + hence "deriv (monom P a n \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)) b = deriv (monom P a n) b \ (to_fun p b \ to_fun q b) + \ to_fun (monom P a n) b \ (deriv p b \ deriv q b)" + by (simp add: A(1) A(2) R.add.m_assoc R.r_distr assms(1) assms(2) deriv_closed to_fun_closed) + thus "deriv (monom P a n \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)) b = deriv (monom P a n) b \ to_fun (p \\<^bsub>P\<^esub> q) b \ to_fun (monom P a n) b \ deriv (p \\<^bsub>P\<^esub> q) b" + by (simp add: A(1) A(2) assms(2) deriv_add to_fun_plus) + qed + show "\c m. c \ carrier R \ deriv (monom P a n \\<^bsub>P\<^esub> monom P c m) b = + deriv (monom P a n) b \ to_fun (monom P c m) b + \ to_fun (monom P a n) b \ deriv (monom P c m) b" + proof- fix c m assume A: "c \ carrier R" + show "deriv (monom P a n \\<^bsub>P\<^esub> monom P c m) b = deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b" + proof(cases "n = 0") + case True + have LHS: "deriv (monom P a n \\<^bsub>P\<^esub> monom P c m) b = deriv (monom P (a \ c) m) b" + by (metis A True add.left_neutral assms(1) monom_mult) + have RHS: "deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b + = a \ deriv (monom P c m) b " + using deriv_const to_fun_monom A True assms(1) assms(2) deriv_closed by auto + show ?thesis using A assms LHS RHS deriv_monom + by (smt R.add.nat_pow_closed R.add_pow_rdistr R.m_assoc R.m_closed R.nat_pow_closed) + next + case False + show ?thesis + proof(cases "m = 0") + case True + have LHS: "deriv (monom P a n \\<^bsub>P\<^esub> monom P c m) b = deriv (monom P (a \ c) n) b" + by (metis A True add.comm_neutral assms(1) monom_mult) + have RHS: "deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b + = c \ deriv (monom P a n) b " + by (metis (no_types, lifting) A lcf_monom(1) P_def R.m_closed R.m_comm R.r_null + R.r_zero True UP_cring.to_fun_ctrm UP_cring_axioms assms(1) assms(2) deg_const + deriv_closed deriv_const to_fun_closed monom_closed) + show ?thesis using LHS RHS deriv_monom A assms + by (smt R.add.nat_pow_closed R.add_pow_ldistr R.m_assoc R.m_closed R.m_comm R.nat_pow_closed) + next + case F: False + have pos: "n > 0" "m >0" + using F False by auto + have RHS: "deriv (monom P a n \\<^bsub>P\<^esub> monom P c m) b = [(n + m)] \ (a \ c) \ b [^] (n + m - 1)" + using deriv_monom[of "a \ c" b "n + m"] monom_mult[of a c n m] + by (simp add: A assms(1) assms(2)) + have LHS: "deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b + = [n]\a \(b[^](n-1)) \ c \ b[^]m \ a \ b[^]n \ [m]\c \(b[^](m-1))" + using deriv_monom[of a b n] to_fun_monom[of a b n] + deriv_monom[of c b m] to_fun_monom[of c b m] A assms + by (simp add: R.m_assoc) + have 0: "[n]\a \ (b[^](n-1)) \ c \ b[^]m = [n]\a \ c \ b[^](n + m -1) " + proof- + have "[n]\a \ (b[^](n-1)) \ c \ b[^]m = [n]\a \ c \ (b[^](n-1)) \ b[^]m" + by (simp add: A R.m_lcomm R.semiring_axioms assms(1) assms(2) semiring.semiring_simprules(8)) + hence "[n]\a \ (b[^](n-1)) \ c \ b[^]m = [n]\a \ c \ ((b[^](n-1)) \ b[^]m)" + by (simp add: A R.m_assoc assms(1) assms(2)) + thus ?thesis + by (simp add: False R.nat_pow_mult add_eq_if assms(2)) + qed + have 1: "a \ b[^]n \ [m]\c \(b[^](m-1)) = a \ [m]\c \ b[^](n + m -1)" + proof- + have "a \ b[^]n \ [m]\c \(b[^](m-1)) = a \ [m]\c \ b[^]n \(b[^](m-1))" + using A R.m_comm R.m_lcomm assms(1) assms(2) by auto + hence "a \ b[^]n \ [m]\c \(b[^](m-1)) = a \ [m]\c \ (b[^]n \(b[^](m-1)))" + by (simp add: A R.m_assoc assms(1) assms(2)) + thus ?thesis + by (simp add: F R.nat_pow_mult add.commute add_eq_if assms(2)) + qed + have LHS: "deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b + = [n]\a \ c \ b[^](n + m -1) \ a \ [m]\c \ b[^](n + m -1)" + using LHS 0 1 + by simp + hence LHS: "deriv (monom P a n) b \ to_fun (monom P c m) b \ to_fun (monom P a n) b \ deriv (monom P c m) b + = [n]\ (a \ c \ b[^](n + m -1)) \ [m]\ (a \ c \ b[^](n + m -1))" + by (simp add: A R.add_pow_ldistr R.add_pow_rdistr assms(1) assms(2)) + show ?thesis using LHS RHS + by (simp add: A R.add.nat_pow_mult R.add_pow_ldistr assms(1) assms(2)) + qed + qed + qed +qed + +lemma deriv_prod_rule: + assumes "p \ carrier P" + assumes "q \ carrier P" + assumes "a \ carrier R" + shows "deriv (p \\<^bsub>P\<^esub> q) a = deriv p a \ (to_fun q a) \ (to_fun p a) \ deriv q a" +proof(rule poly_induct3[of p]) + show "p \ carrier P" + using assms(1) by simp + show " \p qa. + qa \ carrier P \ + p \ carrier P \ + deriv (p \\<^bsub>P\<^esub> q) a = deriv p a \ to_fun q a \ to_fun p a \ deriv q a \ + deriv (qa \\<^bsub>P\<^esub> q) a = deriv qa a \ to_fun q a \ to_fun qa a \ deriv q a \ + deriv ((p \\<^bsub>P\<^esub> qa) \\<^bsub>P\<^esub> q) a = deriv (p \\<^bsub>P\<^esub> qa) a \ to_fun q a \ to_fun (p \\<^bsub>P\<^esub> qa) a \ deriv q a" + proof- fix f g assume A: "f \ carrier P" "g \ carrier P" + "deriv (f \\<^bsub>P\<^esub> q) a = deriv f a \ to_fun q a \ to_fun f a \ deriv q a" + "deriv (g \\<^bsub>P\<^esub> q) a = deriv g a \ to_fun q a \ to_fun g a \ deriv q a" + have "deriv ((f \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q) a = deriv f a \ to_fun q a \ to_fun f a \ deriv q a \ + deriv g a \ to_fun q a \ to_fun g a \ deriv q a" + using A deriv_add + by (simp add: P.l_distr R.add.m_assoc assms(2) assms(3) deriv_closed to_fun_closed) + hence "deriv ((f \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q) a = deriv f a \ to_fun q a \ deriv g a \ to_fun q a \ + to_fun f a \ deriv q a \ to_fun g a \ deriv q a" + using R.a_comm R.a_assoc deriv_closed to_fun_closed assms + by (simp add: A(1) A(2)) + hence "deriv ((f \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q) a = (deriv f a \ to_fun q a \ deriv g a \ to_fun q a) \ + (to_fun f a \ deriv q a \ to_fun g a \ deriv q a)" + by (simp add: A(1) A(2) R.add.m_assoc assms(2) assms(3) deriv_closed to_fun_closed) + thus "deriv ((f \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q) a = deriv (f \\<^bsub>P\<^esub> g) a \ to_fun q a \ to_fun (f \\<^bsub>P\<^esub> g) a \ deriv q a" + by (simp add: A(1) A(2) R.l_distr assms(2) assms(3) deriv_add deriv_closed to_fun_closed to_fun_plus) + qed + show "\aa n. aa \ carrier R \ deriv (monom P aa n \\<^bsub>P\<^esub> q) a = deriv (monom P aa n) a \ to_fun q a \ to_fun (monom P aa n) a \ deriv q a" + using deriv_prod_rule_times_monom + by (simp add: assms(2) assms(3)) +qed + +lemma pderiv_eval_deriv_monom: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "to_fun (pderiv (monom P a n)) b = deriv (monom P a n) b" + using deriv_monom assms pderiv_monom + by (simp add: P_def UP_cring.to_fun_monom UP_cring_axioms) + +lemma pderiv_eval_deriv: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "deriv f a = to_fun (pderiv f) a" + apply(rule poly_induct3[of f]) + apply (simp add: assms(1)) + using assms(2) deriv_add to_fun_plus pderiv_add pderiv_closed apply presburger + using assms(2) pderiv_eval_deriv_monom + by presburger + +text\Taking taylor expansions commutes with taking derivatives:\ + +lemma(in UP_cring) taylor_expansion_pderiv_comm: + assumes "f \ carrier (UP R)" + assumes "c \ carrier R" + shows "pderiv (taylor_expansion R c f) = taylor_expansion R c (pderiv f)" + apply(rule poly_induct3[of f]) + using assms unfolding P_def apply blast +proof- + fix p q assume A: " q \ carrier (UP R)" "p \ carrier (UP R)" + "pderiv (taylor_expansion R c p) = taylor_expansion R c (pderiv p)" + "pderiv (taylor_expansion R c q) = taylor_expansion R c (pderiv q)" + have 0: " pderiv (taylor_expansion R c (p \\<^bsub>UP R\<^esub> q)) = pderiv (taylor_expansion R c p \\<^bsub>UP R\<^esub> taylor_expansion R c q)" + using A P_def taylor_expansion_add assms(2) by presburger + show "pderiv (taylor_expansion R c (p \\<^bsub>UP R\<^esub> q)) = taylor_expansion R c (pderiv (p \\<^bsub>UP R\<^esub> q))" + unfolding 0 + using A(1) A(2) A(3) A(4) taylor_def UP_cring.taylor_closed UP_cring.taylor_expansion_add UP_cring.pderiv_add UP_cring.pderiv_closed UP_cring_axioms assms(2) by fastforce +next + fix a n assume A: "a \ carrier R" + show "pderiv (taylor_expansion R c (up_ring.monom (UP R) a n)) = taylor_expansion R c (pderiv (up_ring.monom (UP R) a n))" + proof(cases "n = 0") + case True + have 0: "deg R (taylor_expansion R c (up_ring.monom (UP R) a n)) = 0" + unfolding True + using P_def A assms taylor_def taylor_deg deg_const is_UP_monomE(1) monom_is_UP_monom(2) by presburger + have 1: "(pderiv (up_ring.monom (UP R) a n)) = \\<^bsub>P\<^esub>" + unfolding True using P_def A assms pderiv_const by blast + show ?thesis unfolding 1 using 0 A assms P_def + by (metis P.add.right_cancel taylor_closed taylor_def taylor_expansion_add UP_l_zero UP_zero_closed monom_closed pderiv_deg_0) + next + case False + have 0: "pderiv (up_ring.monom (UP R) a n) = (up_ring.monom (UP R) ([n]\a) (n-1))" + using A + by (simp add: UP_cring.pderiv_monom UP_cring_axioms) + have 1: "pderiv (taylor_expansion R c (up_ring.monom (UP R) a n)) = (Cring_Poly.compose R (up_ring.monom (UP R) ([n]\a) (n-1)) (X_plus c)) \\<^bsub>P\<^esub> pderiv (X_plus c)" + using chain_rule[of "up_ring.monom (UP R) a n" "X_plus c"] unfolding 0 taylor_expansion_def + using A P_def X_plus_closed assms(2) is_UP_monom_def monom_is_UP_monom(1) by presburger + have 2: "pderiv (X_plus c) = \\<^bsub>P\<^esub>" + using pderiv_add[of "X_poly R" "to_poly c"] P.l_null P.l_one P.r_zero P_def R.one_closed X_closed + X_poly_def X_poly_plus_def assms(2) monom_one pderiv_const to_poly_closed to_polynomial_def + by (metis times_x_product_rule) + show ?thesis unfolding 1 0 2 taylor_expansion_def + by (metis "1" "2" A P.l_one P_def R.add.nat_pow_closed UP_m_comm UP_one_closed X_plus_closed assms(2) monom_closed sub_closed taylor_expansion_def) + qed +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Linear Substitutions\ + (**********************************************************************) + (**********************************************************************) +lemma(in UP_ring) lcoeff_Lcf: + assumes "f \ carrier P" + shows "lcoeff f = lcf f" + unfolding P_def + using assms coeff_simp[of f] by metis + +lemma(in UP_cring) linear_sub_cfs: + assumes "f \ carrier (UP R)" + assumes "d \ carrier R" + assumes "g = compose R f (up_ring.monom (UP R) d 1)" + shows "g i = d[^]i \ f i" +proof- + have 0: "(up_ring.monom (UP R) d 1) \ carrier (UP R)" + using assms by (meson R.ring_axioms UP_ring.intro UP_ring.monom_closed) + have 1: "(\i. compose R f (up_ring.monom (UP R) d 1) i = d[^]i \ f i)" + apply(rule poly_induct3[of f]) + using assms unfolding P_def apply blast + proof- + show "\p q. q \ carrier (UP R) \ + p \ carrier (UP R) \ + \i. Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i = d [^] i \ p i \ + \i. Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i = d [^] i \ q i \ + \i. Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = d [^] i \ (p \\<^bsub>UP R\<^esub> q) i" + proof + fix p q i + assume A: "q \ carrier (UP R)" + "p \ carrier (UP R)" + "\i. Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i = d [^] i \ p i" + "\i. Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i = d [^] i \ q i" + show "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = d [^] i \ (p \\<^bsub>UP R\<^esub> q) i" + proof- + have 1: "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) = + Cring_Poly.compose R p (up_ring.monom (UP R) d 1) \\<^bsub>UP R\<^esub> Cring_Poly.compose R q (up_ring.monom (UP R) d 1)" + using A(1) A(2) sub_add[of "up_ring.monom (UP R) d 1" q p] unfolding P_def + using "0" P_def sub_add by blast + have 2: "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = + Cring_Poly.compose R p (up_ring.monom (UP R) d 1) i \ Cring_Poly.compose R q (up_ring.monom (UP R) d 1) i" + using 1 by (metis (no_types, lifting) "0" A(1) A(2) P_def cfs_add sub_closed) + have 3: "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = d [^] i \ p i \ d [^] i \ q i" + unfolding 2 using A by presburger + have 4: "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = d [^] i \ (p i \ q i)" + using "3" A(1) A(2) R.nat_pow_closed R.r_distr UP_car_memE(1) assms(2) by presburger + thus "Cring_Poly.compose R (p \\<^bsub>UP R\<^esub> q) (up_ring.monom (UP R) d 1) i = d [^] i \ (p \\<^bsub>UP R\<^esub> q) i" + unfolding 4 using A(1) A(2) P_def cfs_add by presburger + qed + qed + show "\a n. a \ carrier R \ + \i. Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) i = d [^] i \ up_ring.monom (UP R) a n i" + proof fix a n i assume A: "a \ carrier R" + have 0: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) = + a \\<^bsub>UP R\<^esub>(up_ring.monom (UP R) d 1)[^]\<^bsub>UP R\<^esub>n" + using assms A 0 P_def monom_sub by blast + have 1: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) = + a \\<^bsub>UP R\<^esub> (d[^]n \\<^bsub>UP R\<^esub>(up_ring.monom (UP R) \ n))" + unfolding 0 using A assms + by (metis P_def R.nat_pow_closed monic_monom_smult monom_pow mult.left_neutral) + have 2: "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) = + (a \d[^]n)\\<^bsub>UP R\<^esub>(up_ring.monom (UP R) \ n)" + unfolding 1 using A assms + by (metis R.nat_pow_closed R.one_closed R.ring_axioms UP_ring.UP_smult_assoc1 UP_ring.intro UP_ring.monom_closed) + show "Cring_Poly.compose R (up_ring.monom (UP R) a n) (up_ring.monom (UP R) d 1) i = d [^] i \ up_ring.monom (UP R) a n i" + apply(cases "i = n") + unfolding 2 using A P_def R.m_closed R.m_comm R.nat_pow_closed assms(2) cfs_monom monic_monom_smult apply presburger + using A P_def R.m_closed R.nat_pow_closed R.r_null assms(2) cfs_monom monic_monom_smult by presburger + qed + qed + show ?thesis using 1 unfolding assms + by blast +qed + +lemma(in UP_cring) linear_sub_deriv: + assumes "f \ carrier (UP R)" + assumes "d \ carrier R" + assumes "g = compose R f (up_ring.monom (UP R) d 1)" + assumes "c \ carrier R" + shows "pderiv g = d \\<^bsub>UP R\<^esub> compose R (pderiv f) (up_ring.monom (UP R) d 1)" + unfolding assms +proof(rule poly_induct3[of f]) + show "f \ carrier P" + using assms unfolding P_def by blast + show "\ p q. q \ carrier P \ + p \ carrier P \ + pderiv (Cring_Poly.compose R p (up_ring.monom (UP R) d 1)) = d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv p) (up_ring.monom (UP R) d 1) \ + pderiv (Cring_Poly.compose R q (up_ring.monom (UP R) d 1)) = d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv q) (up_ring.monom (UP R) d 1) \ + pderiv (Cring_Poly.compose R (p \\<^bsub>P\<^esub> q) (up_ring.monom (UP R) d 1)) = + d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv (p \\<^bsub>P\<^esub> q)) (up_ring.monom (UP R) d 1)" + proof- fix p q assume A: "q \ carrier P" "p \ carrier P" + "pderiv (Cring_Poly.compose R p (up_ring.monom (UP R) d 1)) = d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv p) (up_ring.monom (UP R) d 1)" + "pderiv (Cring_Poly.compose R q (up_ring.monom (UP R) d 1)) = d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv q) (up_ring.monom (UP R) d 1)" + show " pderiv (Cring_Poly.compose R (p \\<^bsub>P\<^esub> q) (up_ring.monom (UP R) d 1)) = + d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv (p \\<^bsub>P\<^esub> q)) (up_ring.monom (UP R) d 1)" + using A assms + by (smt P_def UP_a_closed UP_r_distr monom_closed monom_mult_is_smult pderiv_add pderiv_closed rev_sub_add sub_closed sub_rev_sub) + qed + show "\a n. a \ carrier R \ + pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = + d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)" + proof- fix a n assume A: "a \ carrier R" + have "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = a \\<^bsub>UP R\<^esub> (up_ring.monom P d 1)[^]\<^bsub>UP R\<^esub> n" + using A assms sub_monom(2) P_def is_UP_monomE(1) monom_is_UP_monom(1) by blast + hence 0: "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = a \\<^bsub>UP R\<^esub> (up_ring.monom P (d[^]n) n)" + using A assms P_def monom_pow nat_mult_1 by metis + show "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = + d \\<^bsub>UP R\<^esub> Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)" + proof(cases "n = 0") + case True + have T0: "pderiv (up_ring.monom P a n) = \\<^bsub> UP R\<^esub>" unfolding True + using A P_def pderiv_const by blast + have T1: "(Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = up_ring.monom P a n" + unfolding True + using A assms P_def deg_const is_UP_monomE(1) monom_is_UP_monom(1) sub_const by presburger + thus ?thesis unfolding T0 T1 + by (metis P.nat_pow_eone P_def UP_smult_closed UP_zero_closed X_closed assms(2) deg_zero monom_rep_X_pow smult_r_null sub_const) + next + case False + have F0: "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = (a \\<^bsub>UP R\<^esub> (up_ring.monom P ([n]\\<^bsub>R\<^esub>(d[^]n))(n-1)))" + using A assms pderiv_monom unfolding 0 + using P_def R.nat_pow_closed is_UP_monomE(1) monom_is_UP_monom(1) pderiv_smult by metis + have F1: "(pderiv (up_ring.monom P a n)) = up_ring.monom P ([n] \ a) (n - 1)" + using A assms pderiv_monom[of a n] by blast + hence F2: "(pderiv (up_ring.monom P a n)) = ([n] \ a) \\<^bsub>UP R\<^esub>up_ring.monom P \ (n - 1)" + using A P_def monic_monom_smult by auto + have F3: "Cring_Poly.compose R ((([n] \ a) \\<^bsub>UP R\<^esub> (up_ring.monom P \ (n - 1)))) (up_ring.monom (UP R) d 1) = + ([n] \ a) \\<^bsub>UP R\<^esub> ((up_ring.monom (UP R) d 1)[^]\<^bsub>UP R\<^esub>(n-1))" + using A F1 F2 P_def assms(2) monom_closed sub_monom(2) by fastforce + have F4: "Cring_Poly.compose R ((([n] \ a) \\<^bsub>UP R\<^esub> (up_ring.monom P \ (n - 1)))) (up_ring.monom (UP R) d 1) = + ([n] \ a) \\<^bsub>UP R\<^esub> ((up_ring.monom (UP R) (d[^](n-1)) (n-1)))" + by (metis F3 P_def assms(2) monom_pow nat_mult_1) + have F5: "d \\<^bsub>UP R\<^esub> (Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)) = + (d \([n] \ a)) \\<^bsub>UP R\<^esub> up_ring.monom (UP R) (d [^] (n - 1)) (n - 1)" + unfolding F4 F2 + using A P_def assms(2) monom_closed smult_assoc1 by auto + have F6: "d \\<^bsub>UP R\<^esub> (Cring_Poly.compose R (pderiv (up_ring.monom P a n)) (up_ring.monom (UP R) d 1)) = + (d \ d[^](n-1) \[n] \ a) \\<^bsub>UP R\<^esub> ((up_ring.monom (UP R) \ (n-1)))" + unfolding F5 using False A assms P_def R.m_assoc R.m_closed R.m_comm R.nat_pow_closed monic_monom_smult monom_mult_smult + by (smt R.add.nat_pow_closed) + have F7: "pderiv (Cring_Poly.compose R (up_ring.monom P a n) (up_ring.monom (UP R) d 1)) = (a \ ([n]\\<^bsub>R\<^esub>(d[^]n)) \\<^bsub>UP R\<^esub> (up_ring.monom P \ (n-1)))" + unfolding F0 using A assms P_def R.m_closed R.nat_pow_closed monic_monom_smult monom_mult_smult + by simp + have F8: "a \ [n] \ (d [^] n) = d \ d [^] (n - 1) \ [n] \ a" + proof- + have F80: "d \ d [^] (n - 1) \ [n] \ a = d [^] n \ [n] \ a" + using A assms False by (metis R.nat_pow_Suc2 add.right_neutral add_eq_if) + show ?thesis unfolding F80 + using A R.add_pow_rdistr R.m_comm R.nat_pow_closed assms(2) by presburger + qed + show ?thesis unfolding F6 F7 unfolding F8 P_def by blast + qed + qed +qed + +lemma(in UP_cring) linear_sub_deriv': + assumes "f \ carrier (UP R)" + assumes "d \ carrier R" + assumes "g = compose R f (up_ring.monom (UP R) d 1)" + assumes "c \ carrier R" + shows "pderiv g = compose R (d \\<^bsub>UP R\<^esub> pderiv f) (up_ring.monom (UP R) d 1)" + using assms linear_sub_deriv[of f d g c] P_def is_UP_monomE(1) is_UP_monomI pderiv_closed sub_smult by metis + +lemma(in UP_cring) linear_sub_inv: + assumes "f \ carrier (UP R)" + assumes "d \ Units R" + assumes "g = compose R f (up_ring.monom (UP R) d 1)" + shows "compose R g (up_ring.monom (UP R) (inv d) 1) = f" + unfolding assms +proof fix x + have 0: "Cring_Poly.compose R (Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) (up_ring.monom (UP R) (inv d) 1) x = + (inv d)[^]x \ ((Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) x)" + apply(rule linear_sub_cfs) + using P_def R.Units_closed assms(1) assms(2) monom_closed sub_closed apply auto[1] + apply (simp add: assms(2)) + by blast + show "Cring_Poly.compose R (Cring_Poly.compose R f (up_ring.monom (UP R) d 1)) (up_ring.monom (UP R) (inv d) 1) x = f x " + unfolding 0 using linear_sub_cfs[of f d "Cring_Poly.compose R f (up_ring.monom (UP R) d 1)" x] + assms + by (smt R.Units_closed R.Units_inv_closed R.Units_l_inv R.m_assoc R.m_comm R.nat_pow_closed R.nat_pow_distrib R.nat_pow_one R.r_one UP_car_memE(1)) +qed + +lemma(in UP_cring) linear_sub_deg: + assumes "f \ carrier (UP R)" + assumes "d \ Units R" + assumes "g = compose R f (up_ring.monom (UP R) d 1)" + shows "deg R g = deg R f" +proof(cases "deg R f = 0") + case True + show ?thesis using assms + unfolding True assms using P_def True monom_closed + by (simp add: R.Units_closed sub_const) +next + case False + have 0: "monom (UP R) d 1 (deg R (monom (UP R) d 1)) = d" + using assms lcf_monom(2) by blast + have 1: "d[^](deg R f) \ Units R" + using assms(2) + by (metis Group.comm_monoid.axioms(1) R.units_comm_group R.units_of_pow comm_group_def monoid.nat_pow_closed units_of_carrier) + have 2: "f (deg R f) \ \" + using assms False P_def UP_cring.ltrm_rep_X_pow UP_cring_axioms deg_ltrm degree_monom by fastforce + have "deg R g = deg R f * deg R (up_ring.monom (UP R) d 1)" + unfolding assms + apply(rule cring_sub_deg[of "up_ring.monom (UP R) d 1" f] ) + using assms P_def monom_closed apply blast + unfolding P_def apply(rule assms) + unfolding 0 using 2 1 + by (metis R.Units_closed R.Units_l_cancel R.m_comm R.r_null R.zero_closed UP_car_memE(1) assms(1)) + thus ?thesis using assms unfolding assms + by (metis (no_types, lifting) P_def R.Units_closed deg_monom deg_zero is_UP_monomE(1) linear_sub_inv monom_is_UP_monom(2) monom_zero mult.right_neutral mult_0_right sub_closed sub_const) +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Lemmas About Polynomial Division\ +(**************************************************************************************************) +(**************************************************************************************************) +context UP_cring +begin + + (**********************************************************************) + (**********************************************************************) + subsection\Division by Linear Terms\ + (**********************************************************************) + (**********************************************************************) +definition UP_root_div where +"UP_root_div f a = (poly_shift (T\<^bsub>a\<^esub> f)) of (X_minus a)" + +definition UP_root_rem where +"UP_root_rem f a = ctrm (T\<^bsub>a\<^esub> f)" + +lemma UP_root_div_closed: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "UP_root_div f a \ carrier P" + using assms + unfolding UP_root_div_def + by (simp add: taylor_closed X_minus_closed poly_shift_closed sub_closed) + +lemma rem_closed: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "UP_root_rem f a \ carrier P" + using assms + unfolding UP_root_rem_def + by (simp add: taylor_closed ctrm_is_poly) + +lemma rem_deg: + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "degree (UP_root_rem f a) = 0" + by (simp add: taylor_closed assms(1) assms(2) ctrm_degree UP_root_rem_def) + +lemma remainder_theorem: + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "g = UP_root_div f a" + assumes "r = UP_root_rem f a" + shows "f = r \\<^bsub>P\<^esub> (X_minus a) \\<^bsub>P\<^esub> g" +proof- + have "T\<^bsub>a\<^esub>f = (ctrm (T\<^bsub>a\<^esub>f)) \\<^bsub>P\<^esub> X \\<^bsub>P\<^esub> poly_shift (T\<^bsub>a\<^esub>f)" + using poly_shift_eq[of "T\<^bsub>a\<^esub>f"] assms taylor_closed + by blast + hence 1: "T\<^bsub>a\<^esub>f of (X_minus a) = (ctrm (T\<^bsub>a\<^esub>f)) \\<^bsub>P\<^esub> (X_minus a) \\<^bsub>P\<^esub> (poly_shift (T\<^bsub>a\<^esub>f) of (X_minus a))" + using assms taylor_closed[of f a] X_minus_closed[of a] X_closed + sub_add[of "X_minus a" "ctrm (T\<^bsub>a\<^esub>f)" "X \\<^bsub>P\<^esub> poly_shift (T\<^bsub>a\<^esub>f)"] + sub_const[of "X_minus a"] sub_mult[of "X_minus a" X "poly_shift (T\<^bsub>a\<^esub>f)"] + ctrm_degree ctrm_is_poly P.m_closed poly_shift_closed sub_X + by presburger + have 2: "f = (ctrm (T\<^bsub>a\<^esub>f)) \\<^bsub>P\<^esub> (X_minus a) \\<^bsub>P\<^esub> (poly_shift (T\<^bsub>a\<^esub>f) of (X_minus a))" + using 1 taylor_id[of a f] assms + by simp + then show ?thesis + using assms + unfolding UP_root_div_def UP_root_rem_def + by auto +qed + +lemma remainder_theorem': + assumes "f \ carrier P" + assumes "a \ carrier R" + shows "f = UP_root_rem f a \\<^bsub>P\<^esub> (X_minus a) \\<^bsub>P\<^esub> UP_root_div f a" + using assms remainder_theorem by auto + +lemma factor_theorem: + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "g = UP_root_div f a" + assumes "to_fun f a = \" + shows "f = (X_minus a) \\<^bsub>P\<^esub> g" + using remainder_theorem[of f a g _] assms + unfolding UP_root_rem_def + by (simp add: ctrm_zcf taylor_zcf taylor_closed UP_root_div_closed X_minus_closed) + +lemma factor_theorem': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "to_fun f a = \" + shows "f = (X_minus a) \\<^bsub>P\<^esub> UP_root_div f a" + by (simp add: assms(1) assms(2) assms(3) factor_theorem) + + (**********************************************************************) + (**********************************************************************) + subsection\Geometric Sums\ + (**********************************************************************) + (**********************************************************************) +lemma geom_quot: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "p = monom P \ (Suc n) \\<^bsub>P\<^esub> monom P (b[^](Suc n)) 0 " + assumes "g = UP_root_div p b" + shows "a[^](Suc n) \ b[^] (Suc n) = (a \ b) \ (to_fun g a)" +proof- + have root: "to_fun p b = \" + using assms to_fun_const[of "b[^](Suc n)" b] to_fun_monic_monom[of b "Suc n"] R.nat_pow_closed[of b "Suc n"] + to_fun_diff[of "monom P \ (Suc n)" "monom P (b[^](Suc n)) 0" b] monom_closed + by (metis P.minus_closed P_def R.one_closed R.zero_closed UP_cring.f_minus_ctrm + UP_cring.to_fun_diff UP_cring_axioms zcf_to_fun cfs_monom to_fun_const) + have LHS: "to_fun p a = a[^](Suc n) \ b[^] (Suc n)" + using assms to_fun_const to_fun_monic_monom to_fun_diff + by auto + have RHS: "to_fun ((X_minus b) \\<^bsub>P\<^esub> g) a = (a \ b) \ (to_fun g a)" + using to_fun_mult[of g "X_minus b"] assms X_minus_closed + by (metis P.minus_closed P_def R.nat_pow_closed R.one_closed UP_cring.UP_root_div_closed UP_cring_axioms to_fun_X_minus monom_closed) + show ?thesis + using RHS LHS root factor_theorem' assms(2) assms(3) assms(4) + by auto +qed + +end + +context UP_cring +begin + +definition geometric_series where +"geometric_series n a b = to_fun (UP_root_div (monom P \ (Suc n) \\<^bsub>UP R\<^esub> (monom P (b[^](Suc n)) 0)) b) a" + +lemma geometric_series_id: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "a[^](Suc n) \b[^] (Suc n) = (a \ b) \ (geometric_series n a b)" + using assms geom_quot + by (simp add: P_def geometric_series_def) + +lemma geometric_series_closed: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "(geometric_series n a b) \ carrier R" + unfolding geometric_series_def + using assms P.minus_closed P_def UP_root_div_closed to_fun_closed monom_closed + by auto + +text\Shows that $a^n - b^n$ has $a - b$ as a factor:\ +lemma to_fun_monic_monom_diff: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "\c. c \ carrier R \ to_fun (monom P \ n) a \ to_fun (monom P \ n) b = (a \ b) \ c" +proof(cases "n = 0") + case True + have "to_fun (monom P \ 0) a \ to_fun (monom P \ 0) b = (a \ b) \ \" + unfolding a_minus_def using to_fun_const[of \] assms + by (simp add: R.r_neg) + then show ?thesis + using True by blast +next + case False + then show ?thesis + using Suc_diff_1[of n] geometric_series_id[of a b "n-1"] geometric_series_closed[of a b "n-1"] + assms(1) assms(2) to_fun_monic_monom + by auto +qed + +lemma to_fun_diff_factor: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "f \ carrier P" + shows "\c. c \ carrier R \(to_fun f a) \ (to_fun f b) = (a \ b)\c" +proof(rule poly_induct5[of f]) + show "f \ carrier P" using assms by simp + show "\p q. q \ carrier P \ + p \ carrier P \ + \c. c \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ c \ + \c. c \ carrier R \ to_fun q a \ to_fun q b = (a \ b) \ c \ + \c. c \ carrier R \ to_fun (p \\<^bsub>P\<^esub> q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = (a \ b) \ c" + proof- fix p q assume A: "q \ carrier P" "p \ carrier P" + "\c. c \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ c" + "\c. c \ carrier R \ to_fun q a \ to_fun q b = (a \ b) \ c" + obtain c where c_def: "c \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ c" + using A by blast + obtain c' where c'_def: "c' \ carrier R \ to_fun q a \ to_fun q b = (a \ b) \ c'" + using A by blast + have 0: "(a \ b) \ c \ (a \ b) \ c' = (a \ b)\(c \ c')" + using assms c_def c'_def unfolding a_minus_def + by (simp add: R.r_distr R.r_minus) + have 1: "to_fun (p \\<^bsub>P\<^esub>q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = to_fun p a \ to_fun q a \ to_fun p b \ to_fun q b" + using A to_fun_plus[of p q a] to_fun_plus[of p q b] assms to_fun_closed + R.ring_simprules(19)[of "to_fun p b" "to_fun q b"] + by (simp add: R.add.m_assoc R.minus_eq to_fun_plus) + hence "to_fun (p \\<^bsub>P\<^esub>q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = to_fun p a \ to_fun p b \ to_fun q a \ to_fun q b" + using 0 A assms R.ring_simprules to_fun_closed a_assoc a_comm + unfolding a_minus_def + by smt + hence "to_fun (p \\<^bsub>P\<^esub>q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = to_fun p a \ to_fun p b \ (to_fun q a \ to_fun q b)" + using 0 A assms R.ring_simprules to_fun_closed + unfolding a_minus_def + by metis + hence "to_fun (p \\<^bsub>P\<^esub>q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = (a \ b)\(c \ c')" + using 0 A c_def c'_def + by simp + thus "\c. c \ carrier R \ to_fun (p \\<^bsub>P\<^esub> q) a \ to_fun (p \\<^bsub>P\<^esub> q) b = (a \ b) \ c" + using R.add.m_closed c'_def c_def by blast + qed + show "\n. \c. c \ carrier R \ to_fun (monom P \ n) a \ to_fun (monom P \ n) b = (a \ b) \ c" + by (simp add: assms(1) assms(2) to_fun_monic_monom_diff) + show "\p aa. + aa \ carrier R \ + p \ carrier P \ \c. c \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ c \ \c. c \ carrier R \ to_fun (aa \\<^bsub>P\<^esub> p) a \ to_fun (aa \\<^bsub>P\<^esub> p) b = (a \ b) \ c" + proof- fix p c assume A: "c \ carrier R" " p \ carrier P" + "\e. e \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ e" + then obtain d where d_def: "d \ carrier R \ to_fun p a \ to_fun p b = (a \ b) \ d" + by blast + have "to_fun (c \\<^bsub>P\<^esub> p) a \ to_fun (c \\<^bsub>P\<^esub> p) b = c \ (to_fun p a \ to_fun p b)" + using A d_def assms to_fun_smult[of p a c] to_fun_smult[of p b c] + to_fun_closed[of p a] to_fun_closed[of p b] R.ring_simprules + by smt + hence "c\d \ carrier R \ to_fun (c \\<^bsub>P\<^esub> p) a \ to_fun (c \\<^bsub>P\<^esub> p) b = (a \ b) \ (c \d)" + by (simp add: A(1) R.m_lcomm assms(1) assms(2) d_def) + thus "\e. e \ carrier R \ to_fun (c \\<^bsub>P\<^esub> p) a \ to_fun (c \\<^bsub>P\<^esub> p) b = (a \ b) \ e" + by blast + qed +qed + +text\Any finite set over a domain is the zero set of a polynomial:\ +lemma(in UP_domain) fin_set_poly_roots: + assumes "F \ carrier R" + assumes "finite F" + shows "\ P \ carrier (UP R). \ x \ carrier R. to_fun P x = \ \ x \ F" + apply(rule finite.induct) + apply (simp add: assms(2)) +proof- + show "\P\carrier (UP R). \x\carrier R. (to_fun P x = \) = (x \ {})" + proof- + have "\x\carrier R. (to_fun (\\<^bsub>UP R\<^esub>) x = \) = (x \ {})" + proof + fix x + assume A: "x \ carrier R" + then have "(to_fun (\\<^bsub>UP R\<^esub>)) x = \" + by (metis P_def R.one_closed UP_cring.to_fun_to_poly UP_cring_axioms ring_hom_one to_poly_is_ring_hom) + then show "(to_fun \\<^bsub>UP R\<^esub> x = \) = (x \ {})" + by simp + qed + then show ?thesis + using P_def UP_one_closed + by blast + qed + show "\A a. finite A \ + \P\carrier (UP R). \x\carrier R. (to_fun P x = \) = (x \ A) \ \P\carrier (UP R). \x\carrier R. (to_fun P x = \) = (x \ insert a A)" + proof- + fix A :: "'a set" fix a + assume fin_A: "finite A" + assume IH: "\P\carrier (UP R). \x\carrier R. (to_fun P x = \) = (x \ A)" + then obtain p where p_def: "p \carrier (UP R) \ (\x\carrier R. (to_fun p x = \) = (x \ A))" + by blast + show "\P\carrier (UP R). \x\carrier R. (to_fun P x = \) = (x \ insert a A)" + proof(cases "a \ carrier R") + case True + obtain Q where Q_def: "Q = p \\<^bsub>UP R\<^esub> (X \\<^bsub>UP R\<^esub> to_poly a)" + by blast + have "\x\carrier R. (to_fun Q x = \) = (x \ insert a A)" + proof fix x + assume P: "x \ carrier R" + have P0: "to_fun (X \\<^bsub>UP R\<^esub> to_poly a) x = x \ a" + using to_fun_plus[of X "\\<^bsub>UP R\<^esub> to_poly a" x] True P + unfolding a_minus_def + by (metis X_poly_minus_def a_minus_def to_fun_X_minus) + then have "to_fun Q x = (to_fun p x) \ (x \ a)" + proof- + have 0: " p \ carrier P" + by (simp add: P_def p_def) + have 1: " X \\<^bsub>UP R\<^esub> to_poly a \ carrier P" + using P.minus_closed P_def True X_closed to_poly_closed by auto + have 2: "x \ carrier R" + by (simp add: P) + then show ?thesis + using to_fun_mult[of p "(X \\<^bsub>UP R\<^esub> to_poly a)" x] P0 0 1 2 Q_def True P_def to_fun_mult + by auto + qed + then show "(to_fun Q x = \) = (x \ insert a A)" + using p_def + by (metis P R.add.inv_closed R.integral_iff R.l_neg R.minus_closed R.minus_unique True UP_cring.to_fun_closed UP_cring_axioms a_minus_def insert_iff) + qed + then have "Q \ carrier (UP R) \ (\x\carrier R. (to_fun Q x = \) = (x \ insert a A))" + using P.minus_closed P_def Q_def True UP_mult_closed X_closed p_def to_poly_closed by auto + then show ?thesis + by blast + next + case False + then show ?thesis + using IH subsetD by auto + qed + qed +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Polynomial Evaluation at Multiplicative Inverses\ + (**********************************************************************) + (**********************************************************************) +text\For every polynomial $p(x)$ of degree $n$, there is a unique polynomial $q(x)$ which satisfies the equation $q(x) = x^n p(1/x)$. This section defines this polynomial and proves this identity.\ +definition(in UP_cring) one_over_poly where +"one_over_poly p = (\ n. if n \ degree p then p ((degree p) - n) else \)" + +lemma(in UP_cring) one_over_poly_closed: + assumes "p \ carrier P" + shows "one_over_poly p \ carrier P" + apply(rule UP_car_memI[of "degree p" ]) + unfolding one_over_poly_def using assms apply simp + by (simp add: assms cfs_closed) + +lemma(in UP_cring) one_over_poly_monom: + assumes "a \ carrier R" + shows "one_over_poly (monom P a n) = monom P a 0" + apply(rule ext) + unfolding one_over_poly_def using assms + by (metis cfs_monom deg_monom diff_diff_cancel diff_is_0_eq diff_self_eq_0 zero_diff) + +lemma(in UP_cring) one_over_poly_monom_add: + assumes "a \ carrier R" + assumes "a \ \" + assumes "p \ carrier P" + assumes "degree p < n" + shows "one_over_poly (p \\<^bsub>P\<^esub> monom P a n) = monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - degree p) \\<^bsub>P\<^esub> one_over_poly p" +proof- + have 0: "degree (p \\<^bsub>P\<^esub> monom P a n) = n" + by (simp add: assms(1) assms(2) assms(3) assms(4) equal_deg_sum) + show ?thesis + proof(rule ext) fix x show "one_over_poly (p \\<^bsub>P\<^esub> monom P a n) x = + (monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - deg R p) \\<^bsub>P\<^esub> one_over_poly p) x" + proof(cases "x = 0") + case T: True + have T0: "one_over_poly (p \\<^bsub>P\<^esub> monom P a n) 0 = a" + unfolding one_over_poly_def + by (metis lcf_eq lcf_monom(1) ltrm_of_sum_diff_deg P.add.m_closed assms(1) assms(2) assms(3) assms(4) diff_zero le0 monom_closed) + have T1: "(monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - degree p) \\<^bsub>P\<^esub> one_over_poly p) 0 = a" + using one_over_poly_closed + by (metis (no_types, lifting) lcf_monom(1) R.one_closed R.r_zero UP_m_comm UP_mult_closed assms(1) assms(3) assms(4) cfs_add cfs_monom_mult deg_const monom_closed zero_less_diff) + show ?thesis using T0 T1 T by auto + next + case F: False + show ?thesis + proof(cases "x < n - degree p") + case True + then have T0: "degree p < n - x \ n - x < n" + using F by auto + then have T1: "one_over_poly (p \\<^bsub>P\<^esub> monom P a n) x = \" + using True F 0 unfolding one_over_poly_def + using assms(1) assms(3) coeff_of_sum_diff_degree0 + by (metis ltrm_cfs ltrm_of_sum_diff_deg P.add.m_closed P.add.m_comm assms(2) assms(4) monom_closed nat_neq_iff) + have "(monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - degree p) \\<^bsub>P\<^esub> one_over_poly p) x = \" + using True F 0 one_over_poly_def one_over_poly_closed + by (metis (no_types, lifting) P.add.m_comm P.m_closed R.one_closed UP_m_comm assms(1) + assms(3) cfs_monom_mult coeff_of_sum_diff_degree0 deg_const monom_closed neq0_conv) + then show ?thesis using T1 by auto + next + case False + then have "n - degree p \ x" + by auto + then obtain k where k_def: "k + (n - degree p) = x" + using le_Suc_ex diff_add + by blast + have F0: "(monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - deg R p) \\<^bsub>P\<^esub> one_over_poly p) x + = one_over_poly p k" + using k_def one_over_poly_closed assms + times_X_pow_coeff[of "one_over_poly p" "n - deg R p" k] + P.m_closed + by (metis (no_types, lifting) P.add.m_comm R.one_closed add_gr_0 coeff_of_sum_diff_degree0 deg_const monom_closed zero_less_diff) + show ?thesis + proof(cases "x \ n") + case True + have T0: "n - x = degree p - k" + using assms(4) k_def by linarith + have T1: "n - x < n" + using True F + by linarith + then have F1: "(p \\<^bsub>P\<^esub> monom P a n) (n - x) = p (degree p - k)" + using True False F0 0 k_def cfs_add + by (simp add: F0 T0 assms(1) assms(3) cfs_closed cfs_monom) + then show ?thesis + using "0" F0 assms(1) assms(2) assms(3) degree_of_sum_diff_degree k_def one_over_poly_def + by auto + next + case False + then show ?thesis + using "0" F0 assms(1) assms(2) assms(3) degree_of_sum_diff_degree k_def one_over_poly_def + by auto + qed + qed + qed + qed +qed + +lemma( in UP_cring) one_over_poly_eval: + assumes "p \ carrier P" + assumes "x \ carrier R" + assumes "x \ Units R" + shows "to_fun (one_over_poly p) x = (x[^](degree p)) \ (to_fun p (inv\<^bsub>R\<^esub> x))" +proof(rule poly_induct6[of p]) + show " p \ carrier P" + using assms by simp + show "\a n. a \ carrier R \ + to_fun (one_over_poly (monom P a 0)) x = x [^] deg R (monom P a 0) \ to_fun (monom P a 0) (inv x)" + using assms to_fun_const one_over_poly_monom by auto + show "\a n p. + a \ carrier R \ + a \ \ \ + p \ carrier P \ + deg R p < n \ + to_fun (one_over_poly p) x = x [^] deg R p \ to_fun p (inv x) \ + to_fun (one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = x [^] deg R (p \\<^bsub>P\<^esub> monom P a n) \ to_fun (p \\<^bsub>P\<^esub> monom P a n) (inv x)" + proof- fix a n p assume A: "a \ carrier R" "a \ \" "p \ carrier P" "deg R p < n" + "to_fun (one_over_poly p) x = x [^] deg R p \ to_fun p (inv x)" + have "one_over_poly (p \\<^bsub>P\<^esub> monom P a n) = monom P a 0 \\<^bsub>P\<^esub> monom P \ (n - degree p) \\<^bsub>P\<^esub> one_over_poly p" + using A by (simp add: one_over_poly_monom_add) + hence "to_fun ( one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = + a \ to_fun ( monom P \ (n - degree p) \\<^bsub>P\<^esub> one_over_poly p) x" + using A to_fun_plus one_over_poly_closed cfs_add + by (simp add: assms(2) to_fun_const) + hence "to_fun (one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = a \ x[^](n - degree p) \ x [^] degree p \ to_fun p (inv x)" + by (simp add: A(3) A(5) R.m_assoc assms(2) assms(3) to_fun_closed to_fun_monic_monom to_fun_mult one_over_poly_closed) + hence 0:"to_fun (one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = a \ x[^]n \ to_fun p (inv x)" + using A R.nat_pow_mult assms(2) + by auto + have 1: "to_fun (one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = x[^]n \ (a \ inv x [^]n \ to_fun p (inv x))" + proof- + have "x[^]n \ a \ inv x [^]n = a" + by (metis (no_types, hide_lams) A(1) R.Units_inv_closed R.Units_r_inv R.m_assoc + R.m_comm R.nat_pow_closed R.nat_pow_distrib R.nat_pow_one R.r_one assms(2) assms(3)) + thus ?thesis + using A R.ring_simprules(23)[of _ _ "x[^]n"] 0 R.m_assoc assms(2) assms(3) to_fun_closed + by auto + qed + have 2: "degree (p \\<^bsub>P\<^esub> monom P a n) = n" + by (simp add: A(1) A(2) A(3) A(4) equal_deg_sum) + show " to_fun (one_over_poly (p \\<^bsub>P\<^esub> monom P a n)) x = x [^] deg R (p \\<^bsub>P\<^esub> monom P a n) \ to_fun (p \\<^bsub>P\<^esub> monom P a n) (inv x)" + using 1 2 + by (metis (no_types, lifting) A(1) A(3) P_def R.Units_inv_closed R.add.m_comm + UP_cring.to_fun_monom UP_cring_axioms assms(3) to_fun_closed to_fun_plus monom_closed) + qed +qed + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Lifting Homomorphisms of Rings to Polynomial Rings by Application to Coefficients\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition poly_lift_hom where +"poly_lift_hom R S \ = eval R (UP S) (to_polynomial S \ \) (X_poly S)" + +context UP_ring +begin + +lemma(in UP_cring) pre_poly_lift_hom_is_hom: + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "ring_hom_ring R (UP S) (to_polynomial S \ \)" + apply(rule ring_hom_ringI) + apply (simp add: R.ring_axioms) + apply (simp add: UP_ring.UP_ring UP_ring.intro assms(1) cring.axioms(1)) + using UP_cring.intro UP_cring.to_poly_closed assms(1) assms(2) ring_hom_closed apply fastforce + using assms UP_cring.to_poly_closed[of S] ring_hom_closed[of \ R S] comp_apply[of "to_polynomial S" \] + unfolding UP_cring_def + apply (metis UP_cring.to_poly_mult UP_cring_def ring_hom_mult) + using assms UP_cring.to_poly_closed[of S] ring_hom_closed[of \ R S] comp_apply[of "to_polynomial S" \] + unfolding UP_cring_def + apply (metis UP_cring.to_poly_add UP_cring_def ring_hom_add) + using assms UP_cring.to_poly_closed[of S] ring_hom_one[of \ R S] comp_apply[of "to_polynomial S" \] + unfolding UP_cring_def + by (simp add: \\ \ ring_hom R S \ \ \ = \\<^bsub>S\<^esub>\ UP_cring.intro UP_cring.to_poly_is_ring_hom ring_hom_one) + +lemma(in UP_cring) poly_lift_hom_is_hom: + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "poly_lift_hom R S \ \ ring_hom (UP R) (UP S)" + unfolding poly_lift_hom_def + apply( rule UP_pre_univ_prop.eval_ring_hom[of R "UP S" ]) + unfolding UP_pre_univ_prop_def + apply (simp add: R_cring RingHom.ring_hom_cringI UP_cring.UP_cring UP_cring_def assms(1) assms(2) pre_poly_lift_hom_is_hom) + by (simp add: UP_cring.X_closed UP_cring.intro assms(1)) + +lemma(in UP_cring) poly_lift_hom_closed: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + shows "poly_lift_hom R S \ p \ carrier (UP S)" + by (metis assms(1) assms(2) assms(3) poly_lift_hom_is_hom ring_hom_closed) + +lemma(in UP_cring) poly_lift_hom_add: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + assumes "q \ carrier (UP R)" + shows "poly_lift_hom R S \ (p \\<^bsub>UP R\<^esub> q) = poly_lift_hom R S \ p \\<^bsub>UP S\<^esub> poly_lift_hom R S \ q" + using assms poly_lift_hom_is_hom[of S \] ring_hom_add[of "poly_lift_hom R S \" "UP R" "UP S" p q] + by blast + +lemma(in UP_cring) poly_lift_hom_mult: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + assumes "q \ carrier (UP R)" + shows "poly_lift_hom R S \ (p \\<^bsub>UP R\<^esub> q) = poly_lift_hom R S \ p \\<^bsub>UP S\<^esub> poly_lift_hom R S \ q" + using assms poly_lift_hom_is_hom[of S \] ring_hom_mult[of "poly_lift_hom R S \" "UP R" "UP S" p q] + by blast + +lemma(in UP_cring) poly_lift_hom_extends_hom: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "r \ carrier R" + shows "poly_lift_hom R S \ (to_polynomial R r) = to_polynomial S (\ r)" + using UP_pre_univ_prop.eval_const[of R "UP S" "to_polynomial S \ \" "X_poly S" r ] assms + comp_apply[of "\a. monom (UP S) a 0" \ r] pre_poly_lift_hom_is_hom[of S \] + unfolding poly_lift_hom_def to_polynomial_def UP_pre_univ_prop_def + by (simp add: R_cring RingHom.ring_hom_cringI UP_cring.UP_cring UP_cring.X_closed UP_cring.intro) + +lemma(in UP_cring) poly_lift_hom_extends_hom': + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "r \ carrier R" + shows "poly_lift_hom R S \ (monom P r 0) = monom (UP S) (\ r) 0" + using poly_lift_hom_extends_hom[of S \ r] assms + unfolding to_polynomial_def P_def + by blast + +lemma(in UP_cring) poly_lift_hom_smult: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (a \\<^bsub>UP R\<^esub> p) = \ a \\<^bsub>UP S\<^esub> (poly_lift_hom R S \ p)" + using assms poly_lift_hom_is_hom[of S \] poly_lift_hom_extends_hom'[of S \ a] + poly_lift_hom_mult[of S \ "monom P a 0" p] ring_hom_closed[of \ R S a] + UP_ring.monom_mult_is_smult[of S "\ a" "poly_lift_hom R S \ p"] + monom_mult_is_smult[of a p] monom_closed[of a 0] poly_lift_hom_closed[of S \ p] + unfolding to_polynomial_def UP_ring_def P_def cring_def + by simp + +lemma(in UP_cring) poly_lift_hom_monom: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "r \ carrier R" + shows "poly_lift_hom R S \ (monom (UP R) r n) = (monom (UP S) (\ r) n)" +proof- + have "eval R (UP S) (to_polynomial S \ \) (X_poly S) (monom (UP R) r n) = (to_polynomial S \ \) r \\<^bsub>UP S\<^esub> X_poly S [^]\<^bsub>UP S\<^esub> n" + using assms UP_pre_univ_prop.eval_monom[of R "UP S" "to_polynomial S \ \" r "X_poly S" n] + unfolding UP_pre_univ_prop_def UP_cring_def ring_hom_cring_def + by (meson UP_cring.UP_cring UP_cring.X_closed UP_cring.pre_poly_lift_hom_is_hom UP_cring_axioms + UP_cring_def ring_hom_cring_axioms.intro ring_hom_ring.homh) + then have "eval R (UP S) (to_polynomial S \ \) (X_poly S) (monom (UP R) r n) = (to_polynomial S (\ r)) \\<^bsub>UP S\<^esub> X_poly S [^]\<^bsub>UP S\<^esub> n" + by simp + then show ?thesis + unfolding poly_lift_hom_def + using assms UP_cring.monom_rep_X_pow[of S "\ r" n] ring_hom_closed[of \ R S r] + by (metis UP_cring.X_closed UP_cring.intro UP_cring.monom_sub UP_cring.sub_monom(1)) +qed + +lemma(in UP_cring) poly_lift_hom_X_var: + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "poly_lift_hom R S \ (monom (UP R) \\<^bsub>R\<^esub> 1) = (monom (UP S) \\<^bsub>S\<^esub> 1)" + using assms(1) assms(2) poly_lift_hom_monom ring_hom_one by fastforce + +lemma(in UP_cring) poly_lift_hom_X_var': + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "poly_lift_hom R S \ (X_poly R) = (X_poly S)" + unfolding X_poly_def + using assms(1) assms(2) poly_lift_hom_X_var by blast + +lemma(in UP_cring) poly_lift_hom_X_var'': + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "poly_lift_hom R S \ (monom (UP R) \\<^bsub>R\<^esub> n) = (monom (UP S) \\<^bsub>S\<^esub> n)" + using assms(1) assms(2) poly_lift_hom_monom ring_hom_one by fastforce + +lemma(in UP_cring) poly_lift_hom_X_var''': + assumes "cring S" + assumes "\ \ ring_hom R S" + shows "poly_lift_hom R S \ (X_poly R [^]\<^bsub>UP R\<^esub> (n::nat)) = (X_poly S) [^]\<^bsub>UP S\<^esub> (n::nat)" + using assms + by (smt ltrm_of_X P.nat_pow_closed P_def R.ring_axioms UP_cring.to_fun_closed UP_cring.intro + UP_cring.monom_pow UP_cring.poly_lift_hom_monom UP_cring_axioms X_closed cfs_closed + cring.axioms(1) to_fun_X_pow poly_lift_hom_X_var' ring_hom_closed ring_hom_nat_pow) + +lemma(in UP_cring) poly_lift_hom_X_plus: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (X_poly_plus R a) = X_poly_plus S (\ a)" + using ring_hom_add + unfolding X_poly_plus_def + using P_def X_closed assms(1) assms(2) assms(3) poly_lift_hom_X_var' poly_lift_hom_add poly_lift_hom_extends_hom to_poly_closed by fastforce + +lemma(in UP_cring) poly_lift_hom_X_plus_nat_pow: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (X_poly_plus R a [^]\<^bsub>UP R\<^esub> (n::nat)) = X_poly_plus S (\ a) [^]\<^bsub>UP S\<^esub> (n::nat)" + using assms poly_lift_hom_X_plus[of S \ a] + ring_hom_nat_pow[of "UP R" "UP S" "poly_lift_hom R S \" "X_poly_plus R a" n] + poly_lift_hom_is_hom[of S \] X_plus_closed[of a] UP_ring.UP_ring[of S] + unfolding P_def cring_def UP_cring_def + using P_def UP_ring UP_ring.intro + by (simp add: UP_ring.intro) + +lemma(in UP_cring) X_poly_plus_nat_pow_closed: + assumes "a \ carrier R" + shows " X_poly_plus R a [^]\<^bsub>UP R\<^esub> (n::nat) \ carrier (UP R)" + using assms P.nat_pow_closed P_def X_plus_closed by auto + +lemma(in UP_cring) poly_lift_hom_X_plus_nat_pow_smult: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "poly_lift_hom R S \ (b \\<^bsub>UP R\<^esub> X_poly_plus R a [^]\<^bsub>UP R\<^esub> (n::nat)) = \ b \\<^bsub>UP S \<^esub>X_poly_plus S (\ a) [^]\<^bsub>UP S\<^esub> (n::nat)" + by (simp add: X_poly_plus_nat_pow_closed assms(1) assms(2) assms(3) assms(4) poly_lift_hom_X_plus_nat_pow poly_lift_hom_smult) + +lemma(in UP_cring) poly_lift_hom_X_minus: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (X_poly_minus R a) = X_poly_minus S (\ a)" + using poly_lift_hom_X_plus[of S \ "\ a"] X_minus_plus[of a] UP_cring.X_minus_plus[of S "\ a"] + R.ring_hom_a_inv[of S \ a] + unfolding UP_cring_def P_def + by (metis R.add.inv_closed assms(1) assms(2) assms(3) cring.axioms(1) ring_hom_closed) + +lemma(in UP_cring) poly_lift_hom_X_minus_nat_pow: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (X_poly_minus R a [^]\<^bsub>UP R\<^esub> (n::nat)) = X_poly_minus S (\ a) [^]\<^bsub>UP S\<^esub> (n::nat)" + using assms poly_lift_hom_X_minus ring_hom_nat_pow X_minus_plus UP_cring.X_minus_plus + poly_lift_hom_X_plus poly_lift_hom_X_plus_nat_pow by fastforce + +lemma(in UP_cring) X_poly_minus_nat_pow_closed: + assumes "a \ carrier R" + shows "X_poly_minus R a [^]\<^bsub>UP R\<^esub> (n::nat) \ carrier (UP R)" + using assms monoid.nat_pow_closed[of "UP R" "X_poly_minus R a" n] + P.nat_pow_closed P_def X_minus_closed by auto + +lemma(in UP_cring) poly_lift_hom_X_minus_nat_pow_smult: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "poly_lift_hom R S \ (b \\<^bsub>UP R\<^esub> X_poly_minus R a [^]\<^bsub>UP R\<^esub> (n::nat)) = \ b \\<^bsub>UP S \<^esub>X_poly_minus S (\ a) [^]\<^bsub>UP S\<^esub> (n::nat)" + by (simp add: X_poly_minus_nat_pow_closed assms(1) assms(2) assms(3) assms(4) poly_lift_hom_X_minus_nat_pow poly_lift_hom_smult) + +lemma(in UP_cring) poly_lift_hom_cf: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier P" + shows "poly_lift_hom R S \ p k = \ (p k)" + apply(rule poly_induct3[of p]) + apply (simp add: assms(3)) +proof- + show "\p q. q \ carrier P \ + p \ carrier P \ + poly_lift_hom R S \ p k = \ (p k) \ poly_lift_hom R S \ q k = \ (q k) \ poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q) k = \ ((p \\<^bsub>P\<^esub> q) k)" + proof- fix p q assume A: "p \ carrier P" "q \ carrier P" + "poly_lift_hom R S \ p k = \ (p k)" "poly_lift_hom R S \ q k = \ (q k)" + show "poly_lift_hom R S \ q k = \ (q k) \ poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q) k = \ ((p \\<^bsub>P\<^esub> q) k)" + using A assms poly_lift_hom_add[of S \ p q] + poly_lift_hom_closed[of S \ p] poly_lift_hom_closed[of S \ q] + UP_ring.cfs_closed[of S "poly_lift_hom R S \ q " k] UP_ring.cfs_closed[of S "poly_lift_hom R S \ p" k] + UP_ring.cfs_add[of S "poly_lift_hom R S \ p" "poly_lift_hom R S \ q" k] + unfolding P_def UP_ring_def + by (metis (full_types) P_def cfs_add cfs_closed cring.axioms(1) ring_hom_add) + qed + show "\a n. a \ carrier R \ poly_lift_hom R S \ (monom P a n) k = \ (monom P a n k)" + proof- fix a m assume A: "a \ carrier R" + show "poly_lift_hom R S \ (monom P a m) k = \ (monom P a m k)" + apply(cases "m = k") + using cfs_monom[of a m k] assms poly_lift_hom_monom[of S \ a m] UP_ring.cfs_monom[of S "\ a" m k] + unfolding P_def UP_ring_def + apply (simp add: A cring.axioms(1) ring_hom_closed) + using cfs_monom[of a m k] assms poly_lift_hom_monom[of S \ a m] UP_ring.cfs_monom[of S "\ a" m k] + unfolding P_def UP_ring_def + by (metis A P_def R.ring_axioms cring.axioms(1) ring_hom_closed ring_hom_zero) + qed +qed + +lemma(in ring) ring_hom_monom_term: + assumes "a \ carrier R" + assumes "c \ carrier R" + assumes "ring S" + assumes "h \ ring_hom R S" + shows "h (a \ c[^](n::nat)) = h a \\<^bsub>S\<^esub> (h c)[^]\<^bsub>S\<^esub>n" + apply(induction n) + using assms ringE(2) ring_hom_closed apply fastforce + by (metis assms(1) assms(2) assms(3) assms(4) local.ring_axioms nat_pow_closed ring_hom_mult ring_hom_nat_pow) + +lemma(in UP_cring) poly_lift_hom_eval: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "UP_cring.to_fun S (poly_lift_hom R S \ p) (\ a) = \ (to_fun p a) " + apply(rule poly_induct3[of p]) + apply (simp add: assms(3)) +proof- + show "\p q. q \ carrier P \ + p \ carrier P \ + UP_cring.to_fun S (poly_lift_hom R S \ p) (\ a) = \ (to_fun p a) \ + UP_cring.to_fun S (poly_lift_hom R S \ q) (\ a) = \ (to_fun q a) \ + UP_cring.to_fun S (poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q)) (\ a) = \ (to_fun (p \\<^bsub>P\<^esub> q) a)" + proof- fix p q assume A: "q \ carrier P" "p \ carrier P" + "UP_cring.to_fun S (poly_lift_hom R S \ p) (\ a) = \ (to_fun p a)" + "UP_cring.to_fun S (poly_lift_hom R S \ q) (\ a) = \ (to_fun q a)" + have "(poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q)) = poly_lift_hom R S \ p \\<^bsub>UP S\<^esub> poly_lift_hom R S \ q" + using A(1) A(2) P_def assms(1) assms(2) poly_lift_hom_add by auto + hence "UP_cring.to_fun S (poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q)) (\ a) = + UP_cring.to_fun S (poly_lift_hom R S \ p) (\ a) \\<^bsub>S\<^esub> UP_cring.to_fun S (poly_lift_hom R S \ q) (\ a)" + using UP_cring.to_fun_plus[of S] assms + unfolding UP_cring_def + by (metis (no_types, lifting) A(1) A(2) P_def poly_lift_hom_closed ring_hom_closed) + thus "UP_cring.to_fun S (poly_lift_hom R S \ (p \\<^bsub>P\<^esub> q)) (\ a) = \ (to_fun (p \\<^bsub>P\<^esub> q) a)" + using A to_fun_plus assms ring_hom_add[of \ R S] + poly_lift_hom_closed[of S \] UP_cring.to_fun_def[of S] to_fun_def + unfolding P_def UP_cring_def + using UP_cring.to_fun_closed UP_cring_axioms + by metis + qed + show "\c n. c \ carrier R \ UP_cring.to_fun S (poly_lift_hom R S \ (monom P c n)) (\ a) = \ (to_fun (monom P c n) a)" + unfolding P_def + proof - fix c n assume A: "c \ carrier R" + have 0: "\ (a [^]\<^bsub>R\<^esub> (n::nat)) = \ a [^]\<^bsub>S\<^esub> n" + using assms ring_hom_nat_pow[of R S \ a n] + unfolding cring_def + using R.ring_axioms by blast + have 1: "\ (c \\<^bsub>R\<^esub> a [^]\<^bsub>R\<^esub> n) = \ c \\<^bsub>S\<^esub> \ a [^]\<^bsub>S\<^esub> n" + using ring_hom_mult[of \ R S c "a [^]\<^bsub>R\<^esub> n" ] 0 assms A monoid.nat_pow_closed [of R a n] + by (simp add: cring.axioms(1) ringE(2)) + show "UP_cring.to_fun S (poly_lift_hom R S \ (monom (UP R) c n)) (\ a) = \ (to_fun(monom (UP R) c n) a)" + using assms A poly_lift_hom_monom[of S \ c n] UP_cring.to_fun_monom[of S "\ c" "\ a" n] + to_fun_monom[of c a n] 0 1 ring_hom_closed[of \ R S] unfolding UP_cring_def + by (simp add: P_def to_fun_def) + qed +qed + +lemma(in UP_cring) poly_lift_hom_sub: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier P" + assumes "q \ carrier P" + shows "poly_lift_hom R S \ (compose R p q) = compose S (poly_lift_hom R S \ p) (poly_lift_hom R S \ q)" + apply(rule poly_induct3[of p]) + apply (simp add: assms(3)) +proof- + show " \p qa. + qa \ carrier P \ + p \ carrier P \ + poly_lift_hom R S \ (Cring_Poly.compose R p q) = Cring_Poly.compose S (poly_lift_hom R S \ p) (poly_lift_hom R S \ q) \ + poly_lift_hom R S \ (Cring_Poly.compose R qa q) = Cring_Poly.compose S (poly_lift_hom R S \ qa) (poly_lift_hom R S \ q) \ + poly_lift_hom R S \ (Cring_Poly.compose R (p \\<^bsub>P\<^esub> qa) q) = Cring_Poly.compose S (poly_lift_hom R S \ (p \\<^bsub>P\<^esub> qa)) (poly_lift_hom R S \ q)" + proof- fix a b assume A: "a \ carrier P" + "b \ carrier P" + "poly_lift_hom R S \ (Cring_Poly.compose R a q) = Cring_Poly.compose S (poly_lift_hom R S \ a) (poly_lift_hom R S \ q)" + "poly_lift_hom R S \ (Cring_Poly.compose R b q) = Cring_Poly.compose S (poly_lift_hom R S \ b) (poly_lift_hom R S \ q)" + show "poly_lift_hom R S \ (Cring_Poly.compose R (a \\<^bsub>P\<^esub> b) q) = Cring_Poly.compose S (poly_lift_hom R S \ (a \\<^bsub>P\<^esub> b)) (poly_lift_hom R S \ q)" + using assms UP_cring.sub_add[of R q a b ] UP_cring.sub_add[of S ] + unfolding UP_cring_def + by (metis A(1) A(2) A(3) A(4) P_def R_cring UP_cring.sub_closed UP_cring_axioms poly_lift_hom_add poly_lift_hom_closed) + qed + show "\a n. a \ carrier R \ + poly_lift_hom R S \ (Cring_Poly.compose R (monom P a n) q) = + Cring_Poly.compose S (poly_lift_hom R S \ (monom P a n)) (poly_lift_hom R S \ q)" + proof- + fix a n assume A: "a \ carrier R" + have 0: "(poly_lift_hom R S \ (monom (UP R) a n)) = monom (UP S) (\ a) n" + by (simp add: A assms(1) assms(2) assms(3) assms(4) poly_lift_hom_monom) + have 1: " q [^]\<^bsub>UP R\<^esub> n \ carrier (UP R)" + using monoid.nat_pow_closed[of "UP R" q n] UP_ring.UP_ring UP_ring.intro assms(1) assms + P.monoid_axioms P_def by blast + have 2: "poly_lift_hom R S \ (to_polynomial R a \\<^bsub>UP R\<^esub> q [^]\<^bsub>UP R\<^esub> n) = + to_polynomial S (\ a) \\<^bsub>UP S\<^esub> (poly_lift_hom R S \ q) [^]\<^bsub>UP S\<^esub> n" + using poly_lift_hom_mult[of S \ "to_polynomial R a" "q [^]\<^bsub>UP R\<^esub> n"] poly_lift_hom_is_hom[of S \] + ring_hom_nat_pow[of P "UP S" "poly_lift_hom R S \" q n] UP_cring.UP_cring[of S] + UP_cring poly_lift_hom_monom[of S \ a 0] ring_hom_closed[of \ R S a] + monom_closed[of a 0] nat_pow_closed[of q n] assms A + unfolding to_polynomial_def P_def UP_cring_def cring_def + by auto + have 3: "poly_lift_hom R S \ (Cring_Poly.compose R (monom (UP R) a n) q) = to_polynomial S (\ a) \\<^bsub>UP S\<^esub> (poly_lift_hom R S \ q) [^]\<^bsub>UP S\<^esub> n" + using "2" A P_def assms(4) sub_monom(1) by auto + have 4: "Cring_Poly.compose S (poly_lift_hom R S \ (monom (UP R) a n)) (poly_lift_hom R S \ q) + = Cring_Poly.compose S (monom (UP S) (\ a) n) (poly_lift_hom R S \ q)" + by (simp add: "0") + have "poly_lift_hom R S \ q \ carrier (UP S)" + using P_def UP_cring.poly_lift_hom_closed UP_cring_axioms assms(1) assms(2) assms(4) by blast + then have 5: "Cring_Poly.compose S (poly_lift_hom R S \ (monom (UP R) a n)) (poly_lift_hom R S \ q) + = to_polynomial S (\ a) \\<^bsub>UP S\<^esub> (poly_lift_hom R S \ q) [^]\<^bsub>UP S\<^esub> n" + using 4 UP_cring.sub_monom[of S "poly_lift_hom R S \ q" "\ a" n] assms + unfolding UP_cring_def + by (simp add: A ring_hom_closed) + thus "poly_lift_hom R S \ (Cring_Poly.compose R (monom P a n) q) = + Cring_Poly.compose S (poly_lift_hom R S \ (monom P a n)) (poly_lift_hom R S \ q)" + using 0 1 2 3 4 assms A + by (simp add: P_def) + qed +qed + +lemma(in UP_cring) poly_lift_hom_comm_taylor_expansion: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (taylor_expansion R a p) = taylor_expansion S (\ a) (poly_lift_hom R S \ p)" + unfolding taylor_expansion_def + using poly_lift_hom_sub[of S \ p "(X_poly_plus R a)"] poly_lift_hom_X_plus[of S \ a] assms + by (simp add: P_def UP_cring.X_plus_closed UP_cring_axioms) + +lemma(in UP_cring) poly_lift_hom_comm_taylor_expansion_cf: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + assumes "a \ carrier R" + shows "\ (taylor_expansion R a p i) = taylor_expansion S (\ a) (poly_lift_hom R S \ p) i" + using poly_lift_hom_cf assms poly_lift_hom_comm_taylor_expansion P_def + taylor_def UP_cring.taylor_closed UP_cring_axioms by fastforce + +lemma(in UP_cring) taylor_expansion_cf_closed: + assumes "p \ carrier P" + assumes "a \ carrier R" + shows "taylor_expansion R a p i \ carrier R" + using assms taylor_closed + by (simp add: taylor_def cfs_closed) + +lemma(in UP_cring) poly_lift_hom_comm_taylor_term: + assumes "cring S" + assumes "\ \ ring_hom R S" + assumes "p \ carrier (UP R)" + assumes "a \ carrier R" + shows "poly_lift_hom R S \ (taylor_term a p i) = UP_cring.taylor_term S (\ a) (poly_lift_hom R S \ p) i" + using poly_lift_hom_X_minus_nat_pow_smult[of S \ a "taylor_expansion R a p i" i] + poly_lift_hom_comm_taylor_expansion[of S \ p a] + poly_lift_hom_comm_taylor_expansion_cf[of S \ p a i] + assms UP_cring.taylor_term_def[of S] + unfolding taylor_term_def UP_cring_def P_def + by (simp add: UP_cring.taylor_expansion_cf_closed UP_cring_axioms) + +lemma(in UP_cring) poly_lift_hom_degree_bound: + assumes "cring S" + assumes "h \ ring_hom R S" + assumes "f \ carrier (UP R)" + shows "deg S (poly_lift_hom R S h f) \ deg R f" + using poly_lift_hom_closed[of S h f] UP_cring.deg_leqI[of S "poly_lift_hom R S h f" "deg R f"] assms ring_hom_zero[of h R S] deg_aboveD[of f] coeff_simp[of f] + unfolding P_def UP_cring_def + by (simp add: P_def R.ring_axioms cring.axioms(1) poly_lift_hom_cf) + +lemma(in UP_cring) deg_eqI: + assumes "f \ carrier (UP R)" + assumes "deg R f \ n" + assumes "f n \ \" + shows "deg R f = n" + using assms coeff_simp[of f] P_def deg_leE le_neq_implies_less by blast + +lemma(in UP_cring) poly_lift_hom_degree_eq: + assumes "cring S" + assumes "h \ ring_hom R S" + assumes "h (lcf f) \ \\<^bsub>S\<^esub>" + assumes "f \ carrier (UP R)" + shows "deg S (poly_lift_hom R S h f) = deg R f" + apply(rule UP_cring.deg_eqI) + using assms unfolding UP_cring_def apply blast + using poly_lift_hom_closed[of S h f] assms apply blast + using poly_lift_hom_degree_bound[of S h f] assms apply blast + using assms poly_lift_hom_cf[of S h f] + by (metis P_def) + +lemma(in UP_cring) poly_lift_hom_lcoeff: + assumes "cring S" + assumes "h \ ring_hom R S" + assumes "h (lcf f) \ \\<^bsub>S\<^esub>" + assumes "f \ carrier (UP R)" + shows "UP_ring.lcf S (poly_lift_hom R S h f) = h (lcf f)" + using poly_lift_hom_degree_eq[of S h f] assms + by (simp add: P_def poly_lift_hom_cf) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Coefficient List Constructor for Polynomials\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition list_to_poly where +"list_to_poly R as n = (if n < length as then as!n else \\<^bsub>R\<^esub>)" + +context UP_ring +begin + +lemma(in UP_ring) list_to_poly_closed: + assumes "set as \ carrier R" + shows "list_to_poly R as \ carrier P" + apply(rule UP_car_memI[of "length as"]) + apply (simp add: list_to_poly_def) + by (metis R.zero_closed assms in_mono list_to_poly_def nth_mem) + +lemma(in UP_ring) list_to_poly_zero[simp]: +"list_to_poly R [] = \\<^bsub>UP R\<^esub>" + unfolding list_to_poly_def + apply auto + by(simp add: UP_def) + +lemma(in UP_domain) list_to_poly_singleton: + assumes "a \ carrier R" + shows "list_to_poly R [a] = monom P a 0" + apply(rule ext) + unfolding list_to_poly_def using assms + by (simp add: cfs_monom) +end + +definition cf_list where +"cf_list R p = map p [(0::nat)..< Suc (deg R p)]" + +lemma cf_list_length: +"length (cf_list R p) = Suc (deg R p)" + unfolding cf_list_def + by simp + +lemma cf_list_entries: + assumes "i \ deg R p" + shows "(cf_list R p)!i = p i" + unfolding cf_list_def + by (metis add.left_neutral assms diff_zero less_Suc_eq_le map_eq_map_tailrec nth_map_upt) + +lemma(in UP_ring) list_to_poly_cf_list_inv: + assumes "p \ carrier (UP R)" + shows "list_to_poly R (cf_list R p) = p" +proof + fix x + show "list_to_poly R (cf_list R p) x = p x" + apply(cases "x < degree p") + unfolding list_to_poly_def + using assms cf_list_length[of R p] cf_list_entries[of _ R p] + apply simp + by (metis P_def UP_ring.coeff_simp UP_ring_axioms \\i. i \ deg R p \ cf_list R p ! i = p i\ \length (cf_list R p) = Suc (deg R p)\ assms deg_belowI less_Suc_eq_le) +qed + +section\Polynomial Rings over a Subring\ + +subsection\Characterizing the Carrier of a Polynomial Ring over a Subring\ +lemma(in ring) carrier_update: +"carrier (R\carrier := S\) = S" +"\\<^bsub>(R\carrier := S\)\<^esub> = \" +"\\<^bsub>(R\carrier := S\)\<^esub> = \" +"(\\<^bsub>(R\carrier := S\)\<^esub>) = (\)" +"(\\<^bsub>(R\carrier := S\)\<^esub>) = (\)" +by auto + + +lemma(in UP_cring) poly_cfs_subring: + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "\n. g n \ S" + shows "g \ carrier (UP (R \ carrier := S \))" + apply(rule UP_cring.UP_car_memI') + using R.subcringI' R.subcring_iff UP_cring.intro assms(1) subringE(1) apply blast +proof- + have "carrier (R\carrier := S\) = S" + using ring.carrier_update by simp + then show 0: "\x. g x \ carrier (R\carrier := S\)" + using assms by blast + have 0: "\\<^bsub>R\carrier := S\\<^esub> = \" + using R.carrier_update(2) by blast + then show "\x. (deg R g) < x \ g x = \\<^bsub>R\carrier := S\\<^esub>" + using UP_car_memE assms(2) by presburger +qed + +lemma(in UP_cring) UP_ring_subring: + assumes "subring S R" + shows "UP_cring (R \ carrier := S \)" "UP_ring (R \ carrier := S \)" + using assms unfolding UP_cring_def + using R.subcringI' R.subcring_iff subringE(1) apply blast + using assms unfolding UP_ring_def + using R.subcringI' R.subcring_iff subringE(1) + by (simp add: R.subring_is_ring) + +lemma(in UP_cring) UP_ring_subring_is_ring: + assumes "subring S R" + shows "cring (UP (R \ carrier := S \))" + using assms UP_ring_subring[of S] UP_cring.UP_cring[of "R\carrier := S\"] + by blast + +lemma(in UP_cring) UP_ring_subring_add_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "f \ carrier (UP (R \ carrier := S \))" + shows "f \\<^bsub>UP (R \ carrier := S \)\<^esub>g \ carrier (UP (R \ carrier := S \))" + using assms UP_ring_subring_is_ring[of S] + by (meson cring.cring_simprules(1)) + +lemma(in UP_cring) UP_ring_subring_mult_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "f \ carrier (UP (R \ carrier := S \))" + shows "f \\<^bsub>UP (R \ carrier := S \)\<^esub>g \ carrier (UP (R \ carrier := S \))" + using assms UP_ring_subring_is_ring[of S] + by (meson cring.carrier_is_subcring subcringE(6)) + +lemma(in UP_cring) UP_ring_subring_car: + assumes "subring S R" + shows "carrier (UP (R \ carrier := S \)) = {h \ carrier (UP R). \n. h n \ S}" +proof + show "carrier (UP (R\carrier := S\)) \ {h \ carrier (UP R). \n. h n \ S}" + proof + fix h assume A: "h \ carrier (UP (R\carrier := S\))" + have "h \ carrier P" + apply(rule UP_car_memI[of "deg (R\carrier := S\) h"]) unfolding P_def + using UP_cring.UP_car_memE[of "R\carrier := S\" h] R.carrier_update[of S] + assms UP_ring_subring A apply presburger + using UP_cring.UP_car_memE[of "R\carrier := S\" h] assms + by (metis A R.ring_axioms UP_cring_def \carrier (R\carrier := S\) = S\ cring.subcringI' is_UP_cring ring.subcring_iff subringE(1) subsetD) + then show "h \ {h \ carrier (UP R). \n. h n \ S}" + unfolding P_def + using assms A UP_cring.UP_car_memE[of "R\carrier := S\" h] R.carrier_update[of S] + UP_ring_subring by blast + qed + show "{h \ carrier (UP R). \n. h n \ S} \ carrier (UP (R\carrier := S\))" + proof fix h assume A: "h \ {h \ carrier (UP R). \n. h n \ S}" + have 0: "h \ carrier (UP R)" + using A by blast + have 1: "\n. h n \ S" + using A by blast + show "h \ carrier (UP (R\carrier := S\))" + apply(rule UP_ring.UP_car_memI[of _ "deg R h"]) + using assms UP_ring_subring[of S] UP_cring.axioms UP_ring.intro cring.axioms(1) apply blast + using UP_car_memE[of h] carrier_update 0 R.carrier_update(2) apply presburger + using assms 1 R.carrier_update(1) by blast + qed +qed + +lemma(in UP_cring) UP_ring_subring_car_subset: + assumes "subring S R" + shows "carrier (UP (R \ carrier := S \)) \ carrier (UP R)" +proof fix h assume "h \ carrier (UP (R \ carrier := S \))" + then show "h \ carrier (UP R)" + using assms UP_ring_subring_car[of S] by blast +qed + +lemma(in UP_cring) UP_ring_subring_car_subset': + assumes "subring S R" + assumes "h \ carrier (UP (R \ carrier := S \))" + shows "h \ carrier (UP R)" + using assms UP_ring_subring_car_subset[of S] by blast + +lemma(in UP_cring) UP_ring_subring_add: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "f \ carrier (UP (R \ carrier := S \))" + shows "g \\<^bsub>UP R\<^esub> f = g \\<^bsub>UP (R \ carrier := S \)\<^esub>f" +proof(rule ext) fix x show "(g \\<^bsub>UP R\<^esub> f) x = (g \\<^bsub>UP (R\carrier := S\)\<^esub> f) x" + proof- + have 0: " (g \\<^bsub>P\<^esub> f) x = g x \ f x" + using assms cfs_add[of g f x] unfolding P_def + using UP_ring_subring_car_subset' by blast + have 1: "(g \\<^bsub>UP (R\carrier := S\)\<^esub> f) x = g x \\<^bsub>R\carrier := S\\<^esub> f x" + using UP_ring.cfs_add[of "R \ carrier := S \" g f x] UP_ring_subring[of S] assms + unfolding UP_ring_def UP_cring_def + using R.subring_is_ring by blast + show ?thesis using 0 1 R.carrier_update(4)[of S] + by (simp add: P_def) + qed +qed + +lemma(in UP_cring) UP_ring_subring_deg: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + shows "deg R g = deg (R \ carrier := S \) g" +proof- + have 0: "g \ carrier (UP R)" + using assms UP_ring_subring_car[of S] by blast + have 1: "deg R g \ deg (R \ carrier := S \) g" + using 0 assms UP_cring.UP_car_memE[of "R \ carrier := S \" g] + UP_car_memE[of g] P_def R.carrier_update(2) UP_ring_subring deg_leqI by presburger + have 2: "deg (R \ carrier := S \) g \ deg R g" + using 0 assms UP_cring.UP_car_memE[of "R \ carrier := S \" g] + UP_car_memE[of g] P_def R.carrier_update(2) UP_ring_subring UP_cring.deg_leqI + by metis + + show ?thesis using 1 2 by presburger +qed + + +lemma(in UP_cring) UP_subring_monom: + assumes "subring S R" + assumes "a \ S" + shows "up_ring.monom (UP R) a n = up_ring.monom (UP (R \ carrier := S \)) a n" +proof fix x + have 0: "a \ carrier R" + using assms subringE(1) by blast + have 1: "a \ carrier (R\carrier := S\)" + using assms by (simp add: assms(2)) + have 2: " up_ring.monom (UP (R\carrier := S\)) a n x = (if n = x then a else \\<^bsub>R\carrier := S\\<^esub>)" + using 1 assms UP_ring_subring[of S] UP_ring.cfs_monom[of "R\carrier := S\" a n x] UP_cring.axioms UP_ring.intro cring.axioms(1) + by blast + show "up_ring.monom (UP R) a n x = up_ring.monom (UP (R\carrier := S\)) a n x" + using 0 1 2 cfs_monom[of a n x] R.carrier_update(2)[of S] unfolding P_def by presburger +qed + +lemma(in UP_cring) UP_ring_subring_mult: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "f \ carrier (UP (R \ carrier := S \))" + shows "g \\<^bsub>UP R\<^esub> f = g \\<^bsub>UP (R \ carrier := S \)\<^esub>f" +proof(rule UP_ring.poly_induct3[of "R \ carrier := S \" f]) + show "UP_ring (R\carrier := S\)" + by (simp add: UP_ring_subring(2) assms(1)) + show " f \ carrier (UP (R\carrier := S\))" + by (simp add: assms(3)) + show " \p q. q \ carrier (UP (R\carrier := S\)) \ + p \ carrier (UP (R\carrier := S\)) \ + g \\<^bsub>UP R\<^esub> p = g \\<^bsub>UP (R\carrier := S\)\<^esub> p \ + g \\<^bsub>UP R\<^esub> q = g \\<^bsub>UP (R\carrier := S\)\<^esub> q \ g \\<^bsub>UP R\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = g \\<^bsub>UP (R\carrier := S\)\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q)" + proof- fix p q + assume A: " q \ carrier (UP (R\carrier := S\))" + "p \ carrier (UP (R\carrier := S\))" + "g \\<^bsub>UP R\<^esub> p = g \\<^bsub>UP (R\carrier := S\)\<^esub> p" + "g \\<^bsub>UP R\<^esub> q = g \\<^bsub>UP (R\carrier := S\)\<^esub> q" + have 0: "p \\<^bsub>UP (R\carrier := S\)\<^esub> q = p \\<^bsub>UP R\<^esub> q" + using A UP_ring_subring_add[of S p q] + by (simp add: assms(1)) + have 1: "g \\<^bsub>UP R\<^esub> (p \\<^bsub>UP R\<^esub> q) = g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q" + using 0 A assms P.r_distr P_def UP_ring_subring_car_subset' by auto + hence 2:"g \\<^bsub>UP R\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q" + using 0 by simp + have 3: "g \\<^bsub>UP (R\carrier := S\)\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = + g \\<^bsub>UP (R\carrier := S\)\<^esub> p \\<^bsub>UP (R\carrier := S\)\<^esub> g \\<^bsub>UP (R\carrier := S\)\<^esub> q" + using 0 A assms semiring.r_distr[of "UP (R\carrier := S\)"] UP_ring_subring_car_subset' + using UP_ring.UP_r_distr \UP_ring (R\carrier := S\)\ by blast + hence 4: "g \\<^bsub>UP (R\carrier := S\)\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = + g \\<^bsub>UP R\<^esub> p \\<^bsub>UP (R\carrier := S\)\<^esub> g \\<^bsub>UP R\<^esub> q" + using A by simp + hence 5: "g \\<^bsub>UP (R\carrier := S\)\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = + g \\<^bsub>UP R\<^esub> p \\<^bsub>UP R\<^esub> g \\<^bsub>UP R\<^esub> q" + using UP_ring_subring_add[of S] + by (simp add: A(1) A(2) A(3) A(4) UP_ring.UP_mult_closed \UP_ring (R\carrier := S\)\ assms(1) assms(2)) + show "g \\<^bsub>UP R\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) = g \\<^bsub>UP (R\carrier := S\)\<^esub> (p \\<^bsub>UP (R\carrier := S\)\<^esub> q)" + by (simp add: "2" "5") + qed + show "\a n. a \ carrier (R\carrier := S\) \ g \\<^bsub>UP R\<^esub> monom (UP (R\carrier := S\)) a n = g \\<^bsub>UP (R\carrier := S\)\<^esub> monom (UP (R\carrier := S\)) a n" + proof fix a n x assume A: "a \ carrier (R\carrier := S\)" + have 0: "monom (UP (R\carrier := S\)) a n = monom (UP R) a n" + using A UP_subring_monom assms(1) by auto + have 1: "g \ carrier (UP R)" + using assms UP_ring_subring_car_subset' by blast + have 2: "a \ carrier R" + using A assms subringE(1)[of S R] R.carrier_update[of S] by blast + show "(g \\<^bsub>UP R\<^esub> monom (UP (R\carrier := S\)) a n) x = (g \\<^bsub>UP (R\carrier := S\)\<^esub> monom (UP (R\carrier := S\)) a n) x" + proof(cases "x < n") + case True + have T0: "(g \\<^bsub>UP R\<^esub> monom (UP R) a n) x = \" + using 1 2 True cfs_monom_mult[of g a x n] A assms unfolding P_def by blast + then show ?thesis using UP_cring.cfs_monom_mult[of "R\carrier := S\" g a x n] 0 A True + UP_ring_subring(1) assms(1) assms(2) by auto + next + case False + have F0: "(g \\<^bsub>UP R\<^esub> monom (UP R) a n) x = a \ (g (x - n))" + using 1 2 False cfs_monom_mult_l[of g a n "x - n"] A assms unfolding P_def by simp + have F1: "(g \\<^bsub>UP (R\carrier := S\)\<^esub> monom (UP (R\carrier := S\)) a n) (x - n + n) = a \\<^bsub>R\carrier := S\\<^esub> g (x - n)" + using 1 2 False UP_cring.cfs_monom_mult_l[of "R\carrier := S\" g a n "x - n"] A assms + UP_ring_subring(1) by blast + hence F2: "(g \\<^bsub>UP (R\carrier := S\)\<^esub> monom (UP R) a n) (x - n + n) = a \ g (x - n)" + using UP_subring_monom[of S a n] R.carrier_update[of S] assms 0 by metis + show ?thesis using F0 F1 1 2 assms + by (simp add: "0" False add.commute add_diff_inverse_nat) + qed + qed +qed + +lemma(in UP_cring) UP_ring_subring_one: + assumes "subring S R" + shows "\\<^bsub>UP R\<^esub> = \\<^bsub>UP (R \ carrier := S \)\<^esub>" + using UP_subring_monom[of S \ 0] assms P_def R.subcringI' UP_ring.monom_one UP_ring_subring(2) monom_one subcringE(3) by force + +lemma(in UP_cring) UP_ring_subring_zero: + assumes "subring S R" + shows "\\<^bsub>UP R\<^esub> = \\<^bsub>UP (R \ carrier := S \)\<^esub>" + using UP_subring_monom[of S \ 0] UP_ring.monom_zero[of "R \ carrier := S \" 0] assms monom_zero[of 0] + UP_ring_subring[of S] subringE(2)[of S R] + unfolding P_def + by (simp add: P_def R.carrier_update(2)) + +lemma(in UP_cring) UP_ring_subring_nat_pow: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + shows "g[^]\<^bsub>UP R\<^esub>n = g[^]\<^bsub>UP (R \ carrier := S \)\<^esub>(n::nat)" + apply(induction n) + using assms apply (simp add: UP_ring_subring_one) +proof- + fix n::nat + assume A: "g [^]\<^bsub>UP R\<^esub> n = g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n" + have "Group.monoid (UP (R\carrier := S\)) " + using assms UP_ring_subring[of S] UP_ring.UP_ring[of "R\carrier := S\"] ring.is_monoid by blast + hence 0 : " g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n \ carrier (UP (R\carrier := S\))" + using monoid.nat_pow_closed[of "UP (R \ carrier := S \)" g n] assms UP_ring_subring + unfolding UP_ring_def ring_def by blast + have 1: "g [^]\<^bsub>UP R\<^esub> n \ carrier (UP R)" + using 0 assms UP_ring_subring_car_subset'[of S] by (simp add: A) + then have 2: "g [^]\<^bsub>UP R\<^esub> n \\<^bsub>UP R\<^esub> g = g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n \\<^bsub>UP (R\carrier := S\)\<^esub> g" + using assms UP_ring_subring_mult[of S "g [^]\<^bsub>UP R\<^esub> n" g] + by (simp add: "0" A) + then show "g [^]\<^bsub>UP R\<^esub> Suc n = g [^]\<^bsub>UP (R\carrier := S\)\<^esub> Suc n" + by simp +qed + +lemma(in UP_cring) UP_subring_compose_monom: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "compose R (up_ring.monom (UP R) a n) g = compose (R \ carrier := S \) (up_ring.monom (UP (R \ carrier := S \)) a n) g" +proof- + have g_closed: "g \ carrier (UP R)" + using assms UP_ring_subring_car by blast + have 0: "a \ carrier R" + using assms subringE(1) by blast + have 1: "compose R (up_ring.monom (UP R) a n) g = a \\<^bsub>UP R\<^esub> (g[^]\<^bsub>UP R\<^esub>n)" + using monom_sub[of a g n] unfolding P_def + using "0" assms(2) g_closed by blast + have 2: "compose (R\carrier := S\) (up_ring.monom (UP (R\carrier := S\)) a n) g = a \\<^bsub>UP (R\carrier := S\)\<^esub> g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n" + using assms UP_cring.monom_sub[of "R \ carrier := S \" a g n] UP_ring_subring[of S] R.carrier_update[of S] + by blast + have 3: " g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n = g[^]\<^bsub>UP R\<^esub>n" + using UP_ring_subring_nat_pow[of S g n] + by (simp add: assms(1) assms(2)) + have 4: "a \\<^bsub>UP R\<^esub> (g[^]\<^bsub>UP R\<^esub>n) = a \\<^bsub>UP (R\carrier := S\)\<^esub> g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n" + proof fix x + show "(a \\<^bsub>UP R\<^esub> g [^]\<^bsub>UP R\<^esub> n) x = (a \\<^bsub>UP (R\carrier := S\)\<^esub> g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n) x" + proof- + have LHS: "(a \\<^bsub>UP R\<^esub> g [^]\<^bsub>UP R\<^esub> n) x = a \ ((g [^]\<^bsub>UP R\<^esub> n) x)" + using "0" P.nat_pow_closed P_def cfs_smult g_closed by auto + have RHS: "(a \\<^bsub>UP (R\carrier := S\)\<^esub> g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n) x = a \\<^bsub>R\carrier := S\\<^esub> ((g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n) x)" + proof- + have "Group.monoid (UP (R\carrier := S\)) " + using assms UP_ring_subring[of S] UP_ring.UP_ring[of "R\carrier := S\"] ring.is_monoid by blast + hence 0 : " g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n \ carrier (UP (R\carrier := S\))" + using monoid.nat_pow_closed[of "UP (R \ carrier := S \)" g n] assms UP_ring_subring + unfolding UP_ring_def ring_def by blast + have 1: "g [^]\<^bsub>UP (R\carrier := S\)\<^esub> n \ carrier (UP (R\carrier := S\))" + using assms UP_ring_subring[of S] R.carrier_update[of S] 0 by blast + then show ?thesis using UP_ring.cfs_smult UP_ring_subring assms + by (simp add: UP_ring.cfs_smult) + qed + show ?thesis using R.carrier_update RHS LHS 3 assms + by simp + qed + qed + show ?thesis using 0 1 2 3 4 + by simp +qed + +lemma(in UP_cring) UP_subring_compose: + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "f \ carrier (UP R)" + assumes "\n. g n \ S" + assumes "\n. f n \ S" + shows "compose R f g = compose (R \ carrier := S \) f g" +proof- + have g_closed: "g \ carrier (UP (R \ carrier := S \))" + using assms poly_cfs_subring by blast + have 0: "\n. (\ h. h \ carrier (UP R) \ deg R h \ n \ h \ carrier (UP (R \ carrier := S \)) \ compose R h g = compose (R \ carrier := S \) h g)" + proof- fix n show "(\ h. h \ carrier (UP R) \ deg R h \ n \ h \ carrier (UP (R \ carrier := S \)) \ compose R h g = compose (R \ carrier := S \) h g)" + proof(induction n) + show "\h. h \ carrier (UP R) \ deg R h \ 0 \ h \ carrier (UP (R\carrier := S\)) \ Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + proof fix h + show "h \ carrier (UP R) \ deg R h \ 0 \ h \ carrier (UP (R\carrier := S\)) \ Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + proof + assume A: "h \ carrier (UP R) \ deg R h \ 0 \ h \ carrier (UP (R\carrier := S\))" + then have 0: "deg R h = 0" + by linarith + then have 1: "deg (R \ carrier := S \) h = 0" + using A assms UP_ring_subring_deg[of S h] + by linarith + show "Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + using 0 1 g_closed assms sub_const[of g h] UP_cring.sub_const[of "R\carrier := S\" g h] A P_def UP_ring_subring + by presburger + qed + qed + show "\n. \h. h \ carrier (UP R) \ deg R h \ n \ h \ carrier (UP (R\carrier := S\)) \ + Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g \ + \h. h \ carrier (UP R) \ deg R h \ Suc n \ h \ carrier (UP (R\carrier := S\)) \ + Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + proof fix n h + assume IH: "\h. h \ carrier (UP R) \ deg R h \ n \ h \ carrier (UP (R\carrier := S\)) \ + Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + show "h \ carrier (UP R) \ deg R h \ Suc n \ h \ carrier (UP (R\carrier := S\)) \ + Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + proof assume A: "h \ carrier (UP R) \ deg R h \ Suc n \ h \ carrier (UP (R\carrier := S\))" + show "Cring_Poly.compose R h g = Cring_Poly.compose (R\carrier := S\) h g" + proof(cases "deg R h \ n") + case True + then show ?thesis using A IH by blast + next + case False + then have F0: "deg R h = Suc n" + using A by (simp add: A le_Suc_eq) + then have F1: "deg (R\carrier := S\) h = Suc n" + using UP_ring_subring_deg[of S h] A + by (simp add: \h \ carrier (UP R) \ deg R h \ Suc n \ h \ carrier (UP (R\carrier := S\))\ assms(1)) + obtain j where j_def: "j \ carrier (UP (R\carrier := S\)) \ + h = j \\<^bsub>UP (R\carrier := S\)\<^esub> up_ring.monom (UP (R\carrier := S\)) (h (deg (R\carrier := S\) h)) (deg (R\carrier := S\) h) \ + deg (R\carrier := S\) j < deg (R\carrier := S\) h" + using A UP_ring.ltrm_decomp[of "R\carrier := S\" h] assms UP_ring_subring[of S] + F1 by (metis (mono_tags, lifting) F0 False zero_less_Suc) + have j_closed: "j \ carrier (UP R)" + using j_def assms UP_ring_subring_car_subset by blast + have F2: "deg R j < deg R h" + using j_def assms + by (metis (no_types, lifting) F0 F1 UP_ring_subring_deg) + have F3: "(deg (R\carrier := S\) h) = deg R h" + by (simp add: F0 F1) + have F30: "h (deg (R\carrier := S\) h) \ S " + using A UP_cring.UP_car_memE[of "R\carrier := S\" h "deg (R\carrier := S\) h"] + by (metis R.carrier_update(1) UP_ring_subring(1) assms(1)) + hence F4: "up_ring.monom P (h (deg R h)) (deg R h) = + up_ring.monom (UP (R\carrier := S\)) (h (deg (R\carrier := S\) h)) (deg (R\carrier := S\) h)" + using F3 g_closed j_def UP_subring_monom[of S "h (deg (R\carrier := S\) h)"] assms + unfolding P_def by metis + have F5: "compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g = + compose (R \ carrier := S \) (up_ring.monom (UP (R \ carrier := S \)) (h (deg (R\carrier := S\) h)) (deg (R\carrier := S\) h)) g" + using F0 F1 F2 F3 F4 UP_subring_compose_monom[of S] assms P_def \h (deg (R\carrier := S\) h) \ S\ + by (metis g_closed) + have F5: "compose R j g = compose (R \ carrier := S \) j g" + using F0 F2 IH UP_ring_subring_car_subset' assms(1) j_def by auto + have F6: "h = j \\<^bsub>UP R\<^esub> monom (UP R) (h (deg R h)) (deg R h)" + using j_def F4 UP_ring_subring_add[of S j "up_ring.monom (UP (R\carrier := S\)) (h (deg (R\carrier := S\) h)) (deg (R\carrier := S\) h)"] + UP_ring.monom_closed[of "R\carrier := S\" "h (deg (R\carrier := S\) h)" "deg (R\carrier := S\) h"] + using P_def UP_ring_subring(2) \h (deg (R\carrier := S\) h) \ S\ assms(1) by auto + have F7: "compose R h g =compose R j g \\<^bsub>UP R\<^esub> + compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g" + proof- + show ?thesis + using assms(2) j_closed F5 sub_add[of g j "up_ring.monom P (h (deg R h)) (deg R h)" ] + F4 F3 F2 F1 g_closed unfolding P_def + by (metis A F6 ltrm_closed P_def) + qed + have F8: "compose (R \ carrier := S \) h g = compose (R \ carrier := S \) j g \\<^bsub>UP (R \ carrier := S \)\<^esub> + compose (R \ carrier := S \) (up_ring.monom (UP (R \ carrier := S \)) (h (deg (R \ carrier := S \) h)) (deg (R \ carrier := S \) h)) g" + proof- + have 0: " UP_cring (R\carrier := S\)" + by (simp add: UP_ring_subring(1) assms(1)) + have 1: "monom (UP (R\carrier := S\)) (h (deg R h)) (deg R h) \ carrier (UP (R\carrier := S\))" + using assms 0 F30 UP_ring.monom_closed[of "R\carrier := S\" "h (deg R h)" "deg R h"] R.carrier_update[of S] + unfolding UP_ring_def UP_cring_def + by (simp add: F3 cring.axioms(1)) + show ?thesis + using 0 1 g_closed j_def UP_cring.sub_add[of "R \ carrier := S \" g j "monom (UP (R\carrier := S\)) (h (deg R h)) (deg R h)" ] + using F3 by auto + qed + have F9: "compose R j g \ carrier (UP R)" + by (simp add: UP_cring.sub_closed assms(2) is_UP_cring j_closed) + have F10: "compose (R \ carrier := S \) j g \ carrier (UP (R \ carrier := S \))" + using assms j_def UP_cring.sub_closed[of "R \ carrier := S \"] UP_ring_subring(1) g_closed by blast + have F11: " compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g \ carrier (UP R)" + using assms j_def UP_cring.sub_closed[of "R \ carrier := S \"] + UP_ring.monom_closed[of "R \ carrier := S \"] + by (simp add: A UP_car_memE(1) UP_cring.rev_sub_closed UP_ring.monom_closed is_UP_cring is_UP_ring sub_rev_sub) + have F12: " compose (R \ carrier := S \) (up_ring.monom (UP (R \ carrier := S \)) (h (deg (R \ carrier := S \) h)) (deg (R \ carrier := S \) h)) g \ carrier (UP (R \ carrier := S \))" + using assms j_def UP_cring.sub_closed[of "R \ carrier := S \"] + UP_ring.monom_closed[of "R \ carrier := S \"] UP_ring_subring[of S] + using A UP_ring.ltrm_closed g_closed by fastforce + show ?thesis using F9 F10 F11 F12 F7 F8 F5 UP_ring_subring_add[of S "compose R j g" "compose R (up_ring.monom (UP R) (h (deg R h)) (deg R h)) g"] + assms + using F3 F30 UP_subring_compose_monom g_closed by auto + qed + qed + qed + qed + qed + show ?thesis using 0[of "deg R f"] + by (simp add: assms(1) assms(3) assms(5) poly_cfs_subring) +qed + + +subsection\Evaluation over a Subring\ + +lemma(in UP_cring) UP_subring_eval: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "to_function R g a = to_function (R \ carrier := S \) g a" + apply(rule UP_ring.poly_induct3[of "R \ carrier := S \" g] ) + apply (simp add: UP_ring_subring(2) assms(1)) + apply (simp add: assms(2)) +proof- + show "\p q. q \ carrier (UP (R\carrier := S\)) \ + p \ carrier (UP (R\carrier := S\)) \ + to_function R p a = to_function (R\carrier := S\) p a \ + to_function R q a = to_function (R\carrier := S\) q a \ + to_function R (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) a = to_function (R\carrier := S\) (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) a" + proof- fix p q assume A: "q \ carrier (UP (R\carrier := S\))" + "p \ carrier (UP (R\carrier := S\))" + " to_function R p a = to_function (R\carrier := S\) p a" + " to_function R q a = to_function (R\carrier := S\) q a" + have a_closed: "a \ carrier R" + using assms R.carrier_update[of S] subringE(1) by blast + have 0: "UP_cring (R\carrier := S\)" + using assms by (simp add: UP_ring_subring(1)) + have 1: "to_function (R\carrier := S\) p a \ S" + using A 0 UP_cring.to_fun_closed[of "R\carrier := S\"] + by (simp add: UP_cring.to_fun_def assms(3)) + have 2: "to_function (R\carrier := S\) q a \ S" + using A 0 UP_cring.to_fun_closed[of "R\carrier := S\"] + by (simp add: UP_cring.to_fun_def assms(3)) + have 3: "p \ carrier (UP R)" + using A assms 0 UP_ring_subring_car_subset' by blast + have 4: "q \ carrier (UP R)" + using A assms 0 UP_ring_subring_car_subset' by blast + have 5: "to_fun p a \ to_fun q a = UP_cring.to_fun (R\carrier := S\) p a \\<^bsub>R\carrier := S\\<^esub> UP_cring.to_fun (R\carrier := S\) q a" + using 1 2 A R.carrier_update[of S] assms by (simp add: "0" UP_cring.to_fun_def to_fun_def) + have 6: "UP_cring.to_fun (R\carrier := S\) (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) a = + UP_cring.to_fun (R\carrier := S\) p a \\<^bsub>R\carrier := S\\<^esub> UP_cring.to_fun (R\carrier := S\) q a" + using UP_cring.to_fun_plus[of "R \ carrier := S \" q p a] + by (simp add: "0" A(1) A(2) assms(3)) + have 7: "to_fun (p \\<^bsub>P\<^esub> q) a = to_fun p a \ to_fun q a" + using to_fun_plus[of q p a] 3 4 a_closed by (simp add: P_def) + have 8: "p \\<^bsub>UP (R\carrier := S\)\<^esub> q = p \\<^bsub>P\<^esub> q" + unfolding P_def using assms A R.carrier_update[of S] UP_ring_subring_add[of S p q] by simp + show "to_function R (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) a = to_function (R\carrier := S\) (p \\<^bsub>UP (R\carrier := S\)\<^esub> q) a" + using UP_ring_subring_car_subset'[of S ] 0 1 2 3 4 5 6 7 8 A R.carrier_update[of S] + unfolding P_def by (simp add: UP_cring.to_fun_def to_fun_def) + qed + show "\b n. + b \ carrier (R\carrier := S\) \ + to_function R (monom (UP (R\carrier := S\)) b n) a = to_function (R\carrier := S\) (monom (UP (R\carrier := S\)) b n) a" + proof- fix b n assume A: "b \ carrier (R\carrier := S\)" + have 0: "UP_cring (R\carrier := S\)" + by (simp add: UP_ring_subring(1) assms(1)) + have a_closed: "a \ carrier R" + using assms subringE by blast + have 1: "UP_cring.to_fun (R\carrier := S\) (monom (UP (R\carrier := S\)) b n) a = b \\<^bsub>R\carrier := S\\<^esub> a [^]\<^bsub>R\carrier := S\\<^esub> n" + using assms A UP_cring.to_fun_monom[of "R\carrier := S\" b a n] + by (simp add: "0") + have 2: "UP_cring.to_fun (R\carrier := S\) (monom (UP (R\carrier := S\)) b n) \ to_function (R\carrier := S\) (monom (UP (R\carrier := S\)) b n)" + using UP_cring.to_fun_def[of "R\carrier := S\" "monom (UP (R\carrier := S\)) b n"] 0 by linarith + have 3: "(monom (UP (R\carrier := S\)) b n) = monom P b n" + using A assms unfolding P_def using UP_subring_monom by auto + have 4: " b \ a [^] n = b \\<^bsub>R\carrier := S\\<^esub> a [^]\<^bsub>R\carrier := S\\<^esub> n" + apply(induction n) using R.carrier_update[of S] apply simp + using R.carrier_update[of S] R.nat_pow_consistent by auto + hence 5: "to_function R (monom (UP (R\carrier := S\)) b n) a = b \\<^bsub>R\carrier := S\\<^esub> a[^]\<^bsub>R\carrier := S\\<^esub>n" + using 0 1 2 3 assms A UP_cring.to_fun_monom[of "R\carrier := S\" b a n] UP_cring.to_fun_def[of "R\carrier := S\" "monom (UP (R\carrier := S\)) b n"] + R.carrier_update[of S] subringE[of S R] a_closed UP_ring.monom_closed[of "R\carrier := S\" a n] + to_fun_monom[of b a n] + unfolding P_def UP_cring.to_fun_def to_fun_def by (metis subsetD) + thus " to_function R (monom (UP (R\carrier := S\)) b n) a = to_function (R\carrier := S\) (monom (UP (R\carrier := S\)) b n) a" + using "1" "2" by auto + qed +qed + +lemma(in UP_cring) UP_subring_eval': + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "to_fun g a = to_function (R \ carrier := S \) g a" + unfolding to_fun_def using assms + by (simp add: UP_subring_eval) + +lemma(in UP_cring) UP_subring_eval_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "to_fun g a \ S" + using assms UP_subring_eval'[of S g a] UP_cring.to_fun_closed UP_cring.to_fun_def R.carrier_update(1) UP_ring_subring(1) by fastforce + +subsection\Derivatives and Taylor Expansions over a Subring\ + + +lemma(in UP_cring) UP_subring_taylor: + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "\n. g n \ S" + assumes "a \ S" + shows "taylor_expansion R a g = taylor_expansion (R \ carrier := S \) a g" +proof- + have a_closed: "a \ carrier R" + using assms subringE by blast + have 0: "X_plus a \ carrier (UP R)" + using assms X_plus_closed unfolding P_def + using local.a_closed by auto + have 1: "\n. X_plus a n \ S" + proof- fix n + have "X_plus a n = (if n = 0 then a else + (if n = 1 then \ else \))" + using a_closed + by (simp add: cfs_X_plus) + then show "X_plus a n \ S" using subringE assms + by (simp add: subringE(2) subringE(3)) + qed + have 2: "(X_poly_plus (R\carrier := S\) a) = X_plus a" + proof- + have 20: "(X_poly_plus (R\carrier := S\) a) = (\k. if k = (0::nat) then a else + (if k = 1 then \ else \))" + using a_closed assms UP_cring.cfs_X_plus[of "R\carrier := S\" a] R.carrier_update + UP_ring_subring(1) by auto + have 21: "X_plus a = (\k. if k = (0::nat) then a else + (if k = 1 then \ else \))" + using cfs_X_plus[of a] a_closed + by blast + show ?thesis apply(rule ext) using 20 21 + by auto + qed + show ?thesis + unfolding taylor_expansion_def using 0 1 2 assms UP_subring_compose[of S g "X_plus a"] + by (simp add: UP_subring_compose) +qed + +lemma(in UP_cring) UP_subring_taylor_closed: + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "\n. g n \ S" + assumes "a \ S" + shows "taylor_expansion R a g \ carrier (UP (R \ carrier := S \))" +proof- + have "g \ carrier (UP (R\carrier := S\))" + by (metis P_def R.carrier_update(1) R.carrier_update(2) UP_cring.UP_car_memI' UP_ring_subring(1) assms(1) assms(2) assms(3) deg_leE) + then show ?thesis + using assms UP_cring.taylor_def[of "R\carrier := S\"] UP_subring_taylor[of S g a] + UP_cring.taylor_closed[of "R \ carrier := S \" g a] UP_ring_subring(1)[of S] by simp +qed + +lemma(in UP_cring) UP_subring_taylor_closed': + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "taylor_expansion R a g \ carrier (UP (R \ carrier := S \))" + using UP_subring_taylor_closed assms UP_cring.UP_car_memE[of "R \ carrier := S \" g] R.carrier_update[of S] + UP_ring_subring(1) UP_ring_subring_car_subset' by auto + +lemma(in UP_cring) UP_subring_taylor': + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "\n. g n \ S" + assumes "a \ S" + shows "taylor_expansion R a g n \ S" + using assms UP_subring_taylor R.carrier_update[of S] UP_cring.taylor_closed[of "R \ carrier := S \"] + using UP_cring.taylor_expansion_cf_closed UP_ring_subring(1) poly_cfs_subring by metis + + +lemma(in UP_cring) UP_subring_deriv: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "deriv g a= UP_cring.deriv (R \ carrier := S \) g a" +proof- + have 0: "(\n. g n \ S)" + using assms UP_ring_subring_car by blast + thus ?thesis + unfolding derivative_def using 0 UP_ring_subring_car_subset[of S] assms UP_subring_taylor[of S g a] + by (simp add: subset_iff) +qed + +lemma(in UP_cring) UP_subring_deriv_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + shows "deriv g a \ S" + using assms UP_cring.deriv_closed[of "R \ carrier := S \" g a] UP_subring_deriv[of S g a] + UP_ring_subring_car_subset[of S] UP_ring_subring[of S] + by simp + +lemma(in UP_cring) poly_shift_subring_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + shows "poly_shift g \ carrier (UP (R \ carrier := S \))" + using UP_cring.poly_shift_closed[of "R \ carrier := S \" g] assms UP_ring_subring[of S] + by simp + +lemma(in UP_cring) UP_subring_taylor_appr: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + assumes "b \ S" + shows "\c \ S. to_fun g a= to_fun g b \ (deriv g b)\ (a \ b) \ (c \ (a \ b)[^](2::nat))" +proof- + have a_closed: "a \ carrier R" + using assms subringE by blast + have b_closed: "b \ carrier R" + using assms subringE by blast + have g_closed: " g \ carrier (UP R)" + using UP_ring_subring_car_subset[of S] assms by blast + have 0: "to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b) = to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b)" + by simp + have 1: "to_fun g b = to_fun g b" + by simp + have 2: "deriv g b = deriv g b" + by simp + have 3: "to_fun g a = to_fun g b \ deriv g b \ (a \ b) \ to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b) \ (a \ b) [^] (2::nat)" + using taylor_deg_1_expansion[of g b a "to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b)" "to_fun g b" "deriv g b"] + assms a_closed b_closed g_closed 0 1 2 unfolding P_def by blast + have 4: "to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b) \ S" + proof- + have 0: "(2::nat) = Suc (Suc 0)" + by simp + have 1: "a \ b \ S" + using assms unfolding a_minus_def + by (simp add: subringE(5) subringE(7)) + have 2: "poly_shift (T\<^bsub>b\<^esub> g) \ carrier (UP (R\carrier := S\))" + using poly_shift_subring_closed[of S "taylor_expansion R b g"] UP_ring_subring[of S] + UP_subring_taylor_closed'[of S g b] assms unfolding taylor_def + by blast + hence 3: "poly_shift (poly_shift (T\<^bsub>b\<^esub> g)) \ carrier (UP (R\carrier := S\))" + using UP_cring.poly_shift_closed[of "R\carrier := S\" "(poly_shift (T\<^bsub>b\<^esub> g))"] + unfolding taylor_def + using assms(1) poly_shift_subring_closed by blast + have 4: "to_fun (poly_shift (poly_shift (T\<^bsub>b\<^esub> g))) (a \ b) \ S" + using 1 2 3 0 UP_subring_eval_closed[of S "poly_shift (poly_shift (T\<^bsub>b\<^esub> g))" "a \ b"] + UP_cring.poly_shift_closed[of "R\carrier := S\"] assms + by blast + then show ?thesis + by (simp add: numeral_2_eq_2) + qed + obtain c where c_def: "c = to_fun (shift 2 (T\<^bsub>b\<^esub> g)) (a \ b)" + by blast + have 5: "c \ S \ to_fun g a = to_fun g b \ deriv g b \ (a \ b) \ c \ (a \ b) [^] (2::nat)" + unfolding c_def using 3 4 by blast + thus ?thesis using c_def 4 by blast +qed + +lemma(in UP_cring) UP_subring_taylor_appr': + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + assumes "a \ S" + assumes "b \ S" + shows "\c c' c''. c \ S \ c' \ S \ c'' \ S \ to_fun g a= c \ c'\ (a \ b) \ (c'' \ (a \ b)[^](2::nat))" + using UP_subring_taylor_appr[of S g a b] assms UP_subring_deriv_closed[of S g b] UP_subring_eval_closed[of S g b] + by blast + +lemma (in UP_cring) pderiv_cfs: + assumes"g \ carrier (UP R)" + shows "pderiv g n = [Suc n]\(g (Suc n))" + unfolding pderiv_def + using n_mult_closed[of g] assms poly_shift_cfs[of "n_mult g" n] + unfolding P_def n_mult_def by blast + +lemma(in ring) subring_add_pow: + assumes "subring S R" + assumes "a \ S" + shows "[(n::nat)] \\<^bsub>R\carrier := S\\<^esub> a = [(n::nat)] \a" +proof- + have 0: "a \ carrier R" + using assms(1) assms(2) subringE(1) by blast + have 1: "a \ carrier (R\carrier := S\)" + by (simp add: assms(2)) + show ?thesis + apply(induction n) + using assms 0 1 carrier_update[of S] + apply (simp add: add_pow_def) + using assms 0 1 carrier_update[of S] + by (simp add: add_pow_def) +qed + +lemma(in UP_cring) UP_subring_pderiv_equal: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + shows "pderiv g = UP_cring.pderiv (R\carrier := S\) g" +proof fix n + show "pderiv g n = UP_cring.pderiv (R\carrier := S\) g n" + using UP_cring.pderiv_cfs[of "R \ carrier := S \" g n] pderiv_cfs[of g n] + assms R.subring_add_pow[of S "g (Suc n)" "Suc n"] + by (simp add: UP_ring_subring(1) UP_ring_subring_car) +qed + +lemma(in UP_cring) UP_subring_pderiv_closed: + assumes "subring S R" + assumes "g \ carrier (UP (R \ carrier := S \))" + shows "pderiv g \ carrier (UP (R \ carrier := S \))" + using assms UP_cring.pderiv_closed[of "R \ carrier := S \" g] R.carrier_update(1) UP_ring_subring(1) +UP_subring_pderiv_equal by auto + +lemma(in UP_cring) UP_subring_pderiv_closed': + assumes "subring S R" + assumes "g \ carrier (UP R)" + assumes "\n. g n \ S" + shows "\n. pderiv g n \ S" + using assms UP_subring_pderiv_closed[of S g] poly_cfs_subring[of S g] UP_ring_subring_car + by blast + +lemma(in UP_cring) taylor_deg_one_expansion_subring: + assumes "f \ carrier (UP R)" + assumes "subring S R" + assumes "\i. f i \ S" + assumes "a \ S" + assumes "b \ S" + shows "\c \ S. to_fun f b = (to_fun f a) \ (deriv f a) \ (b \ a) \ (c \ (b \ a)[^](2::nat))" + apply(rule UP_subring_taylor_appr, rule assms) + using assms poly_cfs_subring apply blast + by(rule assms, rule assms) + +lemma(in UP_cring) taylor_deg_one_expansion_subring': + assumes "f \ carrier (UP R)" + assumes "subring S R" + assumes "\i. f i \ S" + assumes "a \ S" + assumes "b \ S" + shows "\c \ S. to_fun f b = (to_fun f a) \ (to_fun (pderiv f) a) \ (b \ a) \ (c \ (b \ a)[^](2::nat))" +proof- + have "S \ carrier R" + using assms subringE(1) by blast + hence 0: "deriv f a = to_fun (pderiv f) a" + using assms pderiv_eval_deriv[of f a] unfolding P_def by blast + show ?thesis + using assms taylor_deg_one_expansion_subring[of f S a b] + unfolding 0 by blast +qed + +end diff --git a/thys/Padic_Ints/Extended_Int.thy b/thys/Padic_Ints/Extended_Int.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Extended_Int.thy @@ -0,0 +1,931 @@ + +section \Extended integers (i.e. with infinity)\ + +text\ + This section formalizes the extended integers, which serve as the codomain for the $p$-adic + valuation. The element $\infty$ is added to the integers to serve as a maximxal element in the + order, which is the valuation of $0$. +\ + +theory Extended_Int +imports Main "HOL-Library.Countable" "HOL-Library.Order_Continuity" "HOL-Library.Extended_Nat" +begin + +text\ + The following is based very closely on the theory \<^theory>\HOL-Library.Extended_Nat\ from the + standard Isabelle distribution, with adaptations made to formalize the integers extended with + an element for infinity. This is the standard range for the valuation function on a discretely + valued ring such as the field of $p$-adic numbers, such as in \cite{engler2005valued}. +\ + +context + fixes f :: "nat \ 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}" +begin + +lemma sums_SUP[simp, intro]: "f sums (SUP n. \iiType definition\ + +text \ + We extend the standard natural numbers by a special value indicating + infinity. +\ + +typedef eint = "UNIV :: int option set" .. + + +definition eint :: "int \ eint" where + "eint n = Abs_eint (Some n)" + +instantiation eint :: infinity +begin + +definition "\ = Abs_eint None" +instance .. + +end + +fun int_option_enumeration :: "int option \ nat" where +"int_option_enumeration (Some n) = (if n \ 0 then nat (2*(n + 1)) else nat (2*(-n) + 1))"| +"int_option_enumeration None = (0::nat)" + +lemma int_option_enumeration_inj: +"inj int_option_enumeration" +proof + have pos_even: "\n::int. n \ 0 \ even (int_option_enumeration (Some n)) \ (int_option_enumeration (Some n))> 0" + proof- + fix n::int assume "n \0" then have "(int_option_enumeration (Some n)) = nat (2*(n + 1))" + by simp + then show "even (int_option_enumeration (Some n)) \ 0 < int_option_enumeration (Some n)" + by (smt \0 \ n\ even_of_nat int_nat_eq oddE zero_less_nat_eq) + qed + have neg_odd: "\n::int. n < 0 \ odd (int_option_enumeration (Some n))" + by (smt evenE even_of_nat int_nat_eq int_option_enumeration.simps(1)) + fix x y assume A: "x \ UNIV" "y \ UNIV" "int_option_enumeration x = int_option_enumeration y" + show "x = y" + apply(cases "x = None") + using A pos_even neg_odd + apply (metis dvd_0_right int_option_enumeration.elims int_option_enumeration.simps(2) not_gr0 not_le) + proof- + assume "x \None" + then obtain n where n_def: "x = Some n" + by auto + then have y_not_None: "y \ None" + using A + by (metis \\thesis. (\n. x = Some n \ thesis) \ thesis\ + add_cancel_right_right even_add int_option_enumeration.simps(2) + linorder_not_less neg_odd neq0_conv pos_even) + then obtain m where m_def: "y = Some m" + by blast + show ?thesis + proof(cases "n \0") + case True + then show ?thesis + using n_def A neg_odd pos_even m_def int_option_enumeration.simps(1) + by (smt int_nat_eq) + next + case False + then show ?thesis + using n_def A neg_odd pos_even m_def int_option_enumeration.simps(1) + by (smt int_nat_eq) + qed + qed +qed + +definition eint_enumeration where +"eint_enumeration = int_option_enumeration \ Rep_eint" + +lemma eint_enumeration_inj: +"inj eint_enumeration" + unfolding eint_enumeration_def + using int_option_enumeration_inj Rep_eint_inject + by (metis (mono_tags, lifting) comp_apply injD inj_on_def) + +instance eint :: countable +proof + show "\to_int::eint \ nat. inj to_int" + using eint_enumeration_inj by blast +qed + +old_rep_datatype eint "\ :: eint" +proof - + fix P i assume "\j. P (eint j)" "P \" + then show "P i" + proof induct + case (Abs_eint y) then show ?case + by (cases y rule: option.exhaust) + (auto simp: eint_def infinity_eint_def) + qed +qed (auto simp add: eint_def infinity_eint_def Abs_eint_inject) + +declare [[coercion "eint::int\eint"]] + +lemmas eint2_cases = eint.exhaust[case_product eint.exhaust] +lemmas eint3_cases = eint.exhaust[case_product eint.exhaust eint.exhaust] + +lemma not_infinity_eq [iff]: "(x \ \) = (\i. x = eint i)" + by (cases x) auto + +lemma not_eint_eq [iff]: "(\y. x \ eint y) = (x = \)" + by (cases x) auto + +lemma eint_ex_split: "(\c::eint. P c) \ P \ \ (\c::int. P c)" + by (metis eint.exhaust) + +primrec the_eint :: "eint \ int" + where "the_eint (eint n) = n" + + +subsection \Constructors and numbers\ + +instantiation eint :: zero_neq_one +begin + +definition + "0 = eint 0" + +definition + "1 = eint 1" + +instance + proof qed (simp add: zero_eint_def one_eint_def) + +end + + +lemma eint_0 [code_post]: "eint 0 = 0" + by (simp add: zero_eint_def) + +lemma eint_1 [code_post]: "eint 1 = 1" + by (simp add: one_eint_def) + +lemma eint_0_iff: "eint x = 0 \ x = 0" "0 = eint x \ x = 0" + by (auto simp add: zero_eint_def) + +lemma eint_1_iff: "eint x = 1 \ x = 1" "1 = eint x \ x = 1" + by (auto simp add: one_eint_def) + +lemma infinity_ne_i0 [simp]: "(\::eint) \ 0" + by (simp add: zero_eint_def) + +lemma i0_ne_infinity [simp]: "0 \ (\::eint)" + by (simp add: zero_eint_def) + +lemma zero_one_eint_neq: + "\ 0 = (1::eint)" + "\ 1 = (0::eint)" + unfolding zero_eint_def one_eint_def by simp_all + +lemma infinity_ne_i1 [simp]: "(\::eint) \ 1" + by (simp add: one_eint_def) + +lemma i1_ne_infinity [simp]: "1 \ (\::eint)" + by (simp add: one_eint_def) + +subsection \Addition\ + +instantiation eint :: comm_monoid_add +begin + +definition [nitpick_simp]: + "m + n = (case m of \ \ \ | eint m \ (case n of \ \ \ | eint n \ eint (m + n)))" + +lemma plus_eint_simps [simp, code]: + fixes q :: eint + shows "eint m + eint n = eint (m + n)" + and "\ + q = \" + and "q + \ = \" + by (simp_all add: plus_eint_def split: eint.splits) + +instance +proof + fix n m q :: eint + show "n + m + q = n + (m + q)" + by (cases n m q rule: eint3_cases) auto + show "n + m = m + n" + by (cases n m rule: eint2_cases) auto + show "0 + n = n" + by (cases n) (simp_all add: zero_eint_def) +qed + +end + +lemma eSuc_eint: "(eint n) + 1 = eint (n + 1)" + by (simp add: one_eint_def) + +lemma eSuc_infinity [simp]: " \ + (1::eint) = \" + unfolding plus_eint_def + by auto + +lemma eSuc_inject [simp]: " m + (1::eint)= n + 1 \ m = n" + unfolding plus_eint_def + apply(cases "m = \") + apply (metis (no_types, lifting) eSuc_eint + eint.distinct(2) eint.exhaust eint.simps(4) eint.simps(5) plus_eint_def) + apply(cases "n = \") + using eSuc_eint plus_eint_def apply auto[1] + unfolding one_eint_def + using add.commute eSuc_eint + by auto + +lemma eSuc_eint_iff: "x + 1 = eint y \ (\n. y = n + 1 \ x = eint n)" + apply(cases "x = \") + apply simp + unfolding plus_eint_def one_eint_def + using eSuc_eint + by auto + +lemma enat_eSuc_iff: "eint y = x + 1 \ (\n. y = n + 1 \ eint n = x)" + using eSuc_eint_iff + by metis + +lemma iadd_Suc: "((m::eint) + 1) + n = (m + n) + 1" + by (metis ab_semigroup_add_class.add_ac(1) add.assoc add.commute) + + +lemma iadd_Suc_right: "(m::eint) + (n + 1) = (m + n) + 1" + using add.assoc[of m n 1] by auto + +subsection \Multiplication\ + +instantiation eint :: "{comm_semiring}" +begin + +definition times_eint_def [nitpick_simp]: + "m * n = (case m of \ \ \ | eint m \ + (case n of \ \ \ | eint n \ eint (m * n)))" + +lemma times_eint_simps [simp, code]: + "eint m * eint n = eint (m * n)" + "\ * \ = (\::eint)" + "\ * eint n = \" + "eint m * \ = \" + unfolding times_eint_def zero_eint_def + by (simp_all split: eint.split) + +lemma sum_infinity_imp_summand_infinity: + assumes "a + b = (\::eint)" + shows "a = \ \ b = \" + using assms + by (metis not_eint_eq plus_eint_simps(1)) + +lemma sum_finite_imp_summands_finite: + assumes "a + b \ (\::eint)" + shows "a \ \" "b \ \" + using assms eint.simps(5) apply fastforce + using assms eint.simps(5) by fastforce + + +instance +proof + fix a b c :: eint + show "(a * b) * c = a * (b * c)" + unfolding times_eint_def zero_eint_def + by (simp split: eint.split) + show comm: "a * b = b * a" + unfolding times_eint_def zero_eint_def + by (simp split: eint.split) + show distr: "(a + b) * c = a * c + b * c" + unfolding times_eint_def plus_eint_def + apply(cases "a + b = \") + apply(cases "a = \") + apply simp + using sum_infinity_imp_summand_infinity[of a b] + apply (metis eint.simps(5) plus_eint_def plus_eint_simps(3)) + apply(cases "c = \") + apply (metis eint.exhaust plus_eint_def plus_eint_simps(3) times_eint_def times_eint_simps(4)) + using sum_finite_imp_summands_finite[of a b] + apply auto + by (simp add: semiring_normalization_rules(1)) +qed + + +end + +lemma mult_one_right[simp]: +"(n::eint)*1 = n" + apply(cases "n = \") + apply (simp add: one_eint_def) + by (metis eint2_cases mult_cancel_left2 one_eint_def times_eint_simps(1)) + +lemma mult_one_left[simp]: +"1*(n::eint) = n" + by (metis mult.commute mult_one_right) + +lemma mult_eSuc: "((m::eint) + 1) * n = m * n + n" + by (simp add: distrib_right) + +lemma mult_eSuc': "((m::eint) + 1) * n = n + m * n" + using mult_eSuc add.commute by simp + +lemma mult_eSuc_right: "(m::eint) * (n + 1) = m * n + m " + by(simp add: distrib_left) + +lemma mult_eSuc_right': "(m::eint) * (n + 1) = m + m * n " + using mult_eSuc_right add.commute by simp + + +subsection \Numerals\ + +lemma numeral_eq_eint: + "numeral k = eint (numeral k)" + by simp + +lemma eint_numeral [code_abbrev]: + "eint (numeral k) = numeral k" + using numeral_eq_eint .. + +lemma infinity_ne_numeral [simp]: "(\::eint) \ numeral k" + by auto + +lemma numeral_ne_infinity [simp]: "numeral k \ (\::eint)" + by simp + +subsection \Subtraction\ + +instantiation eint :: minus +begin + +definition diff_eint_def: +"a - b = (case a of (eint x) \ (case b of (eint y) \ eint (x - y) | \ \ \) + | \ \ \)" + +instance .. + +end + + + + +lemma idiff_eint_eint [simp, code]: "eint a - eint b = eint (a - b)" + by (simp add: diff_eint_def) + +lemma idiff_infinity [simp, code]: "\ - n = (\::eint)" + by (simp add: diff_eint_def) + +lemma idiff_infinity_right [simp, code]: "eint a - \ = \" + by (simp add: diff_eint_def) + +lemma idiff_0 [simp]: "(0::eint) - n = -n" + by (cases n, simp_all add: zero_eint_def) + +lemmas idiff_eint_0 [simp] = idiff_0 [unfolded zero_eint_def] + +lemma idiff_0_right [simp]: "(n::eint) - 0 = n" + by (cases n) (simp_all add: zero_eint_def) + +lemmas idiff_eint_0_right [simp] = idiff_0_right [unfolded zero_eint_def] + +lemma idiff_self [simp]: "n \ \ \ (n::eint) - n = 0" + by (auto simp: zero_eint_def) + + +lemma eSuc_minus_eSuc [simp]: "((n::eint) + 1) - (m + 1) = n - m" + apply(cases "n = \") + apply simp + apply(cases "m = \") + apply (metis eSuc_infinity eint.exhaust idiff_infinity_right infinity_ne_i1 sum_infinity_imp_summand_infinity) +proof- + assume A: "n \\" "m \ \" + obtain a where a_def: "n = eint a" + using A + by auto + obtain b where b_def: "m = eint b" + using A + by auto + show ?thesis + using idiff_eint_eint[of "a + 1" "b + 1"] + idiff_eint_eint[of a b] + by (simp add: a_def b_def eSuc_eint) +qed + +lemma eSuc_minus_1 [simp]: "((n::eint)+ 1) - 1 = n" + using eSuc_minus_eSuc[of n 0] + by auto + + +(*lemmas idiff_self_eq_0_eint = idiff_self_eq_0[unfolded zero_eint_def]*) + +subsection \Ordering\ + +instantiation eint :: linordered_ab_semigroup_add +begin + +definition [nitpick_simp]: + "m \ n = (case n of eint n1 \ (case m of eint m1 \ m1 \ n1 | \ \ False) + | \ \ True)" + +definition [nitpick_simp]: + "m < n = (case m of eint m1 \ (case n of eint n1 \ m1 < n1 | \ \ True) + | \ \ False)" + +lemma eint_ord_simps [simp]: + "eint m \ eint n \ m \ n" + "eint m < eint n \ m < n" + "q \ (\::eint)" + "q < (\::eint) \ q \ \" + "(\::eint) \ q \ q = \" + "(\::eint) < q \ False" + by (simp_all add: less_eq_eint_def less_eint_def split: eint.splits) + +lemma numeral_le_eint_iff[simp]: + shows "numeral m \ eint n \ numeral m \ n" + by auto + +lemma numeral_less_eint_iff[simp]: + shows "numeral m < eint n \ numeral m < n" + by simp + +lemma eint_ord_code [code]: + "eint m \ eint n \ m \ n" + "eint m < eint n \ m < n" + "q \ (\::eint) \ True" + "eint m < \ \ True" + "\ \ eint n \ False" + "(\::eint) < q \ False" + by simp_all + +lemma eint_ord_plus_one[simp]: + assumes "eint n \ x" + assumes "x < y" + shows "eint (n + 1) \ y" +proof- + obtain m where "x = eint m" + using assms(2) + by fastforce + show ?thesis apply(cases "y = \") + apply simp + using \x = eint m\ assms(1) assms(2) + by force +qed + +instance + by standard (auto simp add: less_eq_eint_def less_eint_def plus_eint_def split: eint.splits) + +end + +instance eint :: "{strict_ordered_comm_monoid_add}" +proof + show "a < b \ c < d \ a + c < b + d" for a b c d :: eint + by (cases a b c d rule: eint2_cases[case_product eint2_cases]) auto +qed + +(* BH: These equations are already proven generally for any type in +class linordered_semidom. However, eint is not in that class because +it does not have the cancellation property. Would it be worthwhile to +a generalize linordered_semidom to a new class that includes eint? *) + +lemma add_diff_assoc_eint: "z \ y \ x + (y - z) = x + y - (z::eint)" +by(cases x)(auto simp add: diff_eint_def split: eint.split) + +lemma eint_ord_number [simp]: + "(numeral m :: eint) \ numeral n \ (numeral m :: nat) \ numeral n" + "(numeral m :: eint) < numeral n \ (numeral m :: nat) < numeral n" + apply simp + by simp + +lemma infinity_ileE [elim!]: "\ \ eint m \ R" + by simp + +lemma infinity_ilessE [elim!]: "\ < eint m \ R" + by simp + +lemma imult_infinity: "(0::eint) < n \ \ * n = \" + by (simp add: zero_eint_def less_eint_def split: eint.splits) + +lemma imult_infinity_right: "(0::eint) < n \ n * \ = \" + by (simp add: zero_eint_def less_eint_def split: eint.splits) + +lemma min_eint_simps [simp]: + "min (eint m) (eint n) = eint (min m n)" + "min q (\::eint) = q" + "min (\::eint) q = q" + by (auto simp add: min_def) + +lemma max_eint_simps [simp]: + "max (eint m) (eint n) = eint (max m n)" + "max q \ = (\::eint)" + "max \ q = (\::eint)" + by (simp_all add: max_def) + +lemma eint_ile: "n \ eint m \ \k. n = eint k" + by (cases n) simp_all + +lemma eint_iless: "n < eint m \ \k. n = eint k" + by (cases n) simp_all + +lemma iadd_le_eint_iff: + "x + y \ eint n \ (\y' x'. x = eint x' \ y = eint y' \ x' + y' \ n)" +by(cases x y rule: eint.exhaust[case_product eint.exhaust]) simp_all + +lemma chain_incr: "\i. \j. Y i < Y j \ \j. eint k < Y j" +proof- + assume A: "\i. \j. Y i < Y j" + then have "\i. \n::int. Y i = eint n" + by (metis eint.exhaust eint_ord_simps(6)) + then obtain i n where in_def: "Y (i::'a) = eint n" + by blast + show "\j. eint k < Y j" + proof(rule ccontr) + assume C: "\(\j. eint k < Y j)" + then have C':"\j. Y j \ eint k" + using le_less_linear + by blast + then have "Y (i::'a) \ eint k" + by simp + have "\m::nat. \j::'a. Y j \ eint (n + int m)" + proof- fix m + show "\j::'a. Y j \ eint (n + int m)" + apply(induction m) + apply (metis in_def int_ops(1) order_refl plus_int_code(1)) + proof- fix m + assume "\j. eint (n + int m) \ Y j" + then obtain j where j_def: "eint (n + int m) \ Y j" + by blast + obtain j' where j'_def: "Y j < Y j'" + using A by blast + have "eint (n + int (Suc m)) = eint (n + m + 1)" + by auto + then have "eint (n + int (Suc m)) \ Y j'" + using j_def j'_def eint_ord_plus_one[of "n + m" "Y j" "Y j'"] + by presburger + then show "\j. eint (n + int (Suc m)) \ Y j" + by blast + qed + qed + then show False + by (metis A C \Y i \ eint k\ eint_ord_simps(1) in_def + order.not_eq_order_implies_strict zle_iff_zadd) + qed +qed + +lemma eint_ord_Suc: + assumes "(x::eint) < y" + shows "x + 1 < y + 1" + apply(cases "y = \") + using assms i1_ne_infinity sum_infinity_imp_summand_infinity + apply fastforce + by (metis add_mono_thms_linordered_semiring(3) assms eSuc_inject order_less_le) + +lemma eSuc_ile_mono [simp]: "(n::eint) + 1 \ m+ 1 \ n \ m" + by (meson add_mono_thms_linordered_semiring(3) eint_ord_Suc linorder_not_le) + +lemma eSuc_mono [simp]: "(n::eint) + 1 < m+ 1 \ n < m" + by (meson add_mono_thms_linordered_semiring(3) eint_ord_Suc linorder_not_le) + +lemma ile_eSuc [simp]: "(n::eint) \ n + 1" + by (metis add.right_neutral add_left_mono eint_1_iff(2) eint_ord_code(1) linear not_one_le_zero zero_eint_def) + +lemma ileI1: "(m::eint) < n \ m + 1 \ n" + by (metis eSuc_eint eint.exhaust eint_ex_split eint_iless eint_ord_Suc eint_ord_code(6) + eint_ord_plus_one eint_ord_simps(3) less_le_trans linear ) + +lemma Suc_ile_eq: "eint (m +1) \ n \ eint m < n" + by (cases n) auto + +lemma iless_Suc_eq [simp]: "eint m < n + 1 \ eint m \ n" + by (metis Suc_ile_eq eSuc_eint eSuc_ile_mono) + +lemma eSuc_max: "(max (x::eint) y) + 1 = max (x+1) (y+1)" + by (simp add: max_def) + +lemma eSuc_Max: + assumes "finite A" "A \ ({}::eint set)" + shows " (Max A) + 1 = Max ((+)1 ` A)" +using assms proof induction + case (insert x A) + thus ?case + using Max_insert[of A x] Max_singleton[of x] add.commute[of 1] eSuc_max finite_imageI + image_insert image_is_empty + by (simp add: add.commute hom_Max_commute) +qed simp + +instantiation eint :: "{order_top}" +begin + +definition top_eint :: eint where "top_eint = \" + +instance + by standard (simp add: top_eint_def) + +end + +lemma finite_eint_bounded: + assumes le_fin: "\y. y \ A \ eint m \ y \ y \ eint n" + shows "finite A" +proof (rule finite_subset) + show "finite (eint ` {m..n})" by blast + have "A \ {eint m..eint n}" using le_fin by fastforce + also have "\ \ eint ` {m..n}" + apply (rule subsetI) + subgoal for x by (cases x) auto + done + finally show "A \ eint ` {m..n}" . +qed + + +subsection \Cancellation simprocs\ + +lemma add_diff_cancel_eint[simp]: "x \ \ \ x + y - x = (y::eint)" +by (metis add.commute add.right_neutral add_diff_assoc_eint idiff_self order_refl) + +lemma eint_add_left_cancel: "a + b = a + c \ a = (\::eint) \ b = c" + unfolding plus_eint_def by (simp split: eint.split) + +lemma eint_add_left_cancel_le: "a + b \ a + c \ a = (\::eint) \ b \ c" + unfolding plus_eint_def by (simp split: eint.split) + +lemma eint_add_left_cancel_less: "a + b < a + c \ a \ (\::eint) \ b < c" + unfolding plus_eint_def by (simp split: eint.split) + +lemma plus_eq_infty_iff_eint: "(m::eint) + n = \ \ m=\ \ n=\" +using eint_add_left_cancel by fastforce + +ML \ +structure Cancel_Enat_Common = +struct + (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) + fun find_first_t _ _ [] = raise TERM("find_first_t", []) + | find_first_t past u (t::terms) = + if u aconv t then (rev past @ terms) + else find_first_t (t::past) u terms + + fun dest_summing (Const (\<^const_name>\Groups.plus\, _) $ t $ u, ts) = + dest_summing (t, dest_summing (u, ts)) + | dest_summing (t, ts) = t :: ts + + val mk_sum = Arith_Data.long_mk_sum + fun dest_sum t = dest_summing (t, []) + val find_first = find_first_t [] + val trans_tac = Numeral_Simprocs.trans_tac + val norm_ss = + simpset_of (put_simpset HOL_basic_ss \<^context> + addsimps @{thms ac_simps add_0_left add_0_right}) + fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) + fun simplify_meta_eq ctxt cancel_th th = + Arith_Data.simplify_meta_eq [] ctxt + ([th, cancel_th] MRS trans) + fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) +end + +structure Eq_Enat_Cancel = ExtractCommonTermFun +(open Cancel_Enat_Common + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ \<^typ>\eint\ + fun simp_conv _ _ = SOME @{thm eint_add_left_cancel} +) + +structure Le_Enat_Cancel = ExtractCommonTermFun +(open Cancel_Enat_Common + val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ + val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ \<^typ>\eint\ + fun simp_conv _ _ = SOME @{thm eint_add_left_cancel_le} +) + +structure Less_Enat_Cancel = ExtractCommonTermFun +(open Cancel_Enat_Common + val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ + val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ \<^typ>\eint\ + fun simp_conv _ _ = SOME @{thm eint_add_left_cancel_less} +) +\ + +simproc_setup eint_eq_cancel + ("(l::eint) + m = n" | "(l::eint) = m + n") = + \fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + +simproc_setup eint_le_cancel + ("(l::eint) + m \ n" | "(l::eint) \ m + n") = + \fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + +simproc_setup eint_less_cancel + ("(l::eint) + m < n" | "(l::eint) < m + n") = + \fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + +text \TODO: add regression tests for these simprocs\ + +text \TODO: add simprocs for combining and cancelling numerals\ + +subsection \Well-ordering\ + +lemma less_eintE: + "[| n < eint m; !!k. n = eint k ==> k < m ==> P |] ==> P" +by (induct n) auto + +lemma less_infinityE: + "[| n < \; !!k. n = eint k ==> P |] ==> P" +by (induct n) auto + + +subsection \Traditional theorem names\ + +lemmas eint_defs = zero_eint_def one_eint_def + plus_eint_def less_eq_eint_def less_eint_def + +instantiation eint :: uminus +begin + +definition +"- b = (case b of \ \ \ | eint m \ eint (-m))" + +lemma eint_uminus_eq: +"(a::eint) + (-a) = a - a" + apply(induction a) + apply (simp add: uminus_eint_def) + by simp + + + + +instance.. +end + +section\Additional Lemmas (Useful for the Proof of Hensel's Lemma)\ + +lemma eint_mult_mono: + assumes "(c::eint) > 0 \ c \ \" + assumes "k > n" + shows "k*c > n*c" + using assms apply(induction k, induction n, induction c) + by(auto simp add: zero_eint_def) + +lemma eint_mult_mono': + assumes "(c::eint) \ 0 \ c \ \" + assumes "k > n" + shows "k*c \ n*c" + apply(cases "c = 0") + apply (metis add.right_neutral assms(2) eint_add_left_cancel eint_ord_code(3) + eint_ord_simps(4) eq_iff less_le_trans mult.commute mult_eSuc_right' + mult_one_right not_less times_eint_simps(4) zero_eint_def) + using assms eint_mult_mono + by (simp add: le_less) + +lemma eint_minus_le: + assumes "(b::eint) < c" + shows "c - b > 0" + using assms apply(induction b, induction c) + by (auto simp add: zero_eint_def) + +lemma eint_nat_times: + assumes "(c::eint) > 0" + shows "(Suc n)*(c::eint) > 0" + using assms apply(induction c) + apply (simp add: zero_eint_def) + by simp + +lemma eint_pos_times_is_pos: + assumes "(c::eint) > 0" + assumes "b > 0" + shows "b*c > 0" + using assms apply(induction c, induction b) + by(auto simp add: zero_eint_def imult_infinity_right) + +lemma eint_nat_is_pos: +"eint (Suc n) > 0" + by (simp add: zero_eint_def) + +lemma eint_pow_int_is_pos: + assumes "n > 0" + shows "eint n > 0" + using assms by (simp add: zero_eint_def) + +lemma eint_nat_times': + assumes "(c::eint) \ 0" + shows "(Suc n)*c \ 0" + using assms zero_eint_def by fastforce + +lemma eint_pos_int_times_ge: + assumes "(c::eint) \ 0" + assumes "n > 0" + shows "eint n * c \ c" + using assms apply(induction c) + apply (smt eSuc_eint eint.distinct(2) eint_mult_mono' eint_pow_int_is_pos eq_iff ileI1 less_le mult.commute mult_one_right one_eint_def zero_eint_def) + by simp + +lemma eint_pos_int_times_gt: + assumes "(c::eint) > 0" + assumes "c \\" + assumes "n > 1" + shows "eint n * c > c" + using assms eint_mult_mono[of c 1 "eint n"] + by (metis eint_ord_simps(2) mult_one_left one_eint_def) + +lemma eint_add_cancel_fact[simp]: + assumes "(c::eint) \ \" + shows "c + (b - c) = b" + using assms apply(induction c, induction b) + by auto + +lemma nat_mult_not_infty[simp]: + assumes "c \ \" + shows "(eint n) * c \ \" + using assms by auto + +lemma eint_minus_distl: + assumes "(b::eint) \ d" + shows "b*c - d*c = (b-d)*c" + using assms apply(induction c, induction b, induction d) + apply (metis add_diff_cancel_eint distrib_right eint.distinct(2) eint_add_cancel_fact nat_mult_not_infty) + apply simp + apply simp + by (simp add: mult.commute times_eint_def) + +lemma eint_minus_distr: + assumes "(b::eint) \ d" + shows "c*(b - d) = c*b - c*d" + by (metis assms eint_minus_distl mult.commute) + +lemma eint_int_minus_distr: +"(eint n)*c - (eint m)*c = eint (n - m) * c" + by (metis add.right_neutral distrib_right eint_add_left_cancel eint_minus_distl idiff_eint_eint + idiff_infinity idiff_self infinity_ne_i0 nat_mult_not_infty not_eint_eq times_eint_simps(4)) + +lemma eint_2_minus_1_mult[simp]: +"2*(b::eint) - b = b" +proof - + have "\e. (\::eint) * e = \" + by (simp add: times_eint_def) + then show ?thesis + by (metis add_diff_cancel_eint idiff_infinity mult.commute mult_eSuc_right' mult_one_right one_add_one one_eint_def plus_eint_simps(1)) +qed + +lemma eint_minus_comm: +"(d::eint) + b - c = d - c + b" +apply(induction c ) + apply (metis add.assoc add_diff_cancel_eint eint.distinct(2) eint_add_cancel_fact) + apply(induction d) + apply (metis distrib_left eint2_cases eint_minus_distl i1_ne_infinity idiff_infinity_right + mult_one_left plus_eq_infty_iff_eint sum_infinity_imp_summand_infinity times_eint_simps(3)) + apply(induction b) + apply simp + by simp + +lemma ge_plus_pos_imp_gt: + assumes "(c::eint) \\" + assumes "(b::eint) > 0" + assumes "d \ c + b" + shows "d > c" + using assms apply(induction d, induction c) + apply (metis add.comm_neutral assms(2) eint_add_left_cancel_less less_le_trans) + apply blast + by simp + +lemma eint_minus_ineq: + assumes "(c::eint) \\" + assumes "b \ d" + shows "b - c \ d - c" + by (metis add_left_mono antisym assms(1) assms(2) eint_add_cancel_fact linear) + +lemma eint_minus_ineq': + assumes "(c::eint) \\" + assumes "b \ d" + assumes "(e::eint) > 0" + assumes "e \ \" + shows "e*(b - c) \ e*(d - c)" + using assms eint_minus_ineq + by (metis eint_mult_mono' eq_iff less_le mult.commute) + +lemma eint_minus_ineq'': + assumes "(c::eint) \\" + assumes "b \ d" + assumes "(e::eint) > 0" + assumes "e \ \" + shows "e*(b - c) \ e*d - e*c" + using assms eint_minus_ineq' + proof - + have "\e. (0::eint) + e = e" + by simp + then have f1: "e * 0 = 0" + by (metis add_diff_cancel_eint assms(4) idiff_self mult_eSuc_right' mult_one_right) + have "\ \ c * e" + using assms(1) assms(4) eint_pos_times_is_pos by auto + then show ?thesis + using f1 by (metis assms(1) assms(2) assms(3) assms(4) eint_minus_distl eint_minus_ineq' idiff_self mult.commute) + qed + +lemma eint_min_ineq: + assumes "(b::eint) \ min c d" + assumes "c > e" + assumes "d > e" + shows "b > e" + by (meson assms(1) assms(2) assms(3) less_le_trans min_le_iff_disj) + +lemma eint_plus_times: + assumes "(d::eint) \ 0" + assumes "(b::eint) \ c + (eint k)*d" + assumes "k \ l" + shows "b \ c + l*d" +proof- + have "k*d \ l*d" + by (smt assms(1) assms(3) eint_mult_mono' eint_ord_simps(2) eq_iff times_eint_simps(4)) + thus ?thesis + by (meson add_mono_thms_linordered_semiring(2) assms(2) order_subst2) +qed +end + diff --git a/thys/Padic_Ints/Function_Ring.thy b/thys/Padic_Ints/Function_Ring.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Function_Ring.thy @@ -0,0 +1,1250 @@ +theory Function_Ring + imports "HOL-Algebra.Ring" "HOL-Library.FuncSet" "HOL-Algebra.Module" +begin + +text\ + This theory formalizes basic facts about the ring of extensional functions from a fixed set to + a fixed ring. This will be useful for providing a generic framework for various constructions + related to the $p$-adics such as polynomial evaluation and sequences. The rings of semialgebraic + functions will be defined as subrings of these function rings, which will be necessary for the + proof of $p$-adic quantifier elimination. +\ + +(**************************************************************************************************) +(**************************************************************************************************) +section\The Ring of Extensional Functions from a Fixed Base Set to a Fixed Base Ring\ +(**************************************************************************************************) +(**************************************************************************************************) + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Basic Operations on Extensional Functions\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition function_mult:: "'c set \ ('a, 'b) ring_scheme \ ('c \ 'a) \ ('c \ 'a) \ ('c \ 'a)" where +"function_mult S R f g = (\x \ S. (f x) \\<^bsub>R\<^esub> (g x))" + +abbreviation(input) ring_function_mult:: "('a, 'b) ring_scheme \ ('a \ 'a) \ ('a \ 'a) \ ('a \ 'a)" where +"ring_function_mult R f g \ function_mult (carrier R) R f g" + +definition function_add:: "'c set \ ('a, 'b) ring_scheme \ ('c \ 'a) \ ('c \ 'a) \ ('c \ 'a)" where +"function_add S R f g = (\x \ S. (f x) \\<^bsub>R\<^esub> (g x))" + +abbreviation(input) ring_function_add:: "('a, 'b) ring_scheme \ ('a \ 'a) \ ('a \ 'a) \ ('a \ 'a)" where +"ring_function_add R f g \ function_add (carrier R) R f g" + +definition function_one:: "'c set \ ('a, 'b) ring_scheme \ ('c \ 'a)" where +"function_one S R = (\x \ S. \\<^bsub>R\<^esub>)" + +abbreviation(input) ring_function_one :: "('a, 'b) ring_scheme \ ('a \ 'a)" where +"ring_function_one R \ function_one (carrier R) R" + +definition function_zero:: "'c set \ ('a, 'b) ring_scheme \ ('c \ 'a)" where +"function_zero S R = (\x \ S. \\<^bsub>R\<^esub>)" + +abbreviation(input) ring_function_zero :: "('a, 'b) ring_scheme \ ('a \ 'a)" where +"ring_function_zero R \ function_zero (carrier R) R" + +definition function_uminus:: "'c set \ ('a, 'b) ring_scheme \ ('c \ 'a) \ ('c \ 'a)" where +"function_uminus S R a = (\ x \ S. \\<^bsub>R\<^esub> (a x))" + +definition ring_function_uminus:: " ('a, 'b) ring_scheme \ ('a \ 'a) \ ('a \ 'a)" where +"ring_function_uminus R a = function_uminus (carrier R) R a" + +definition function_scalar_mult:: "'c set \ ('a, 'b) ring_scheme \ 'a \ ('c \ 'a) \ ('c \ 'a)" where +"function_scalar_mult S R a f = (\ x \ S. a \\<^bsub>R\<^esub> (f x))" + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Defining the Ring of Extensional Functions\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition function_ring:: "'c set \ ('a, 'b) ring_scheme \ ( 'a, 'c \ 'a) module" where +"function_ring S R = \ + carrier = extensional_funcset S (carrier R), + Group.monoid.mult = (function_mult S R), + one = (function_one S R), + zero = (function_zero S R), + add = (function_add S R), + smult = function_scalar_mult S R \ " + +text\The following locale consists of a struct R, and a distinguished set S which is meant to serve as the domain for a ring of functions $S \to carrier R$. \ +locale struct_functions = + fixes R ::"('a, 'b) partial_object_scheme" (structure) + and S :: "'c set" + +text\The following are locales which fix a ring R (which may be commutative, a domain, or a field) and a function ring F of extensional functions from a fixed set S to $carrier R$\ +locale ring_functions = struct_functions + R?: ring R + + fixes F (structure) + defines F_def: "F \ function_ring S R" + +locale cring_functions = ring_functions + R?: cring R + +locale domain_functions = ring_functions + R?: domain R + +locale field_functions = ring_functions + R?: field R + +sublocale cring_functions < ring_functions + apply (simp add: ring_functions_axioms) + by (simp add: F_def) + +sublocale domain_functions < ring_functions + apply (simp add: ring_functions_axioms) + by (simp add: F_def) + +sublocale domain_functions < cring_functions + apply (simp add: cring_functions_def is_cring ring_functions_axioms) + by (simp add: F_def) + +sublocale field_functions < domain_functions + apply (simp add: domain_axioms domain_functions_def ring_functions_axioms) + by (simp add: F_def) + +sublocale field_functions < ring_functions + apply (simp add: ring_functions_axioms) + by (simp add: F_def) + +sublocale field_functions < cring_functions + apply (simp add: cring_functions_axioms) + by (simp add: F_def) + +abbreviation(input) ring_function_ring:: "('a, 'b) ring_scheme \ ('a, 'a \ 'a) module" ("Fun") where +"ring_function_ring R \ function_ring (carrier R) R" + + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Algebraic Properties of the Basic Operations\ + (**************************************************************************************************) + (**************************************************************************************************) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Basic Carrier Facts\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_ring_defs: +"carrier F = extensional_funcset S (carrier R)" +"(\\<^bsub>F\<^esub>) = (function_mult S R)" +"(\\<^bsub>F\<^esub>) = (function_add S R)" +"\\<^bsub>F\<^esub> = function_one S R" +"\\<^bsub>F\<^esub> = function_zero S R" +"(\\<^bsub>F\<^esub>) = function_scalar_mult S R" + unfolding F_def + by ( auto simp add: function_ring_def) + +lemma(in ring_functions) function_ring_car_memE: + assumes "a \ carrier F" + shows "a \ extensional S" + "a \ S \ carrier R" + using assms function_ring_defs apply auto[1] + using assms function_ring_defs PiE_iff apply blast + using assms function_ring_defs(1) by fastforce + +lemma(in ring_functions) function_ring_car_closed: + assumes "a \ S" + assumes "f \ carrier F" + shows "f a \ carrier R" + using assms unfolding function_ring_def F_def by auto + +lemma(in ring_functions) function_ring_not_car: + assumes "a \ S" + assumes "f \ carrier F" + shows "f a = undefined" + using assms unfolding function_ring_def F_def by auto + +lemma(in ring_functions) function_ring_car_eqI: + assumes "f \ carrier F" + assumes "g \ carrier F" + assumes "\a. a \ S \ f a = g a" + shows "f = g" + using assms(1) assms(2) assms(3) extensionalityI function_ring_car_memE(1) by blast + +lemma(in ring_functions) function_ring_car_memI: + assumes "\a. a \ S \ f a \ carrier R" + assumes "\ a. a \ S\ f a = undefined" + shows "f \ carrier F" + using function_ring_defs assms + unfolding extensional_funcset_def + by (simp add: \\a. a \ S \ f a \ carrier R\ extensional_def) + +lemma(in ring) function_ring_car_memI: + assumes "\a. a \ S \ f a \ carrier R" + assumes "\ a. a \ S\ f a = undefined" + shows "f \ carrier (function_ring S R)" + by (simp add: assms(1) assms(2) local.ring_axioms ring_functions.function_ring_car_memI ring_functions.intro) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Basic Multiplication Facts\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_mult_eval_car: + assumes "a \ S" + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "(f \\<^bsub>F\<^esub> g) a = (f a) \ (g a)" + using assms function_ring_defs + unfolding function_mult_def + by simp + +lemma(in ring_functions) function_mult_eval_closed: + assumes "a \ S" + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "(f \\<^bsub>F\<^esub> g) a \ carrier R" + using assms function_mult_eval_car + using F_def ring_functions.function_ring_car_closed ring_functions_axioms by fastforce + +lemma(in ring_functions) fun_mult_closed: + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "f \\<^bsub>F\<^esub> g \ carrier F" + apply(rule function_ring_car_memI) + apply (simp add: assms(1) assms(2) function_mult_eval_closed) + by (simp add: function_mult_def function_ring_defs(2)) + +lemma(in ring_functions) fun_mult_eval_assoc: + assumes "x \ carrier F" + assumes "y \ carrier F" + assumes " z \ carrier F" + assumes "a \ S" + shows "(x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a = (x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a" +proof- + have 0: "(x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a = (x a) \ (y a) \ (z a) " + by (simp add: assms(1) assms(2) assms(3) assms(4) fun_mult_closed function_mult_eval_car) + have 1: "(x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a = (x a) \ ((y a) \ (z a))" + by (simp add: assms(1) assms(2) assms(3) assms(4) fun_mult_closed function_mult_eval_car) + have 2:"(x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a = (x a) \ (y a) \ (z a)" + using 1 assms + by (simp add: function_ring_car_closed m_assoc) + show ?thesis + using 0 2 by auto +qed + +lemma(in ring_functions) fun_mult_assoc: + assumes "x \ carrier F" + assumes "y \ carrier F" + assumes "z \ carrier F" + shows "(x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) = (x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z))" + using fun_mult_eval_assoc[of x] + by (simp add: assms(1) assms(2) assms(3) fun_mult_closed function_ring_car_eqI) + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Basic Addition Facts\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) fun_add_eval_car: + assumes "a \ S" + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "(f \\<^bsub>F\<^esub> g) a = (f a) \ (g a)" + by (simp add: assms(1) function_add_def function_ring_defs(3)) + +lemma(in ring_functions) fun_add_eval_closed: + assumes "a \ S" + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "(f \\<^bsub>F\<^esub> g) a \ carrier R" + using assms unfolding F_def + using F_def fun_add_eval_car function_ring_car_closed + by auto + +lemma(in ring_functions) fun_add_closed: + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "f \\<^bsub>F\<^esub> g \ carrier F" + apply(rule function_ring_car_memI) + using assms unfolding F_def + using F_def fun_add_eval_closed apply blast + by (simp add: function_add_def function_ring_def) + +lemma(in ring_functions) fun_add_eval_assoc: + assumes "x \ carrier F" + assumes "y \ carrier F" + assumes " z \ carrier F" + assumes "a \ S" + shows "(x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a = (x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a" +proof- + have 0: "(x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a = (x a) \ (y a) \ (z a) " + by (simp add: assms(1) assms(2) assms(3) assms(4) fun_add_closed fun_add_eval_car) + have 1: "(x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a = (x a) \ ((y a) \ (z a))" + by (simp add: assms(1) assms(2) assms(3) assms(4) fun_add_closed fun_add_eval_car) + have 2:"(x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)) a = (x a) \ (y a) \ (z a)" + using 1 assms + by (simp add: add.m_assoc function_ring_car_closed) + show ?thesis + using 0 2 by auto +qed + +lemma(in ring_functions) fun_add_assoc: + assumes "x \ carrier F" + assumes "y \ carrier F" + assumes " z \ carrier F" + shows "x \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z = x \\<^bsub>F\<^esub> (y \\<^bsub>F\<^esub> z)" + apply(rule function_ring_car_eqI) + using assms apply (simp add: fun_add_closed) + apply (simp add: assms(1) assms(2) assms(3) fun_add_closed) + by (simp add: assms(1) assms(2) assms(3) fun_add_eval_assoc) + +lemma(in ring_functions) fun_add_eval_comm: + assumes "a \ S" + assumes "x \ carrier F" + assumes "y \ carrier F" + shows "(x \\<^bsub>F\<^esub> y) a = (y \\<^bsub>F\<^esub> x) a" + by (metis F_def assms(1) assms(2) assms(3) fun_add_eval_car ring.ring_simprules(10) ring_functions.function_ring_car_closed ring_functions_axioms ring_functions_def) + +lemma(in ring_functions) fun_add_comm: + assumes "x \ carrier F" + assumes "y \ carrier F" + shows "x \\<^bsub>F\<^esub> y = y \\<^bsub>F\<^esub> x" + using fun_add_eval_comm assms + by (metis (no_types, hide_lams) fun_add_closed function_ring_car_eqI) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Basic Facts About the Multiplicative Unit\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_one_eval: + assumes "a \ S" + shows "\\<^bsub>F\<^esub> a = \" + using assms function_ring_defs unfolding function_one_def + by simp + +lemma(in ring_functions) function_one_closed: +"\\<^bsub>F\<^esub> \carrier F" + apply(rule function_ring_car_memI) + using function_ring_defs + using function_one_eval apply auto[1] + by (simp add: function_one_def function_ring_defs(4)) + +lemma(in ring_functions) function_times_one_l: + assumes "a \ carrier F" + shows "\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a = a" +proof(rule function_ring_car_eqI) + show "\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a \ carrier F" + using assms fun_mult_closed function_one_closed + by blast + show " a \ carrier F" + using assms by simp + show "\c. c \ S \ (\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a) c = a c " + by (simp add: assms function_mult_eval_car function_one_eval function_one_closed function_ring_car_closed) +qed + +lemma(in ring_functions) function_times_one_r: + assumes "a \ carrier F" + shows "a\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> = a" +proof(rule function_ring_car_eqI) + show "a\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> \ carrier F" + using assms fun_mult_closed function_one_closed + by blast + show " a \ carrier F" + using assms by simp + show "\c. c \ S \ (a\\<^bsub>F\<^esub> \\<^bsub>F\<^esub>) c = a c " + using assms + by (simp add: function_mult_eval_car function_one_eval function_one_closed function_ring_car_closed) +qed + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Basic Facts About the Additive Unit\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_zero_eval: + assumes "a \ S" + shows "\\<^bsub>F\<^esub> a = \" + using assms function_ring_defs + unfolding function_zero_def + by simp + +lemma(in ring_functions) function_zero_closed: +"\\<^bsub>F\<^esub> \carrier F" + apply(rule function_ring_car_memI) + apply (simp add: function_zero_eval) + by (simp add: function_ring_defs(5) function_zero_def) + +lemma(in ring_functions) fun_add_zeroL: + assumes "a \ carrier F" + shows "\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a = a" +proof(rule function_ring_car_eqI) + show "\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a \ carrier F" + using assms fun_add_closed function_zero_closed + by blast + show "a \ carrier F" + using assms by simp + show "\c. c \ S \ (\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> a) c = a c " + using assms F_def fun_add_eval_car function_zero_closed + ring_functions.function_zero_eval ring_functions_axioms + by (simp add: ring_functions.function_zero_eval function_ring_car_closed) +qed + +lemma(in ring_functions) fun_add_zeroR: + assumes "a \ carrier F" + shows "a \\<^bsub>F\<^esub> \\<^bsub>F\<^esub> = a" + using assms fun_add_comm fun_add_zeroL + by (simp add: function_zero_closed) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Distributive Laws\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_mult_r_distr: + assumes "x \ carrier F" + assumes" y \ carrier F" + assumes " z \ carrier F" + shows " (x \\<^bsub>F\<^esub> y) \\<^bsub>F\<^esub> z = x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z" +proof(rule function_ring_car_eqI) + show "(x \\<^bsub>F\<^esub> y) \\<^bsub>F\<^esub> z \ carrier F" + by (simp add: assms(1) assms(2) assms(3) fun_add_closed fun_mult_closed) + show "x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z \ carrier F" + by (simp add: assms(1) assms(2) assms(3) fun_add_closed fun_mult_closed) + show "\a. a \ S \ ((x \\<^bsub>F\<^esub> y) \\<^bsub>F\<^esub> z) a = (x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a" + proof- + fix a + assume A: "a \ S" + show "((x \\<^bsub>F\<^esub> y) \\<^bsub>F\<^esub> z) a = (x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y \\<^bsub>F\<^esub> z) a" + using A assms fun_add_eval_car[of a x y] fun_add_eval_car[of a "x \\<^bsub>F\<^esub>z" "y \\<^bsub>F\<^esub> z"] + function_mult_eval_car[of a "x \\<^bsub>F\<^esub> y" z] semiring_simprules(10) + F_def + by (smt fun_add_closed function_mult_eval_car function_ring_car_closed + ring_functions.fun_mult_closed ring_functions_axioms) + qed +qed + +lemma(in ring_functions) function_mult_l_distr: + assumes "x \ carrier F" + assumes" y \ carrier F" + assumes " z \ carrier F" + shows "z \\<^bsub>F\<^esub> (x \\<^bsub>F\<^esub> y) = z \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y" +proof(rule function_ring_car_eqI) + show "z \\<^bsub>F\<^esub> (x \\<^bsub>F\<^esub> y) \ carrier F" + by (simp add: assms(1) assms(2) assms(3) fun_add_closed fun_mult_closed) + show "z \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y \ carrier F" + by (simp add: assms(1) assms(2) assms(3) fun_add_closed fun_mult_closed) + show "\a. a \ S \ (z \\<^bsub>F\<^esub> (x \\<^bsub>F\<^esub> y)) a = (z \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y) a" + proof- + fix a + assume A: "a \ S" + show "(z \\<^bsub>F\<^esub> (x \\<^bsub>F\<^esub> y)) a = (z \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> z \\<^bsub>F\<^esub> y) a" + using A assms function_ring_defs fun_add_closed fun_mult_closed + function_mult_eval_car[of a z "x \\<^bsub>F\<^esub> y"] + function_mult_eval_car[of a z x] + function_mult_eval_car[of a z y] + fun_add_eval_car[of a x y] + semiring_simprules(13) + fun_add_eval_car function_ring_car_closed by auto + qed +qed + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Additive Inverses\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_uminus_closed: + assumes "f \ carrier F" + shows "function_uminus S R f \ carrier F" +proof(rule function_ring_car_memI) + show "\a. a \ S \ function_uminus S R f a \ carrier R" + using assms function_ring_car_closed[of _ f] unfolding function_uminus_def + by simp + show "\a. a \ S \ function_uminus S R f a = undefined" + by (simp add: function_uminus_def) +qed + +lemma(in ring_functions) function_uminus_eval: + assumes "a \ S" + assumes "f \ carrier F" + shows "(function_uminus S R f) a = \ (f a)" + using assms unfolding function_uminus_def + by simp + +lemma(in ring_functions) function_uminus_add_r: + assumes "a \ S" + assumes "f \ carrier F" + shows "f \\<^bsub>F\<^esub> function_uminus S R f = \\<^bsub>F\<^esub>" + apply(rule function_ring_car_eqI) + using assms fun_add_closed function_uminus_closed apply blast + unfolding F_def using F_def function_zero_closed apply blast + using F_def assms(2) fun_add_eval_car function_ring_car_closed function_uminus_closed + function_uminus_eval function_zero_eval r_neg by auto + +lemma(in ring_functions) function_uminus_add_l: + assumes "a \ S" + assumes "f \ carrier F" + shows "function_uminus S R f \\<^bsub>F\<^esub> f = \\<^bsub>F\<^esub>" + using assms(1) assms(2) fun_add_comm function_uminus_add_r function_uminus_closed by auto + + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Scalar Multiplication\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_smult_eval: + assumes "a \ carrier R" + assumes "f \ carrier F" + assumes "b \ S" + shows "(a \\<^bsub>F\<^esub> f) b = a \ (f b)" + using function_ring_defs(6) unfolding function_scalar_mult_def + by(simp add: assms) + +lemma(in ring_functions) function_smult_closed: + assumes "a \ carrier R" + assumes "f \ carrier F" + shows "a \\<^bsub>F\<^esub> f \ carrier F" + apply(rule function_ring_car_memI) + using function_smult_eval assms + apply (simp add: function_ring_car_closed) + using function_scalar_mult_def F_def + by (metis function_ring_defs(6) restrict_apply) + +lemma(in ring_functions) function_smult_assoc1: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "f \ carrier F" + shows "b \\<^bsub>F\<^esub> (a \\<^bsub>F\<^esub> f) = (b \ a)\\<^bsub>F\<^esub>f" + apply(rule function_ring_car_eqI) + using assms function_smult_closed apply simp + using assms function_smult_closed apply simp + by (metis F_def assms(1) assms(2) assms(3) function_mult_eval_closed function_one_closed + function_smult_eval function_times_one_r m_assoc m_closed ring_functions.function_smult_closed ring_functions_axioms) + +lemma(in ring_functions) function_smult_assoc2: + assumes "a \ carrier R" + assumes "f \ carrier F" + assumes "g \ carrier F" + shows "(a \\<^bsub>F\<^esub> f)\\<^bsub>F\<^esub>g = a \\<^bsub>F\<^esub> (f \\<^bsub>F\<^esub> g)" + apply(rule function_ring_car_eqI) + using assms function_smult_closed apply (simp add: fun_mult_closed) + apply (simp add: assms(1) assms(2) assms(3) fun_mult_closed function_smult_closed) + by (metis (full_types) F_def assms(1) assms(2) assms(3) fun_mult_closed + function_mult_eval_car function_smult_closed function_smult_eval m_assoc ring_functions.function_ring_car_closed ring_functions_axioms) + +lemma(in ring_functions) function_smult_one: + assumes "f \ carrier F" + shows "\\\<^bsub>F\<^esub>f = f" + apply(rule function_ring_car_eqI) + apply (simp add: assms function_smult_closed) + apply (simp add: assms) + by (simp add: assms function_ring_car_closed function_smult_eval) + +lemma(in ring_functions) function_smult_l_distr: +"[| a \ carrier R; b \ carrier R; x \ carrier F |] ==> + (a \ b) \\<^bsub>F\<^esub> x = a \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> b \\<^bsub>F\<^esub> x" + apply(rule function_ring_car_eqI) + apply (simp add: function_smult_closed) + apply (simp add: fun_add_closed function_smult_closed) + using function_smult_eval + by (simp add: fun_add_eval_car function_ring_car_closed function_smult_closed l_distr) + +lemma(in ring_functions) function_smult_r_distr: + "[| a \ carrier R; x \ carrier F; y \ carrier F |] ==> + a \\<^bsub>F\<^esub> (x \\<^bsub>F\<^esub> y) = a \\<^bsub>F\<^esub> x \\<^bsub>F\<^esub> a \\<^bsub>F\<^esub> y" + apply(rule function_ring_car_eqI) + apply (simp add: fun_add_closed function_smult_closed) + apply (simp add: fun_add_closed function_smult_closed) + by (simp add: fun_add_closed fun_add_eval_car function_ring_car_closed function_smult_closed function_smult_eval r_distr) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\The Ring of Functions Forms an Algebra\ + (**************************************************************************************************) + (**************************************************************************************************) + +lemma(in ring_functions) function_ring_is_abelian_group: +"abelian_group F" + apply(rule abelian_groupI) + apply (simp add: fun_add_closed) + apply (simp add: function_zero_closed) + using fun_add_assoc apply simp + apply (simp add: fun_add_comm) + apply (simp add: fun_add_comm fun_add_zeroR function_zero_closed) + using fun_add_zeroL function_ring_car_eqI function_uminus_add_l + function_uminus_closed function_zero_closed by blast + +lemma(in ring_functions) function_ring_is_monoid: +"monoid F" + apply(rule monoidI) + apply (simp add: fun_mult_closed) + apply (simp add: function_one_closed) + apply (simp add: fun_mult_assoc) + apply (simp add: function_times_one_l) + by (simp add: function_times_one_r) + +lemma(in ring_functions) function_ring_is_ring: +"ring F" + apply(rule ringI) + apply (simp add: function_ring_is_abelian_group) + apply (simp add: function_ring_is_monoid) + apply (simp add: function_mult_r_distr) + by (simp add: function_mult_l_distr) + +sublocale ring_functions < F?: ring F + by (rule function_ring_is_ring) + +lemma(in cring_functions) function_mult_comm: + assumes "x \ carrier F" + assumes" y \ carrier F" + shows "x \\<^bsub>F\<^esub> y = y \\<^bsub>F\<^esub> x" + apply(rule function_ring_car_eqI) + apply (simp add: assms(1) assms(2) fun_mult_closed) + apply (simp add: assms(1) assms(2) fun_mult_closed) + by (simp add: assms(1) assms(2) function_mult_eval_car function_ring_car_closed m_comm) + +lemma(in cring_functions) function_ring_is_comm_monoid: +"comm_monoid F" + apply(rule comm_monoidI) + using fun_mult_assoc function_one_closed + apply (simp add: fun_mult_closed) + apply (simp add: function_one_closed) + apply (simp add: fun_mult_assoc) + apply (simp add: function_times_one_l) + by (simp add: function_mult_comm) + +lemma(in cring_functions) function_ring_is_cring: +"cring F" + apply(rule cringI) + apply (simp add: function_ring_is_abelian_group) + apply (simp add: function_ring_is_comm_monoid) + by (simp add: function_mult_r_distr) + +lemma(in cring_functions) function_ring_is_algebra: +"algebra R F" + apply(rule algebraI) + apply (simp add: is_cring) + apply (simp add: function_ring_is_cring) + using function_smult_closed apply blast + apply (simp add: function_smult_l_distr) + apply (simp add: function_smult_r_distr) + apply (simp add: function_smult_assoc1) + apply (simp add: function_smult_one) + by (simp add: function_smult_assoc2) + +lemma(in ring_functions) function_uminus: + assumes "f \ carrier F" + shows "\\<^bsub>F\<^esub> f = (function_uminus S R) f" + using assms a_inv_def[of F] + by (metis F_def abelian_group.a_group abelian_group.r_neg function_uminus_add_r function_uminus_closed group.inv_closed partial_object.select_convs(1) ring.ring_simprules(18) ring_functions.function_ring_car_eqI ring_functions.function_ring_is_abelian_group ring_functions.function_ring_is_ring ring_functions_axioms) + +lemma(in ring_functions) function_uminus_eval': + assumes "f \ carrier F" + assumes "a \ S" + shows "(\\<^bsub>F\<^esub> f) a = (function_uminus S R) f a" + using assms + by (simp add: function_uminus) + +lemma(in ring_functions) function_uminus_eval'': + assumes "f \ carrier F" + assumes "a \ S" + shows "(\\<^bsub>F\<^esub> f) a = \ (f a)" + using assms(1) assms(2) function_uminus + by (simp add: function_uminus_eval) + +sublocale cring_functions < F?: algebra R F + using function_ring_is_algebra by auto + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Constant Functions\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition constant_function where +"constant_function S a =(\x \ S. a)" + +abbreviation(in ring_functions)(input) const where +"const \ constant_function S" + +lemma(in ring_functions) constant_function_closed: + assumes "a \ carrier R" + shows "const a \ carrier F" + apply(rule function_ring_car_memI) + unfolding constant_function_def + apply (simp add: assms) + by simp + +lemma(in ring_functions) constant_functionE: + assumes "a \ carrier R" + assumes "b \ S" + shows "const a b = a" + by (simp add: assms(2) constant_function_def) + +lemma(in ring_functions) constant_function_add: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "const (a \\<^bsub>R\<^esub> b) = (const a) \\<^bsub>F\<^esub> (const b) " + apply(rule function_ring_car_eqI) + apply (simp add: constant_function_closed assms(1) assms(2)) + using assms(1) constant_function_closed assms(2) fun_add_closed apply auto[1] + by (simp add: assms(1) assms(2) constant_function_closed constant_functionE fun_add_eval_car) + +lemma(in ring_functions) constant_function_mult: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "const (a \\<^bsub>R\<^esub> b) = (const a) \\<^bsub>F\<^esub> (const b)" + apply(rule function_ring_car_eqI) + apply (simp add: constant_function_closed assms(1) assms(2)) + using assms(1) constant_function_closed assms(2) fun_mult_closed apply auto[1] + by (simp add: constant_function_closed assms(1) assms(2) constant_functionE function_mult_eval_car) + +lemma(in ring_functions) constant_function_minus: + assumes "a \ carrier R" + shows "\\<^bsub>F\<^esub>(const a) = (const (\\<^bsub>R\<^esub> a)) " +apply(rule function_ring_car_eqI) + apply (simp add: constant_function_closed assms local.function_uminus) + apply (simp add: constant_function_closed assms function_uminus_closed) + apply (simp add: constant_function_closed assms) + by (simp add: constant_function_closed assms constant_functionE function_uminus_eval'') + +lemma(in ring_functions) function_one_is_constant: +"const \ = \\<^bsub>F\<^esub>" + unfolding F_def + apply(rule function_ring_car_eqI) + apply (simp add: constant_function_closed) + using F_def function_one_closed apply auto[1] + using F_def constant_functionE function_one_eval by auto + +lemma(in ring_functions) function_zero_is_constant: +"const \ = \\<^bsub>F\<^esub>" + apply(rule function_ring_car_eqI) + apply (simp add: constant_function_closed) + using F_def function_zero_closed apply auto[1] + using F_def constant_functionE function_zero_eval by auto + + + (**************************************************************************************************) + (**************************************************************************************************) + subsection\Special Examples of Functions Rings\ + (**************************************************************************************************) + (**************************************************************************************************) + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Functions from the Carrier of a Ring to Itself\ + (**************************************************************************************************) + (**************************************************************************************************) + + +locale U_function_ring = ring + +locale U_function_cring = U_function_ring + cring + +sublocale U_function_ring < S?: struct_functions R "carrier R" + done + +sublocale U_function_ring < FunR?: ring_functions R "carrier R" "Fun R" + apply (simp add: local.ring_axioms ring_functions.intro) + by simp + +sublocale U_function_cring < FunR?: cring_functions R "carrier R" "Fun R" + apply (simp add: cring_functions_def is_cring ring_functions_axioms) + by simp + +abbreviation(in U_function_ring)(input) ring_compose :: "('a \ 'a) \ ('a \ 'a) \ ('a \ 'a)" where +"ring_compose \ compose (carrier R)" + +lemma(in U_function_ring) ring_function_ring_comp: + assumes "f \ carrier (Fun R)" + assumes "g \ carrier (Fun R)" + shows "ring_compose f g \ carrier (Fun R)" + apply(rule function_ring_car_memI) + apply (simp add: assms(1) assms(2) compose_eq) + apply (simp add: assms(1) assms(2) function_ring_car_closed) + by (meson compose_extensional extensional_arb) + +abbreviation(in U_function_ring)(input) ring_const ("\\") where +"ring_const \ constant_function (carrier R)" + +lemma(in ring_functions) function_nat_pow_eval: + assumes "f \ carrier F" + assumes "s \ S" + shows "(f[^]\<^bsub>F\<^esub>(n::nat)) s = (f s)[^]n" + apply(induction n) + using assms(2) function_one_eval apply auto[1] + by (simp add: assms(1) assms(2) function_mult_eval_car function_ring_is_monoid monoid.nat_pow_closed) + + +context U_function_ring +begin + +definition a_translate :: "'a \ 'a \ 'a" where +"a_translate = (\ r \ carrier R. restrict ((add R) r) (carrier R))" + +definition m_translate :: "'a \ 'a \ 'a" where +"m_translate = (\ r \ carrier R. restrict ((mult R) r) (carrier R))" + +definition nat_power :: "nat \ 'a \ 'a" where +"nat_power = (\(n::nat). restrict (\a. a[^]\<^bsub>R\<^esub>n) (carrier R)) " + +text\Restricted operations are in Fs\ + +lemma a_translate_functions: + assumes "c \ carrier R" + shows "a_translate c \ carrier (Fun R)" + apply(rule function_ring_car_memI) + using assms a_translate_def + apply simp + using assms a_translate_def + by simp + +lemma m_translate_functions: + assumes "c \ carrier R" + shows "m_translate c \ carrier (Fun R)" + apply(rule function_ring_car_memI) + using assms m_translate_def + apply simp + using assms m_translate_def + by simp + +lemma nat_power_functions: + shows "nat_power n \ carrier (Fun R)" + apply(rule function_ring_car_memI) + using nat_power_def + apply simp + by (simp add: nat_power_def) + +text\Restricted operations simps\ + +lemma a_translate_eq: + assumes "c \ carrier R" + assumes "a \ carrier R" + shows "a_translate c a = c \ a" + by (simp add: a_translate_def assms(1) assms(2)) + +lemma a_translate_eq': + assumes "c \ carrier R" + assumes "a \ carrier R" + shows "a_translate c a = undefined" + by (meson a_translate_functions assms(1) assms(2) function_ring_not_car) + +lemma a_translate_eq'': + assumes "c \ carrier R" + shows "a_translate c = undefined" + by (simp add: a_translate_def assms) + +lemma m_translate_eq: + assumes "c \ carrier R" + assumes "a \ carrier R" + shows "m_translate c a = c \ a" + by (simp add: m_translate_def assms(1) assms(2)) + +lemma m_translate_eq': + assumes "c \ carrier R" + assumes "a \ carrier R" + shows "m_translate c a = undefined " + by (meson m_translate_functions assms(1) assms(2) function_ring_not_car) + +lemma m_translate_eq'': + assumes "c \ carrier R" + shows "m_translate c = undefined" + by (simp add: m_translate_def assms) + +lemma nat_power_eq: + assumes "a \ carrier R" + shows "nat_power n a = a[^]\<^bsub>R\<^esub> n" + by (simp add: assms nat_power_def) + +lemma nat_power_eq': + assumes "a \ carrier R" + shows "nat_power n a = undefined" + by (simp add: assms nat_power_def) + +text\Constant ring\_function properties\ + +lemma constant_function_eq: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "\\<^bsub>a\<^esub> b = a" + using assms + + by (simp add: constant_functionE) + +lemma constant_function_eq': + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "\\<^bsub>a\<^esub> b = undefined" + by (simp add: constant_function_closed assms(1) assms(2) function_ring_not_car) + +text\Compound expressions from algebraic operations\ +end + +definition monomial_function where +"monomial_function R c (n::nat) = (\ x \ carrier R. c \\<^bsub>R\<^esub> (x[^]\<^bsub>R\<^esub>n))" + +context U_function_ring +begin + +abbreviation monomial where +"monomial \ monomial_function R" + +lemma monomial_functions: + assumes "c \ carrier R" + shows "monomial c n \ carrier (Fun R)" + apply(rule function_ring_car_memI) + unfolding monomial_function_def + apply (simp add: assms) + by simp + +definition ring_id where +"ring_id \ restrict (\x. x) (carrier R) " + +lemma ring_id_closed[simp]: +"ring_id \ carrier (Fun R)" + by (simp add: function_ring_car_memI ring_id_def) + +lemma ring_id_eval: + assumes "a \ carrier R" + shows "ring_id a = a" + using assms unfolding ring_id_def + by simp + +lemma constant_a_trans: + assumes "a \carrier R" + shows "m_translate a = \\<^bsub>a\<^esub> \\<^bsub>Fun R\<^esub> ring_id" +proof(rule function_ring_car_eqI) + show "m_translate a \ carrier (Fun R)" + using assms + using m_translate_functions by blast + show "\\<^bsub>a\<^esub> \\<^bsub>Fun R\<^esub> ring_id \ carrier (Fun R)" + unfolding ring_id_def + using assms ring_id_closed ring_id_def + by (simp add: constant_function_closed fun_mult_closed) + show "\x. x \ carrier R \ m_translate a x = (\\<^bsub>a\<^esub> \\<^bsub>Fun R\<^esub> ring_id) x" + by (simp add: constant_function_closed assms constant_function_eq function_mult_eval_car m_translate_eq ring_id_eval) +qed + +text\polynomials in one variable\ + +fun polynomial :: "'a list \ ('a \ 'a)" where +"polynomial [] = \\<^bsub>Fun R\<^esub> "| +"polynomial (a#as) = (\x \ carrier R. a \ x \ (polynomial as x))" + +lemma polynomial_induct_lemma: + assumes "f \ carrier (Fun R)" + assumes "a \ carrier R" + shows "(\x \ carrier R. a \ x \ (f x)) \ carrier (Fun R)" +proof(rule function_ring_car_memI) + show "\aa. aa \ carrier R \ (\x\carrier R. a \ x \ f x) aa \ carrier R" + proof- fix y assume A: "y \ carrier R" + have "a \ y \ f y \ carrier R" + using A assms(1) assms(2) function_ring_car_closed by blast + thus "(\x\carrier R. a \ x \ f x) y \ carrier R" + using A by auto + qed + show "\aa. aa \ carrier R \ (\x\carrier R. a \ x \ f x) aa = undefined" + by auto +qed + +lemma polynomial_function: + shows "set as \ carrier R \ polynomial as \ carrier (Fun R)" +proof(induction as) + case Nil + then show ?case + by (simp add: function_zero_closed) +next + case (Cons a as) + then show "polynomial (a # as) \ carrier (function_ring (carrier R) R)" + using polynomial.simps(2)[of a as] polynomial_induct_lemma[of "polynomial as" a] + by simp +qed + +lemma polynomial_constant: + assumes "a \ carrier R" + shows "polynomial [a] = \\<^bsub>a\<^esub>" + apply(rule function_ring_car_eqI) + using assms polynomial_function + apply (metis (full_types) list.distinct(1) list.set_cases set_ConsD subset_code(1)) + apply (simp add: constant_function_closed assms) + using polynomial.simps(2)[of a "[]"] polynomial.simps(1) assms + by (simp add: constant_function_eq function_zero_eval) + + +end + + (**************************************************************************************************) + (**************************************************************************************************) + subsubsection\Sequences Indexed by the Natural Numbers\ + (**************************************************************************************************) + (**************************************************************************************************) + +definition nat_seqs ("_\<^bsup>\\<^esup>")where +"nat_seqs R \ function_ring (UNIV::nat set) R" + +abbreviation(input) closed_seqs where +"closed_seqs R \ carrier (R\<^bsup>\\<^esup>)" + +lemma closed_seqs_memI: + assumes "\k. s k \ carrier R" + shows "s \ closed_seqs R" + unfolding nat_seqs_def function_ring_def + by (simp add: PiE_UNIV_domain assms) + +lemma closed_seqs_memE: + assumes "s \ closed_seqs R" + shows "s k \ carrier R" + using assms unfolding nat_seqs_def function_ring_def + by (simp add: PiE_iff) + +definition is_constant_fun where +"is_constant_fun R f = (\x \ carrier R. f = constant_function (carrier R) R x)" + +definition is_constant_seq where +"is_constant_seq R s = (\x \ carrier R. s = constant_function (UNIV::nat set) x)" + +lemma is_constant_seqI: + fixes a + assumes "s \ closed_seqs R" + assumes "\k. s k = a" + shows "is_constant_seq R s" + unfolding is_constant_seq_def constant_function_def + by (metis assms(1) assms(2) closed_seqs_memE restrict_UNIV restrict_ext) + +lemma is_constant_seqE: + assumes "is_constant_seq R s" + assumes "s k = a" + shows "s n = a" + using assms unfolding is_constant_seq_def + by (metis constant_function_def restrict_UNIV) + +lemma is_constant_seq_imp_closed: + assumes "is_constant_seq R s" + shows "s \ closed_seqs R" + apply(rule closed_seqs_memI) + using assms unfolding is_constant_seq_def constant_function_def + by auto + +context U_function_ring +begin + +text\Sequence sums and products are closed\ + +lemma seq_plus_closed: + assumes "s \ closed_seqs R" + assumes "s' \ closed_seqs R" + shows "s \\<^bsub>R\<^bsup>\\<^esup>\<^esub> s' \ closed_seqs R" + by (metis assms(1) assms(2) nat_seqs_def ring_functions.fun_add_closed ring_functions_axioms) + +lemma seq_mult_closed: + assumes "s \ closed_seqs R" + assumes "s' \ closed_seqs R" + shows "s \\<^bsub>R\<^bsup>\\<^esup>\<^esub> s' \ closed_seqs R" + apply(rule closed_seqs_memI) + by (metis assms(1) assms(2) closed_seqs_memE nat_seqs_def ring_functions.fun_mult_closed ring_functions_axioms) + +lemma constant_function_comp_is_closed_seq: + assumes "a \ carrier R" + assumes "s \ closed_seqs R" + shows "(const a \ s) \ closed_seqs R" + by (simp add: constant_functionE assms(1) assms(2) closed_seqs_memE closed_seqs_memI) + +lemma constant_function_comp_is_constant_seq: + assumes "a \ carrier R" + assumes "s \ closed_seqs R" + shows "is_constant_seq R ((const a) \ s)" + apply(rule is_constant_seqI[of _ _ a] ) + apply (simp add: assms(1) assms(2) constant_function_comp_is_closed_seq) + using assms(1) assms(2) closed_seqs_memE + by (simp add: closed_seqs_memE constant_functionE) + +lemma function_comp_is_closed_seq: + assumes "s \ closed_seqs R" + assumes "f \ carrier (Fun R)" + shows "f \ s \ closed_seqs R" + apply(rule closed_seqs_memI) + using assms(1) assms(2) closed_seqs_memE + by (metis comp_apply fun_add_eval_closed fun_add_zeroR function_zero_closed) + +lemma function_sum_comp_is_seq_sum: + assumes "s \ closed_seqs R" + assumes "f \ carrier (Fun R)" + assumes "g \ carrier (Fun R)" + shows "(f \\<^bsub>Fun R\<^esub> g) \ s = (f \ s) \\<^bsub>R\<^bsup>\\<^esup>\<^esub> (g \ s)" + apply(rule ring_functions.function_ring_car_eqI[of R _ "UNIV :: nat set"]) + apply (simp add: ring_functions_axioms) + using function_comp_is_closed_seq + apply (metis assms(1) assms(2) assms(3) fun_add_closed nat_seqs_def) + apply (metis assms(1) assms(2) assms(3) function_comp_is_closed_seq nat_seqs_def seq_plus_closed) + by (smt UNIV_eq_I assms(1) assms(2) assms(3) closed_seqs_memE comp_apply function_comp_is_closed_seq nat_seqs_def ring_functions.fun_add_eval_car ring_functions_axioms) + +lemma function_mult_comp_is_seq_mult: + assumes "s \ closed_seqs R" + assumes "f \ carrier (Fun R)" + assumes "g \ carrier (Fun R)" + shows "(f \\<^bsub>Fun R\<^esub> g) \ s = (f \ s) \\<^bsub>R\<^bsup>\\<^esup>\<^esub> (g \ s)" + apply(rule ring_functions.function_ring_car_eqI[of R _ "UNIV :: nat set"]) + apply (simp add: ring_functions_axioms) + using function_comp_is_closed_seq + apply (metis assms(1) assms(2) assms(3) fun_mult_closed nat_seqs_def) + apply (metis assms(1) assms(2) assms(3) function_comp_is_closed_seq nat_seqs_def seq_mult_closed) + by (metis (no_types, lifting) assms(1) assms(2) assms(3) comp_apply function_comp_is_closed_seq nat_seqs_def ring_functions.function_mult_eval_car ring_functions.function_ring_car_closed ring_functions_axioms) + +lemma seq_plus_simp: + assumes "s \ closed_seqs R" + assumes "t \ closed_seqs R" + shows "(s \\<^bsub>R\<^bsup>\\<^esup>\<^esub> t) k = s k \ t k" + using assms unfolding nat_seqs_def + by (simp add: ring_functions.fun_add_eval_car ring_functions_axioms) + +lemma seq_mult_simp: + assumes "s \ closed_seqs R" + assumes "t \ closed_seqs R" + shows "(s \\<^bsub>R\<^bsup>\\<^esup>\<^esub> t) k = s k \ t k" + using assms unfolding nat_seqs_def + by (simp add: ring_functions.function_mult_eval_car ring_functions_axioms) + +lemma seq_one_simp: +"\\<^bsub>R\<^bsup>\\<^esup>\<^esub> k = \" + by (simp add: nat_seqs_def ring_functions.function_one_eval ring_functions_axioms) + +lemma seq_zero_simp: +"\\<^bsub>R\<^bsup>\\<^esup>\<^esub> k = \" + by (simp add: nat_seqs_def ring_functions.function_zero_eval ring_functions_axioms) + +lemma(in U_function_ring) ring_id_seq_comp: + assumes "s \ closed_seqs R" + shows "ring_id \ s = s" + apply(rule ring_functions.function_ring_car_eqI[of R _ "UNIV::nat set"]) + using ring_functions_axioms apply auto[1] + apply (metis assms function_comp_is_closed_seq nat_seqs_def ring_id_closed) + apply (metis assms nat_seqs_def) + by (simp add: assms closed_seqs_memE ring_id_eval) + +lemma(in U_function_ring) ring_seq_smult_closed: + assumes "s \ closed_seqs R" + assumes "a \ carrier R" + shows "a \\<^bsub>R\<^bsup>\\<^esup>\<^esub> s \ closed_seqs R" + apply(rule closed_seqs_memI) + by (metis assms(1) assms(2) closed_seqs_memE nat_seqs_def ring_functions.function_smult_closed ring_functions_axioms) + +lemma(in U_function_ring) ring_seq_smult_eval: + assumes "s \ closed_seqs R" + assumes "a \ carrier R" + shows "(a \\<^bsub>R\<^bsup>\\<^esup>\<^esub> s) k = a \ (s k)" + by (metis UNIV_I assms(1) assms(2) nat_seqs_def ring_functions.function_smult_eval ring_functions_axioms) + +lemma(in U_function_ring) ring_seq_smult_comp_assoc: + assumes "s \ closed_seqs R" + assumes "f \ carrier (Fun R)" + assumes "a \ carrier R" + shows "((a \\<^bsub>Fun R\<^esub> f) \ s) = a \\<^bsub>R\<^bsup>\\<^esup>\<^esub> (f \ s)" + apply(rule ext) + using function_smult_eval[of a f] ring_seq_smult_eval[of "f \ s" a] + by (simp add: assms(1) assms(2) assms(3) closed_seqs_memE function_comp_is_closed_seq) + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Extensional Maps Between the Carriers of two Structures\ +(**************************************************************************************************) +(**************************************************************************************************) + +definition struct_maps :: "('a, 'c) partial_object_scheme \ ('b, 'd) partial_object_scheme + \ ('a \ 'b) set" where +"struct_maps T S = {f. (f \ (carrier T) \ (carrier S)) \ f = restrict f (carrier T) }" + +definition to_struct_map where +"to_struct_map T f = restrict f (carrier T)" + +lemma to_struct_map_closed: + assumes "f \ (carrier T) \ (carrier S)" + shows "to_struct_map T f \ (struct_maps T S)" + by (smt PiE_restrict Pi_iff assms mem_Collect_eq restrict_PiE struct_maps_def to_struct_map_def) + +lemma struct_maps_memI: + assumes "\ x. x \ carrier T \ f x \ carrier S" + assumes "\x. x \ carrier T \ f x = undefined" + shows "f \ struct_maps T S" +proof- + have 0: " (f \ (carrier T) \ (carrier S))" + using assms + by blast + have 1: "f = restrict f (carrier T)" + using assms + by (simp add: extensional_def extensional_restrict) + show ?thesis + using 0 1 + unfolding struct_maps_def + by blast +qed + +lemma struct_maps_memE: + assumes "f \ struct_maps T S" + shows "\ x. x \ carrier T \ f x \ carrier S" + "\x. x \ carrier T \ f x = undefined" + using assms unfolding struct_maps_def + apply blast + using assms unfolding struct_maps_def + by (metis (mono_tags, lifting) mem_Collect_eq restrict_apply) + +text\An abbreviation for restricted composition of function of functions. This is necessary for the composition of two struct maps to again be a struct map.\ +abbreviation(input) rcomp + where "rcomp \ FuncSet.compose" + +lemma struct_map_comp: + assumes "g \ (struct_maps T S)" + assumes "f \ (struct_maps S U)" + shows "rcomp (carrier T) f g \ (struct_maps T U)" +proof(rule struct_maps_memI) + show "\x. x \ carrier T \ rcomp (carrier T) f g x \ carrier U" + using assms struct_maps_memE(1) + by (metis compose_eq) + show " \x. x \ carrier T \ rcomp (carrier T) f g x = undefined" + by (meson compose_extensional extensional_arb) +qed + +lemma r_comp_is_compose: + assumes "g \ (struct_maps T S)" + assumes "f \ (struct_maps S U)" + assumes "a \ (carrier T)" + shows "(rcomp (carrier T) f g) a = (f \ g) a" + by (simp add: FuncSet.compose_def assms(3)) + +lemma r_comp_not_in_car: + assumes "g \ (struct_maps T S)" + assumes "f \ (struct_maps S U)" + assumes "a \ (carrier T)" + shows "(rcomp (carrier T) f g) a = undefined" + by (simp add: FuncSet.compose_def assms(3)) + +text\The reverse composition of two struct maps:\ + +definition pullback :: + "('a, 'd) partial_object_scheme \ ('a \ 'b) \ ('b \ 'c) \ ('a \ 'c)" where +"pullback T f g = rcomp (carrier T) g f" + +lemma pullback_closed: + assumes "f \ (struct_maps T S)" + assumes "g \ (struct_maps S U)" + shows "pullback T f g \ (struct_maps T U)" + by (metis assms(1) assms(2) pullback_def struct_map_comp) + +text\Composition of struct maps which takes the structure itself rather than the carrier as a parameter:\ + +definition pushforward :: + "('a, 'd) partial_object_scheme \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" where +"pushforward T f g \ rcomp (carrier T) f g" + +lemma pushforward_closed: + assumes "g \ (struct_maps T S)" + assumes "f \ (struct_maps S U)" + shows "pushforward T f g \ (struct_maps T U)" + using assms(1) assms(2) struct_map_comp + by (metis pushforward_def) + + +end + diff --git a/thys/Padic_Ints/Hensels_Lemma.thy b/thys/Padic_Ints/Hensels_Lemma.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Hensels_Lemma.thy @@ -0,0 +1,1987 @@ +theory Hensels_Lemma + imports Padic_Int_Polynomials +begin + + +text\ + The following proof of Hensel's Lemma is directly adapted from Keith Conrad's proof which is + given in an online note \cite{keithconrad}. The same note was used as the basis for a + formalization of Hensel's Lemma by Robert Lewis in the Lean proof assistant + \cite{10.1145/3293880.3294089}. \ +(**************************************************************************************************) +(**************************************************************************************************) +section\Auxiliary Lemmas for Hensel's Lemma\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma(in ring) minus_sum: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "\ (a \ b) = \ a \ \ b" + by (simp add: assms(1) assms(2) local.minus_add) + +context padic_integers +begin + + +lemma poly_diff_val: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + shows "val_Zp (f\a \ f\b) \ val_Zp (a \ b)" +proof- + obtain c where c_def: "c \ carrier Zp \ (f\a \ f\b) = (a \ b) \ c" + using assms + by (meson to_fun_diff_factor) + have 1: "val_Zp c \ 0" + using c_def val_pos by blast + have 2: "val_Zp (f\a \ f\b) = val_Zp (a \ b) + (val_Zp c)" + using c_def val_Zp_mult + by (simp add: assms(2) assms(3)) + then show ?thesis + using "1" by auto +qed + +text\Restricted p-adic division\ + +definition divide where +"divide x y = (if x = \ then \ else + (\

[^](nat (ord_Zp x - ord_Zp y)) \ ac_Zp x \ (inv ac_Zp y)))" + +lemma divide_closed: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "y \ \" + shows "divide x y \ carrier Zp" + unfolding divide_def + apply(cases "x = \") + apply auto[1] + using assms ac_Zp_is_Unit + by (simp add: ac_Zp_in_Zp) + +lemma divide_formula: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "y \ \" + assumes "val_Zp x \ val_Zp y" + shows "y \ divide x y = x" + apply(cases "x = \") + apply (simp add: divide_def mult_zero_l) +proof- assume A: "x \ \" + have 0: "y \ divide x y = \

[^] nat (ord_Zp y) \ ac_Zp y \ (\

[^](nat (ord_Zp x - ord_Zp y)) \ ac_Zp x \ (inv ac_Zp y))" + using assms ac_Zp_factors_x[of x] ac_Zp_factors_x[of y] A divide_def + by auto + hence 1: "y \ divide x y = \

[^] nat (ord_Zp y) \ (\

[^](nat (ord_Zp x - ord_Zp y)) \ ac_Zp x \ ac_Zp y \ (inv ac_Zp y))" + using mult_assoc mult_comm by auto + have 2: "(nat (ord_Zp y) + nat (ord_Zp x - ord_Zp y)) = nat (ord_Zp x)" + using assms ord_pos[of x] ord_pos[of y] A val_ord_Zp by auto + have "y \ divide x y = \

[^] nat (ord_Zp y) \ \

[^](nat (ord_Zp x - ord_Zp y)) \ ac_Zp x" + using 1 A assms + by (simp add: ac_Zp_in_Zp ac_Zp_is_Unit mult_assoc) + thus "y \ divide x y = x" + using "2" A ac_Zp_factors_x(1) assms(1) p_natpow_prod by auto +qed + +lemma divide_nonzero: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "val_Zp x \ val_Zp y" + shows "divide x y \ nonzero Zp" + by (metis assms(1) assms(2) assms(3) divide_closed divide_formula mult_zero_l nonzero_closed nonzero_memE(2) nonzero_memI) + +lemma val_of_divide: + assumes "x \ carrier Zp" + assumes "y \ nonzero Zp" + assumes "val_Zp x \ val_Zp y" + shows "val_Zp (divide x y) = val_Zp x - val_Zp y" +proof- + have 0: "y \ divide x y = x" + by (simp add: assms(1) assms(2) assms(3) divide_formula nonzero_closed nonzero_memE(2)) + hence "val_Zp y + val_Zp (divide x y) = val_Zp x" + using assms(1) assms(2) divide_closed nonzero_closed not_nonzero_memI val_Zp_mult by fastforce + thus ?thesis + by (smt Zp_def add.commute add.left_neutral add.right_neutral add_diff_assoc_eint assms(1) + assms(2) divide_nonzero eSuc_minus_eSuc iadd_Suc idiff_0_right mult_zero(1) mult_zero_l + nonzero_closed ord_pos order_refl padic_integers.Zp_int_inc_closed padic_integers.mult_comm + padic_integers.ord_of_nonzero(2) padic_integers_axioms val_Zp_eq_frac_0 val_Zp_mult val_Zp_p) +qed + +lemma val_of_divide': + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "y \ \" + assumes "val_Zp x \ val_Zp y" + shows "val_Zp (divide x y) = val_Zp x - val_Zp y" + using Zp_def assms(1) assms(2) assms(3) assms(4) padic_integers.not_nonzero_Zp + padic_integers.val_of_divide padic_integers_axioms by blast +end + +lemma(in UP_cring) taylor_deg_1_eval''': + assumes "f \ carrier P" + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c = to_fun (shift (2::nat) (T\<^bsub>a\<^esub> f)) (\b)" + assumes "b \ (deriv f a) = (to_fun f a)" + shows "to_fun f (a \ b) = (c \ b[^](2::nat))" +proof- + have 0: "to_fun f (a \ b) = (to_fun f a) \ (deriv f a \ b) \ (c \ b[^](2::nat))" + using assms taylor_deg_1_eval'' + by blast + have 1: "(to_fun f a) \ (deriv f a \ b) = \" + using assms + proof - + have "\f a. f \ carrier P \ a \ carrier R \ to_fun f a \ carrier R" + using to_fun_closed by presburger + then show ?thesis + using R.m_comm R.r_right_minus_eq assms(1) assms(2) assms(3) assms(5) + by (simp add: deriv_closed) + qed + have 2: "to_fun f (a \ b) = \ \ (c \ b[^](2::nat))" + using 0 1 + by simp + then show ?thesis using assms + by (simp add: taylor_closed to_fun_closed shift_closed) +qed + +lemma(in padic_integers) res_diff_zero_fact: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "(a \ b) k = 0" + shows "a k = b k" "a k \\<^bsub>Zp_res_ring k\<^esub> b k = 0" + apply(cases "k = 0") + apply (metis assms(1) assms(2) p_res_ring_0 p_res_ring_0' p_res_ring_car p_residue_padic_int p_residue_range' zero_le) + apply (metis R.add.inv_closed R.add.m_lcomm R.minus_eq R.r_neg R.r_zero Zp_residue_add_zero(2) assms(1) assms(2) assms(3)) + using assms(2) assms(3) residue_of_diff by auto + +lemma(in padic_integers) res_diff_zero_fact': + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a k = b k" + shows "a k \\<^bsub>Zp_res_ring k\<^esub> b k = 0" + by (simp add: assms(3) residue_minus) + +lemma(in padic_integers) res_diff_zero_fact'': + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a k = b k" + shows "(a \ b) k = 0" + by (simp add: assms(2) assms(3) res_diff_zero_fact' residue_of_diff) + +lemma(in padic_integers) is_Zp_cauchyI': +assumes "s \ closed_seqs Zp" +assumes "\n::nat. \ k::int.\m. m \ k \ val_Zp (s (Suc m) \ s m) \ n" +shows "is_Zp_cauchy s" +proof(rule is_Zp_cauchyI) + show A0: "s \ closed_seqs Zp" + by (simp add: assms(1)) + show "\n. \N. \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n" + proof- + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n" + proof(induction n) + case 0 + then show ?case + proof- + have "\n0 n1. 0 < n0 \ 0 < n1 \ s n0 0 = s n1 0" + apply auto + proof- + fix n0 n1::nat + assume A: "n0 > 0" "n1 > 0" + have 0: "s n0 \ carrier Zp" + using A0 + by (simp add: closed_seqs_memE) + have 1: "s n1 \ carrier Zp" + using A0 + by (simp add: closed_seqs_memE) + show " s n0 (0::nat) = s n1 (0::nat)" + using A0 Zp_def 0 1 residues_closed + by (metis p_res_ring_0') + qed + then show ?thesis + by blast + qed + next + case (Suc n) + fix n + assume IH: "\N. \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n" + show " \N. \n0 n1. N < n0 \ N < n1 \ s n0 (Suc n) = s n1 (Suc n)" + proof- + obtain N where N_def: "\n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n" + using IH + by blast + obtain k where k_def: "\m. (Suc m) \ k \ val_Zp (s (Suc (Suc m)) \ s (Suc m)) \ Suc (Suc n)" + using assms Suc_n_not_le_n + by (meson nat_le_iff) + have "\n0 n1. Suc (max N (max n k)) < n0 \ Suc (max N (max n k))< n1 \ s n0 (Suc n) = s n1 (Suc n)" + apply auto + proof- + fix n0 n1 + assume A: "Suc (max N (max n k)) < n0" " Suc (max N (max n k)) < n1" + show "s n0 (Suc n) = s n1 (Suc n) " + proof- + obtain K where K_def: "K = Suc (max N (max n k))" + by simp + have P0: "\m. s ((Suc m)+ K) (Suc n) = s (Suc K) (Suc n)" + apply auto + proof- + fix m + show "s (Suc (m + K)) (Suc n) = s (Suc K) (Suc n)" + apply(induction m) + apply auto + proof- + fix m + assume A0: " s (Suc (m + K)) (Suc n) = s (Suc K) (Suc n)" + show " s (Suc (Suc (m + K))) (Suc n) = s (Suc K) (Suc n)" + proof- + have I: "k < m + K" + using K_def + by linarith + have "val_Zp (s (Suc (Suc (m + K))) \ s (Suc (m + K))) \ Suc (Suc n)" + proof- + have "(Suc (m + K)) > k" + by (simp add: I less_Suc_eq) + then show ?thesis + using k_def less_imp_le_nat + by blast + qed + hence D: "val_Zp (s (Suc (Suc (m + K))) \ s (Suc (m + K))) > (Suc n)" + using Suc_ile_eq by fastforce + have "s (Suc (Suc (m + K))) (Suc n) = s (Suc (m + K)) (Suc n)" + proof- + have "(s (Suc (Suc (m + K))) \ s (Suc (m + K))) (Suc n) = 0" + using D assms(1) res_diff_zero_fact''[of "s (Suc (Suc (m + K)))" "s (Suc (m + K)) " "Suc n"] + val_Zp_dist_res_eq[of "s (Suc (Suc (m + K)))" "s (Suc (m + K)) " "Suc n"] unfolding val_Zp_dist_def + by (simp add: closed_seqs_memE) + hence 0: "(s (Suc (Suc (m + K))) (Suc n) \\<^bsub>Zp_res_ring (Suc n)\<^esub> (s (Suc (m + K))) (Suc n)) = 0" + using res_diff_zero_fact(2)[of "s (Suc (Suc (m + K)))" "s (Suc (m + K))" "Suc n" ] + assms(1) + by (simp add: closed_seqs_memE) + + show ?thesis + proof- + have 00: "cring (Zp_res_ring (Suc n))" + using R_cring by blast + have 01: " s (Suc (Suc (m + K))) (Suc n) \ carrier (Zp_res_ring (Suc n))" + using assms(1) closed_seqs_memE residues_closed by blast + have 02: "(\\<^bsub>Zp_res_ring (Suc n)\<^esub> (s (Suc (m + K)) (Suc n))) \ carrier (Zp_res_ring (Suc n)) " + by (meson "00" assms(1) cring.cring_simprules(3) closed_seqs_memE residues_closed) + show ?thesis + unfolding a_minus_def + using 00 01 02 + cring.sum_zero_eq_neg[of "Zp_res_ring (Suc n)" "s (Suc (Suc (m + K))) (Suc n)" + "\\<^bsub>Zp_res_ring (Suc n)\<^esub>s (Suc (m + K)) (Suc n)"] + by (metis 0 a_minus_def assms(1) cring.cring_simprules(21) closed_seqs_memE + p_res_ring_zero residues_closed) + qed + qed + then show ?thesis using A0 assms(1) + by simp + qed + qed + qed + have "\m0. n0 = (Suc m0) + K" + proof- + have "n0 > K" + by (simp add: A(1) K_def) + then have "n0 = (Suc (n0 - K - 1)) + K" + by auto + then show ?thesis by blast + qed + then obtain m0 where m0_def: "n0 = (Suc m0) + K" + by blast + have "\m0. n1 = (Suc m0) + K" + proof- + have "n1 > K" + by (simp add: A(2) K_def) + then have "n1 = (Suc (n1 - K - 1)) + K" + by auto + then show ?thesis by blast + qed + then obtain m1 where m1_def: "n1 = (Suc m1) + K" + by blast + have 0: "s n0 (Suc n) = s (Suc K) (Suc n)" + using m0_def P0[of "m0"] by auto + have 1: "s n1 (Suc n) = s (Suc K) (Suc n)" + using m1_def P0[of "m1"] by auto + show ?thesis using 0 1 + by auto + qed + qed + then show ?thesis + by blast + qed + qed + qed +qed + +(**************************************************************************************************) +(**************************************************************************************************) +section\The Proof of Hensel's Lemma\ +(**************************************************************************************************) +(**************************************************************************************************) + +subsection\Building a Locale for the Proof of Hensel's Lemma\ + +locale hensel = padic_integers+ + fixes f::padic_int_poly + fixes a::padic_int + assumes f_closed[simp]: "f \ carrier Zp_x" + assumes a_closed[simp]: "a \ carrier Zp" + assumes fa_nonzero[simp]: "f\a \\" + assumes hensel_hypothesis[simp]: "(val_Zp (f\a) > 2* val_Zp ((pderiv f)\a))" + +sublocale hensel < cring Zp + by (simp add: R.is_cring) + +context hensel +begin + +abbreviation f' where +"f' \ pderiv f" + +lemma f'_closed: +"f' \ carrier Zp_x" + using f_closed pderiv_closed by blast + +lemma f'_vals_closed: + assumes "a \ carrier Zp" + shows "f'\a \ carrier Zp" + by (simp add: UP_cring.to_fun_closed Zp_x_is_UP_cring f'_closed) + +lemma fa_closed: +"(f\a) \ carrier Zp" + by (simp add: UP_cring.to_fun_closed Zp_x_is_UP_cring) + +lemma f'a_closed: +"(f'\a) \ carrier Zp" +proof- + have "f' \ carrier Zp_x" + by (simp add: f'_closed) + then show ?thesis + by (simp add: f'_vals_closed) +qed + +lemma fa_nonzero': +"(f\a) \ nonzero Zp" + using fa_closed fa_nonzero not_nonzero_Zp by blast + +lemma f'a_nonzero[simp]: +"(f'\a) \ \" +proof(rule ccontr) + assume "\ (f'\a) \ \" + then have "(f'\a) = \" + by blast + then have "\ < val_Zp (f\a)" using hensel_hypothesis + by (simp add: val_Zp_def) + thus False + using eint_ord_simps(6) by blast +qed + +lemma f'a_nonzero': +"(f'\a) \ nonzero Zp" + using f'a_closed f'a_nonzero not_nonzero_Zp by blast + +lemma f'a_not_infinite[simp]: +"val_Zp (f'\a) \ \" + by (metis eint_ord_code(3) hensel_hypothesis linorder_not_less times_eint_simps(4)) + +lemma f'a_nonneg_val[simp]: +"val_Zp ((f'\a)) \ 0" + using f'a_closed val_pos by blast + +lemma hensel_hypothesis_weakened: +"val_Zp (f\a) > val_Zp (f'\a)" +proof- + have 0: "0 \ val_Zp (f'\a) \ val_Zp (f'\a) \ \" + using f'a_closed val_ord_Zp val_pos by force + have 1: "1 < eint 2 " + by (simp add: one_eint_def) + thus ?thesis using 0 eint_mult_mono'[of "val_Zp (f'\a)" 1 2] hensel_hypothesis + by (metis linorder_not_less mult_one_left order_trans) +qed + +subsection\Constructing the Newton Sequence\ + +definition newton_step :: "padic_int \ padic_int" where +"newton_step x = x \ (divide (f\x) (f'\x))" + +lemma newton_step_closed: + "newton_step a \ carrier Zp" + using divide_closed unfolding newton_step_def + using f'a_closed f'a_nonzero fa_closed local.a_closed by blast + +fun newton_seq :: "padic_int_seq" ("ns") where +"newton_seq 0 = a"| +"newton_seq (Suc n) = newton_step (newton_seq n)" + +subsection\Key Properties of the Newton Sequence\ + +lemma hensel_factor_id: +"(divide (f\a) (f'\a)) \ ((f'\a)) = (f\a)" + using hensel_hypothesis hensel_axioms divide_formula f'a_closed + fa_closed hensel_hypothesis_weakened mult_comm + by auto + +definition hensel_factor ("t") where +"hensel_factor = val_Zp (f\a) - 2*(val_Zp (f'\a))" + +lemma t_pos[simp]: +"t > 0" + using hensel_factor_def hensel_hypothesis + by (simp add: eint_minus_le) + +lemma t_neq_infty[simp]: +"t \ \" + by (simp add: hensel_factor_def val_Zp_def) + +lemma t_times_pow_pos[simp]: +"(2^(n::nat))*t > 0" + apply(cases "n = 0") + using one_eint_def apply auto[1] + using eint_mult_mono'[of t 1 "2^n"] t_pos + by (smt eint_ord_simps(2) linorder_not_less mult_one_left neq0_conv one_eint_def order_less_le order_trans self_le_power t_neq_infty) + +lemma newton_seq_props_induct: +shows "\k. k \ n \ (ns k) \ carrier Zp + \ val_Zp (f'\(ns k)) = val_Zp ((f'\a)) + \ val_Zp (f\(ns k)) \ 2*(val_Zp (f'\a)) + (2^k)*t" +proof(induction n) + case 0 + then have kz: "k = 0" + by simp + have B0: "( ns k) \ carrier Zp" + using kz + by simp + have B1: "val_Zp (f' \ ns k) = (val_Zp (f'\a))" + using kz newton_seq.simps(1) + by presburger + have B2: "val_Zp (f \ (ns k)) \ (2 * (val_Zp (f'\a))) + 2 ^ k * t" + proof- + have B20: "(2 * (val_Zp (f'\a))) + 2 ^ k * t = (2 * (val_Zp (f'\a))) + t" + proof- + have "(2 * (val_Zp (f'\a))) + 2 ^ k * t = (2 * (val_Zp (f'\a))) + t" + using kz one_eint_def by auto + then show ?thesis + by blast + qed + then have "(2 * (val_Zp (f'\a))) + 2 ^ k * t = (2 * (val_Zp (f'\a))) + val_Zp (f\a) - 2*(val_Zp (f'\a))" + unfolding hensel_factor_def + by (simp add: val_Zp_def) + then have "(2 * (val_Zp (f'\a))) + 2 ^ k * t = val_Zp (f\a)" + by (metis add_diff_cancel_eint eint_ord_simps(6) hensel_hypothesis) + thus ?thesis by (simp add: kz) + qed + thus ?case + using B0 B1 by blast +next + case (Suc n) + show ?case + proof(cases "k \ n") + case True + then show ?thesis using Suc.IH + by blast + next + case False + have F1: "(ns n) \ carrier Zp" + using Suc.IH by blast + have F2: "val_Zp (f'\(ns n)) = val_Zp ((f'\a))" + using Suc.IH by blast + have F3: "val_Zp (f\(ns n)) \ 2*(val_Zp (f'\a)) + (2^n)*t" + using Suc.IH by blast + have kval: "k = Suc n" + using False Suc.prems le_Suc_eq by blast + have F6: "val_Zp (f\(ns n)) \ val_Zp (f'\(ns n))" + proof- + have "2*(val_Zp (f'\a)) \ val_Zp (f'\a)" + using f'a_closed val_pos eint_mult_mono'[of "val_Zp (f'\a)" 1 2] + by (metis Groups.add_ac(2) add.right_neutral eSuc_eint eint_0_iff(2) eint_add_left_cancel_le + eint_ord_simps(2) f'a_nonneg_val f'a_not_infinite infinity_ne_i1 linorder_not_less + mult_one_left not_one_less_zero one_add_one one_eint_def order_less_le order_trans zero_one_eint_neq(1)) + hence "2*(val_Zp (f'\a)) + (2^n)*t \ val_Zp (f'\a)" + using t_times_pow_pos[of n] + by (metis (no_types, lifting) add.right_neutral eint_add_left_cancel_le order_less_le order_trans) + then show ?thesis + using F2 F3 by auto + qed + have F5: " divide (f\(ns n))(f'\(ns n)) \ carrier Zp" + proof- + have 00: "f \ ns n \ carrier Zp" + by (simp add: F1 to_fun_closed) + have "val_Zp ((f'\a)) \ val_Zp \" + by (simp add: val_Zp_def) + then have 01: "f' \ ns n \ nonzero Zp" + using F2 F1 Zp_x_is_UP_cring f'_closed nonzero_def + proof - + have "f' \ ns n \ carrier Zp" + using F1 Zp_continuous_is_Zp_closed f'_closed polynomial_is_Zp_continuous + by (simp add: to_fun_closed) + then show ?thesis + using F2 \val_Zp (f'\a) \ val_Zp \\ not_nonzero_Zp by fastforce + qed + then show ?thesis + using F6 + by (metis "00" F2 \val_Zp (f'\a) \ val_Zp \\ divide_closed nonzero_closed) + qed + have F4: "(ns k) \ (ns n) = (\ divide (f\(ns n))(f'\(ns n)))" + using F1 F5 newton_seq.simps(2)[of n] kval + unfolding newton_step_def + by (metis R.l_neg R.minus_closed R.minus_zero R.plus_diff_simp R.r_neg2 R.r_right_minus_eq + a_minus_def local.a_closed minus_a_inv) + have F7: "val_Zp (divide (f\(ns n))(f'\(ns n))) = val_Zp (f\(ns n)) - val_Zp (f'\(ns n))" + apply(rule val_of_divide) + apply (simp add: F1 to_fun_closed) + using F1 f'_closed to_fun_closed F2 not_nonzero_Zp val_Zp_def apply fastforce + by (simp add: F6) + show ?thesis + proof + show P0:"ns k \ carrier Zp" + proof- + have A0: "ns k = ns n \ (divide (f\ (ns n)) (f'\(ns n)))" + by (simp add: kval newton_step_def) + have A1: "val_Zp (f'\(ns n)) = val_Zp (f'\a)" + using Suc.IH + by blast + have A2: "val_Zp (f\(ns n)) \val_Zp (f'\a)" + proof- + have A20: "(2 * val_Zp (f'\a)) + 2 ^ n * (val_Zp (f\a) - 2 * val_Zp (f'\a)) \val_Zp (f'\a)" + proof- + have "val_Zp (f\a) - 2 * val_Zp (f'\a) > 0" + using hensel_hypothesis eint_minus_le by blast + then have " (2 ^ n) * (val_Zp (f\a) - 2 * val_Zp (f'\a)) + \ (val_Zp (f\a) - 2 * val_Zp (f'\a))" + using eint_pos_int_times_ge by auto + then have " ((2 * val_Zp (f'\a)) + 2 ^ n * (val_Zp (f\a) - 2 * val_Zp (f'\a))) + \ (2 * val_Zp (f'\a)) + (val_Zp (f\a) - 2 * val_Zp (f'\a))" + by (simp add: val_Zp_def) + then have " ((2 * val_Zp (f'\a)) + 2 ^ n * (val_Zp (f\a) - 2 * val_Zp (f'\a))) + \ (val_Zp (f\a) )" + by simp + then show " ((2 * val_Zp (f'\a)) + 2 ^ n * (val_Zp (f\a) - 2 * val_Zp (f'\a))) + \ (val_Zp (f'\a) )" + using hensel_hypothesis_weakened by auto + qed + have A21:"val_Zp (f\(ns n)) \ (2 * val_Zp (f'\a)) + 2 ^ n * (val_Zp (f\a) - 2 * val_Zp (f'\a))" + using Suc.IH unfolding hensel_factor_def + by blast + show ?thesis using A21 A20 + by auto + qed + have A3: "ns n \ carrier Zp" + using Suc.IH by blast + have A4: "val_Zp (f\(ns n)) \val_Zp (f'\(ns n))" + using A1 A2 + by presburger + have A5: "f\(ns n) \ carrier Zp" + by (simp add: F1 UP_cring.to_fun_closed Zp_x_is_UP_cring) + have A6: "(f'\(ns n)) \ nonzero Zp" + proof- + have "(f'\(ns n)) \ carrier Zp" + by (simp add: F1 UP_cring.to_fun_closed Zp_x_is_UP_cring f'_closed) + have "val_Zp (f'\(ns n)) \ \" + using A1 + by (simp add: val_Zp_def) + then show ?thesis + using \f' \ ns n \ carrier Zp\ not_nonzero_Zp val_Zp_def + by meson + qed + have A7: " (divide (f\ (ns n)) (f'\(ns n))) \ carrier Zp" + using A5 A6 A4 A3 F5 by linarith + then show ?thesis + using A0 A3 cring.cring_simprules(4) + by (simp add: F1 F5 cring.cring_simprules(4)) + qed + have P1: "val_Zp (f' \ ns k) = val_Zp (f'\a) " + proof(cases "(f' \ ns k) = (f' \ ns n)") + case True + then show ?thesis using Suc.IH + by (metis order_refl) + next + case False + have "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ val_Zp ((ns k) \ (ns n))" + using False P0 f'_closed poly_diff_val Suc.IH + by blast + then have "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ val_Zp (\ divide (f\(ns n))(f'\(ns n)))" + using F4 by metis + then have "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ val_Zp (divide (f\(ns n))(f'\(ns n)))" + using F5 val_Zp_of_minus + by presburger + then have P10: "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ val_Zp (f\(ns n)) - val_Zp (f'\(ns n))" + using F7 by metis + have P11: "val_Zp (f'\(ns n)) \ \" + by (simp add: F2) + then have "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ (2 * val_Zp (f'\a)) + 2 ^ n * t - val_Zp (f'\(ns n))" + using F3 P10 + by (smt eint_add_cancel_fact eint_add_left_cancel_le order_trans) + then have P12: "val_Zp ((f' \ ns k) \ (f' \ ns n)) \ (2 *(val_Zp (f'\a))) + 2 ^ n * t - (val_Zp (f'\a))" + by (simp add: F2) + have P13:"val_Zp ((f' \ ns k) \ (f' \ ns n)) \ (val_Zp (f'\a)) + 2 ^ n * t " + proof- + have "(2 *(val_Zp (f'\a))) + (2 ^ n * t) - (val_Zp (f'\a)) = (2 *(val_Zp (f'\a))) - (val_Zp (f'\a)) + (2 ^ n * t) " + using eint_minus_comm by blast + then show ?thesis using P12 + using f'a_not_infinite by force + qed + then have P14: "val_Zp ((f' \ ns k) \ (f' \ ns n)) > (val_Zp (f'\a))" + using f'a_not_infinite ge_plus_pos_imp_gt t_times_pow_pos by blast + show ?thesis + by (meson F1 F2 P0 P14 equal_val_Zp f'_closed f'a_closed to_fun_closed) + qed + have P2: "val_Zp (f\(ns k)) \ 2*(val_Zp (f'\a)) + (2^k)*t" + proof- + have P23: "2 * (val_Zp (f'\a)) + ((2 ^ k) * t) \ val_Zp (f \ ns k)" + proof- + have 0: "ns n \ carrier Zp" + by (simp add: F1) + have 1: "local.divide (f \ ns n) (f' \ ns n) \ carrier Zp" + using F5 by blast + have 2: "(poly_shift_iter 2 (taylor (ns n) f)) \ \ local.divide (f \ ns n) (f' \ ns n) \ carrier Zp" + using F1 F5 shift_closed 1 + by (simp add: taylor_closed to_fun_closed) + have 3: "divide (f \ ns n) (f' \ ns n) \ deriv f (ns n) = f \ ns n" + by (metis F1 F2 F6 divide_formula f'_closed f'a_not_infinite f_closed mult_comm pderiv_eval_deriv to_fun_closed val_Zp_def) + have 4: "f \ carrier Zp_x" + by simp + obtain c where c_def: "c = poly_shift_iter (2::nat) (taylor (ns n) f) \ \ local.divide (f \ ns n) (f' \ ns n)" + by blast + then have c_def': "c \ carrier Zp \ f \ (ns n \ local.divide (f \ ns n) (f' \ ns n)) = c \ local.divide (f \ ns n) (f' \ ns n) [^] (2::nat)" + using 0 1 2 3 4 UP_cring.taylor_deg_1_eval'''[of Zp f "ns n" "(divide (f\(ns n)) (f'\(ns n)))" c] + Zp_x_is_UP_cring + by blast + have P230: "f\(ns k) = (c \ (divide (f\(ns n)) (f'\(ns n)))[^](2::nat))" + using c_def' + by (simp add: kval newton_step_def) + have P231: "val_Zp (f\(ns k)) = val_Zp c + 2*(val_Zp (f\(ns n)) - val_Zp(f'\(ns n)))" + proof- + have P2310: "val_Zp (f\(ns k)) = val_Zp c + val_Zp ((divide (f\(ns n)) (f'\(ns n)))[^](2::nat))" + by (simp add: F5 P230 c_def' val_Zp_mult) + have P2311: "val_Zp ((divide (f\(ns n)) (f'\(ns n)))[^](2::nat)) + = 2*(val_Zp (f\(ns n)) - val_Zp(f'\(ns n)))" + by (metis F5 F7 R.pow_zero mult.commute not_nonzero_Zp of_nat_numeral times_eint_simps(3) val_Zp_def val_Zp_pow' zero_less_numeral) + thus ?thesis + by (simp add: P2310) + qed + have P232: "val_Zp (f\(ns k)) \ 2*(val_Zp (f\(ns n)) - val_Zp(f'\(ns n)))" + by (simp add: P231 c_def' val_pos) + have P236: "val_Zp (f\(ns k)) \ 2*(2 *val_Zp (f'\a) + 2 ^ n * t) - 2* val_Zp(f'\(ns n))" + using P232 F3 eint_minus_ineq''[of "val_Zp(f'\(ns n))" "(2 *val_Zp (f'\a)) + 2 ^ n * t" "val_Zp (f\(ns n))" 2 ] + F2 eint_pow_int_is_pos by auto + hence P237: "val_Zp (f\(ns k)) \(4*val_Zp (f'\a)) + (2*((2 ^ n)* t)) - 2* val_Zp(f'\(ns n))" + proof- + have "2*(2*val_Zp (f'\a) + 2 ^ n * t) = (4*val_Zp (f'\a)) + 2*(2 ^ n)* t " + using distrib_left[of 2 "2*val_Zp (f'\a)" "2 ^ n * t"] mult.assoc mult_one_right one_eint_def plus_eint_simps(1) + hensel_factor_def val_Zp_def by auto + then show ?thesis + using P236 + by (metis mult.assoc) + qed + hence P237: "val_Zp (f\(ns k)) \ 4*val_Zp (f'\a) + 2*(2 ^ n)* t - 2* val_Zp((f'\a))" + by (metis F2 mult.assoc) + hence P238: "val_Zp (f\(ns k)) \ 2*val_Zp (f'\a) + 2*(2 ^ n)* t" + using eint_minus_comm[of "4*val_Zp (f'\a) " "2*(2 ^ n)* t" "2* val_Zp((f'\a))"] + by (simp add: eint_int_minus_distr) + thus ?thesis + by (simp add: kval) + qed + thus ?thesis + by blast + qed + show "val_Zp (to_fun f' (ns k)) = val_Zp (f'\a) \ + 2 * val_Zp (f'\a) + eint (2 ^ k) * t \ val_Zp (to_fun f (ns k))" + using P1 P2 by blast + qed + qed +qed + +lemma newton_seq_closed: +shows "ns m \ carrier Zp" + using newton_seq_props_induct + by blast + +lemma f_of_newton_seq_closed: +shows "f \ ns m \ carrier Zp" + by (simp add: to_fun_closed newton_seq_closed) + +lemma newton_seq_fact1[simp]: +" val_Zp (f'\(ns k)) = val_Zp ((f'\a))" +using newton_seq_props_induct by blast + +lemma newton_seq_fact2: +"\k. val_Zp (f\(ns k)) \ 2*(val_Zp (f'\a)) + (2^k)*t" + by (meson le_iff_add newton_seq_props_induct) + +lemma newton_seq_fact3: +"val_Zp (f\(ns l)) \ val_Zp (f'\(ns l))" +proof- + have "2*(val_Zp (f'\a)) + (2^l)*t \ (val_Zp (f'\a))" + using f'a_closed ord_pos t_pos + by (smt eint_pos_int_times_ge f'a_nonneg_val f'a_not_infinite ge_plus_pos_imp_gt linorder_not_less nat_mult_not_infty order_less_le t_times_pow_pos) + then show "val_Zp (f \ ns l) \ val_Zp (f' \ ns l) " + using f'a_closed f'a_nonzero newton_seq_fact1[of l] newton_seq_fact2[of l] val_Zp_def + proof - + show ?thesis + using \eint 2 * val_Zp (f'\a) + eint (2 ^ l) * t \ val_Zp (to_fun f (ns l))\ \val_Zp (f'\a) \ eint 2 * val_Zp (f'\a) + eint (2 ^ l) * t\ by force + qed +qed + +lemma newton_seq_fact4[simp]: + assumes "f\(ns l) \\" + shows "val_Zp (f\(ns l)) \ val_Zp (f'\(ns l))" + using newton_seq_fact3 by blast + +lemma newton_seq_fact5: +"divide (f \ ns l) (f' \ ns l) \ carrier Zp" + apply(rule divide_closed) + apply (simp add: to_fun_closed newton_seq_closed) + apply (simp add: f'_closed to_fun_closed newton_seq_closed) + by (metis f'a_not_infinite newton_seq_fact1 val_Zp_def) + +lemma newton_seq_fact6: +"(f'\(ns l)) \ nonzero Zp" + apply(rule ccontr) + using nonzero_memI nonzero_memE + f'a_nonzero newton_seq_fact1 val_Zp_def + by (metis (no_types, lifting) divide_closed f'_closed f'a_closed fa_closed hensel_factor_id + hensel_hypothesis_weakened mult_zero_l newton_seq_closed order_less_le to_fun_closed val_Zp_mult) + +lemma newton_seq_fact7: + "(ns (Suc n)) \ (ns n) = \divide (f\(ns n)) (f'\(ns n))" + using newton_seq.simps(2)[of n] newton_seq_fact5[of n] + newton_seq_closed[of "Suc n"] newton_seq_closed[of n] + R.ring_simprules + unfolding newton_step_def a_minus_def + by smt + +lemma newton_seq_fact8: + assumes "f\(ns l) \\" + shows "divide (f \ ns l) (f' \ ns l) \ nonzero Zp" + using assms divide_nonzero[of "f \ ns l" "f' \ ns l"] + nonzero_memI + using f_of_newton_seq_closed newton_seq_fact3 newton_seq_fact6 by blast + +lemma newton_seq_fact9: + assumes "f\(ns n) \\" + shows "val_Zp((ns (Suc n)) \ (ns n)) = val_Zp (f\(ns n)) - val_Zp (f'\(ns n))" + using newton_seq_fact7 val_of_divide newton_seq_fact6 assms nonzero_memI + f_of_newton_seq_closed newton_seq_fact4 newton_seq_fact5 + by (metis val_Zp_of_minus) + +text\Assuming no element of the Newton sequence is a root of f, the Newton sequence is Cauchy.\ + +lemma newton_seq_is_Zp_cauchy_0: +assumes "\k. f\(ns k) \\" +shows "is_Zp_cauchy ns" +proof(rule is_Zp_cauchyI') + show P0: "ns \ closed_seqs Zp" + proof(rule closed_seqs_memI) + show "\k. ns k \ carrier Zp " + by (simp add: newton_seq_closed) + qed + show "\n. \k. \m. k \ int m \ int n \ val_Zp (ns (Suc m) \ ns m)" + proof + fix n + show "\k. \m. k \ int m \ int n \ val_Zp (ns (Suc m) \ ns m)" + proof(induction "n") + case 0 + have B0: "\n0 n1. 0 < n0 \ 0 < n1 \ ns n0 0 = ns n1 0" + apply auto + proof- + fix n0 n1::nat + assume A: "0 < n0" "0 < n1" + show "ns n0 0 = ns n1 0" + proof- + have 0: "ns n0 \ carrier Zp" + using P0 + by (simp add: newton_seq_closed) + have 1: "ns n1 \ carrier Zp" + using P0 + by (simp add: newton_seq_closed) + show ?thesis + using 0 1 Zp_defs(3) prime + by (metis p_res_ring_0' residue_closed) + qed + qed + have "\m. 1 \ int m \ int 0 \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof + fix m + show "1 \ int m \ int 0 \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof + assume "1 \ int m " + then have C0:"ns (Suc m) 0 = ns m 0" + using B0 + by (metis int_one_le_iff_zero_less int_ops(1) less_Suc_eq_0_disj of_nat_less_iff) + then show "int 0 \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof- + have "(newton_step (ns m)) \(ns m)" + proof- + have A0: "divide (f\(ns m)) (f'\(ns m)) \\" + proof- + have 0: "(f\(ns m)) \ \" + using assms by auto + have 1: " (f'\(ns m)) \ carrier Zp" + by (simp add: UP_cring.to_fun_closed Zp_x_is_UP_cring f'_closed newton_seq_closed) + have 2: "(f'\(ns m)) \ \" + using newton_seq_fact6 not_nonzero_memI by blast + show ?thesis using 0 1 2 + by (metis R.r_null divide_formula f_closed to_fun_closed newton_seq_closed newton_seq_fact4) + qed + have A2: "local.divide (f \ ns m) (f' \ ns m) \ carrier Zp" + using newton_seq_fact5 by blast + have A3: "ns m \ carrier Zp" + by (simp add: newton_seq_closed) + have A4: "newton_step (ns m) \ carrier Zp" + by (metis newton_seq.simps(2) newton_seq_closed) + show ?thesis + apply(rule ccontr) + using A4 A3 A2 A0 newton_step_def[of "(ns m)"] + by (simp add: a_minus_def) + qed + then show ?thesis using C0 + by (metis newton_seq.simps(2) newton_seq_closed val_Zp_dist_res_eq2) + qed + qed + qed + then show ?case + using val_Zp_def val_Zp_dist_def + by (metis int_ops(1) newton_seq.simps(2) zero_eint_def) + next + case (Suc n) + show "\k. \m. k \ int m \ int (Suc n) \ val_Zp (ns (Suc m) \ ns m)" + proof- + obtain k0 where k0_def: "k0 \0 \ (\m. k0 \ int m \ int n \ val_Zp (ns (Suc m) \ ns m))" + using Suc.IH + by (metis int_nat_eq le0 nat_le_iff of_nat_0_eq_iff ) + have I0: "\l. val_Zp (ns (Suc l) \ ns l) = val_Zp (f\ (ns l)) - val_Zp (f'\(ns l))" + proof- + fix l + have I00:"(ns (Suc l) \ ns l) = (\ divide (f\(ns l)) (f'\(ns l)))" + proof- + have "local.divide (f \ ns l) (f' \ ns l) \ carrier Zp" + by (simp add: newton_seq_fact5) + then show ?thesis + using newton_seq.simps(2)[of l] newton_seq_closed R.ring_simprules + unfolding newton_step_def a_minus_def + by (metis add_comm) + qed + have I01: "val_Zp (ns (Suc l) \ ns l) = val_Zp (divide (f\(ns l)) (f'\(ns l)))" + proof- + have I010: "(divide (f\(ns l)) (f'\(ns l))) \carrier Zp" + by (simp add: newton_seq_fact5) + have I011: "(divide (f\(ns l)) (f'\(ns l))) \ \" + proof- + have A: "(f\(ns l)) \\" + by (simp add: assms) + have B: " (f'\(ns l)) \carrier Zp" + using nonzero_memE newton_seq_fact6 by auto + then have C: " (f'\(ns l)) \nonzero Zp" + using f'a_closed fa_closed fa_nonzero hensel_factor_id hensel_hypothesis_weakened + newton_seq_fact1[of l] not_nonzero_Zp val_Zp_def + by fastforce + then show ?thesis using I010 A + by (metis B R.r_null divide_formula f_closed to_fun_closed newton_seq_closed newton_seq_fact4 nonzero_memE(2)) + qed + then have "val_Zp (divide (f\(ns l)) (f'\(ns l))) + = val_Zp (\ divide (f\(ns l)) (f'\(ns l)))" + using I010 not_nonzero_Zp val_Zp_of_minus by blast + then show ?thesis using I00 by metis + qed + have I02: "val_Zp (f\(ns l)) \ val_Zp (f'\(ns l))" + using assms newton_seq_fact4 + by blast + have I03: "(f\(ns l)) \ nonzero Zp" + by (meson UP_cring.to_fun_closed Zp_x_is_UP_cring assms f_closed newton_seq_closed not_nonzero_Zp) + have I04: "f'\(ns l) \ nonzero Zp" + by (simp add: newton_seq_fact6) + have I05 :" val_Zp (divide (f\(ns l)) (f'\(ns l))) = val_Zp (f\ (ns l)) - val_Zp (f'\(ns l))" + using I02 I03 I04 I01 assms newton_seq_fact9 by auto + then show " val_Zp (ns (Suc l) \ ns l) = val_Zp (f\ (ns l)) - val_Zp (f'\(ns l))" + using I01 by simp + qed + have "\m. int(Suc n) + k0 + 1 \ int m \ int (Suc n) \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof + fix m + show "int (Suc n) + k0 + 1 \ int m \ int (Suc n) \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof + assume A: "int (Suc n) + k0 + 1 \ int m " + show " int (Suc n) \ val_Zp_dist (newton_step (ns m)) (ns m)" + proof- + have 0: " val_Zp_dist (newton_step (ns m)) (ns m) = val_Zp (f\ (ns m)) - val_Zp (f'\(ns m))" + using I0 val_Zp_dist_def by auto + have 1: "val_Zp (f\ (ns m)) - val_Zp (f'\(ns m)) > int n" + proof- + have "val_Zp (f\ (ns m)) \ 2*(val_Zp (f'\a)) + (2^m)*t" + by (simp add: newton_seq_fact2) + then have 10:"val_Zp (f\ (ns m)) - val_Zp (f'\(ns m)) \ 2*(val_Zp (f'\a)) + (2^m)*t - val_Zp (f'\(ns m))" + by (simp add: eint_minus_ineq) + have "2^m * t > m" + apply(induction m) + using one_eint_def zero_eint_def apply auto[1] + proof- fix m + assume IH : "int m < 2 ^ m * t " + then have "((2 ^ (Suc m)) * t) = 2* ((2 ^ m) * t)" + by (metis mult.assoc power_Suc times_eint_simps(1)) + then show "int (Suc m) < 2 ^ Suc m * t" + using IH t_neq_infty by force + qed + then have 100: "2^m * t > int m" + by blast + have "int m \2 + (int n + k0)" + using A by simp + hence 1000: "2^m * t > 2 + (int n + k0)" + using 100 + by (meson eint_ord_simps(2) less_le_trans linorder_not_less) + have "2 + (int n + k0) > 1 + int n" + using k0_def by linarith + then have "2^m * t > 1 + int n" + using 1000 eint_ord_simps(2) k0_def less_le_trans linorder_not_less + proof - + have "eint (2 + (int n + k0)) < t * eint (int (2 ^ m))" + by (metis "1000" mult.commute numeral_power_eq_of_nat_cancel_iff) + then have "eint (int (Suc n)) < t * eint (int (2 ^ m))" + by (metis \1 + int n < 2 + (int n + k0)\ eint_ord_simps(2) less_trans of_nat_Suc) + then show ?thesis + by (simp add: mult.commute) + qed + hence "2*val_Zp (f'\a) + eint (2 ^ m) * t \ 2*(val_Zp (f'\a)) + 1 + int n" + by (smt eSuc_eint eint_add_left_cancel_le iadd_Suc iadd_Suc_right order_less_le) + then have 11: "val_Zp (f\ (ns m)) - val_Zp (f'\(ns m)) + \ 2*(val_Zp (f'\a)) + 1 + int n - val_Zp (f'\(ns m))" + using "10" + by (smt \eint 2 * val_Zp (f'\a) + eint (2 ^ m) * t \ val_Zp (to_fun f (ns m))\ + f'a_not_infinite eint_minus_ineq hensel_axioms newton_seq_fact1 order_trans) + have 12: "val_Zp (f'\(ns m)) = val_Zp (f'\a) " + using nonzero_memE newton_seq_fact1 newton_seq_fact6 val_Zp_def val_Zp_def + by auto + then have 13: "val_Zp (f\ (ns m)) - val_Zp (f'\(ns m)) + \ 2*(val_Zp (f'\a)) + (1 + int n) - val_Zp ((f'\a))" + using 11 + by (smt eSuc_eint iadd_Suc iadd_Suc_right) + then have 14:"val_Zp (f\ (ns m)) - val_Zp (f'\(ns m)) + \ 1 + int n + val_Zp ((f'\a))" + using eint_minus_comm[of "2*(val_Zp (f'\a))" "1 + int n" "val_Zp ((f'\a))"] + by (simp add: Groups.add_ac(2)) + then show ?thesis + by (smt Suc_ile_eq add.right_neutral eint.distinct(2) f'a_nonneg_val ge_plus_pos_imp_gt order_less_le) + qed + then show ?thesis + by (smt "0" Suc_ile_eq of_nat_Suc) + qed + qed + qed + then show ?thesis + using val_Zp_def val_Zp_dist_def + by (metis newton_seq.simps(2)) + qed + qed + qed +qed + +lemma eventually_zero: +"f \ ns (k + m) = \ \ f \ ns (k + Suc m) = \" +proof- + assume A: "f \ ns (k + m) = \" + have 0: "ns (k + Suc m) = ns (k + m) \ (divide (f \ ns (k + m)) (f' \ ns (k + m)))" + by (simp add: newton_step_def) + have 1: "(divide (f \ ns (k + m)) (f' \ ns (k + m))) = \" + by (simp add: A divide_def) + show "f \ ns (k + Suc m) = \" + using A 0 1 + by (simp add: a_minus_def newton_seq_closed) +qed + +text\The Newton Sequence is Cauchy:\ + +lemma newton_seq_is_Zp_cauchy: +"is_Zp_cauchy ns" +proof(cases "\k. f\(ns k) \\") + case True + then show ?thesis using newton_seq_is_Zp_cauchy_0 + by blast +next + case False + obtain k where k_def:"f\(ns k) = \" + using False by blast + have 0: "\m. (ns (m + k)) = (ns k)" + proof- + fix m + show "(ns (m + k)) = (ns k)" + proof(induction m) + case 0 + then show ?case + by simp + next + case (Suc m) + show "(ns (Suc m + k)) = (ns k)" + proof- + have "f \ ns (m + k) = \" + by (simp add: Suc.IH k_def) + then have "divide ( f \ ns (m + k)) (f' \ ns (m + k)) = \" + by (simp add: divide_def) + then show ?thesis using newton_step_def + by (simp add: Suc.IH a_minus_def newton_seq_closed) + qed + qed + qed + show "is_Zp_cauchy ns" + apply(rule is_Zp_cauchyI) + apply (simp add: closed_seqs_memI newton_seq_closed) + proof- + show "\n.\n. \N. \n0 n1. N < n0 \ N < n1 \ ns n0 n = ns n1 n" + proof- + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ ns n0 n = ns n1 n" + proof- + have "\n0 n1. k < n0 \ k < n1 \ ns n0 n = ns n1 n" + apply auto + proof- + fix n0 n1 + assume A0: "k < n0" + assume A1: "k < n1" + obtain m0 where m0_def: "n0 = k + m0" + using A0 less_imp_add_positive by blast + obtain m1 where m1_def: "n1 = k + m1" + using A1 less_imp_add_positive by auto + show "ns n0 n = ns n1 n" + using 0 m0_def m1_def + by (metis add.commute) + qed + then show ?thesis by blast + qed + qed + qed +qed + +subsection\The Proof of Hensel's Lemma\ +lemma pre_hensel: +"val_Zp (a \ (ns n)) > val_Zp (f'\a)" +"\N. \n. n> N \ (val_Zp (a \ (ns n)) = val_Zp (divide (f\a) (f'\a)))" +"val_Zp (f'\(ns n)) = val_Zp (f'\a)" +proof- + show "val_Zp (a \ (ns n)) > val_Zp (f'\a)" + proof(induction n) + case 0 + then show ?case + by (simp add: val_Zp_def) + next + case (Suc n) + show "val_Zp (a \ (ns (Suc n))) > val_Zp (f'\a)" + proof- + have I0: "val_Zp ((ns (Suc n)) \ (ns n)) > val_Zp (f'\a)" + proof(cases "(ns (Suc n)) = (ns n)") + case True + then show ?thesis + by (simp add: newton_seq_closed val_Zp_def) + next + case False + have 00:"(ns (Suc n)) \ (ns n) = \divide (f\(ns n)) (f'\(ns n))" + using newton_seq_fact7 by blast + then have 0: "val_Zp((ns (Suc n)) \ (ns n)) = val_Zp (divide (f\(ns n)) (f'\(ns n)))" + using newton_seq_fact5 val_Zp_of_minus by presburger + have 1: "(f\(ns n)) \ nonzero Zp" + by (metis False R.minus_zero R.r_right_minus_eq 00 divide_def f_closed to_fun_closed + newton_seq_closed not_nonzero_Zp) + have 2: "f'\(ns n) \ nonzero Zp" + by (simp add: newton_seq_fact6) + have "val_Zp (f\(ns n)) \ val_Zp (f'\(ns n))" + using nonzero_memE \f \ ns n \ nonzero Zp\ newton_seq_fact4 by blast + then have 3:"val_Zp((ns (Suc n)) \ (ns n)) = val_Zp (f\(ns n)) - val_Zp (f'\(ns n))" + using 0 1 2 newton_seq_fact9 nonzero_memE(2) by blast + have 4: "val_Zp (f \ ns n) \ (2 * val_Zp (f'\a)) + 2 ^ n * t" + using newton_seq_fact2[of n] by metis + then have 5: "val_Zp((ns (Suc n)) \ (ns n)) \ ((2 * val_Zp (f'\a)) + 2 ^ n * t) - val_Zp (f'\(ns n))" + using "3" eint_minus_ineq f'a_not_infinite newton_seq_fact1 by presburger + have 6: "((ns (Suc n)) \ (ns n)) \ nonzero Zp" + using False not_eq_diff_nonzero newton_seq_closed by blast + then have "val_Zp((ns (Suc n)) \ (ns n)) \ (2 * val_Zp (f'\a)) + 2 ^ n * t - val_Zp ((f'\a))" + using "5" by auto + then have 7: "val_Zp((ns (Suc n)) \ (ns n)) \ (val_Zp (f'\a)) + 2 ^ n * t" + by (simp add: eint_minus_comm) + then show "val_Zp((ns (Suc n)) \ (ns n)) > (val_Zp (f'\a))" + using f'a_not_infinite ge_plus_pos_imp_gt t_times_pow_pos by blast + qed + have "val_Zp ((ns (Suc n)) \ (ns n)) = val_Zp ((ns n) \ (ns (Suc n)))" + using newton_seq_closed[of "n"] newton_seq_closed[of "Suc n"] + val_Zp_def val_Zp_dist_def val_Zp_dist_sym val_Zp_def + by auto + then have I1: "val_Zp ((ns n) \ (ns (Suc n))) > val_Zp (f'\a)" + using I0 + by presburger + have I2: " (a \ (ns n)) \ ((ns n) \ (ns (Suc n))) = (a \ (ns (Suc n)))" + by (metis R.plus_diff_simp add_comm local.a_closed newton_seq_closed) + then have "val_Zp (a \ (ns (Suc n))) \ min (val_Zp (a \ ns n)) (val_Zp (ns n \ ns (Suc n)))" + by (metis R.minus_closed local.a_closed newton_seq_closed val_Zp_ultrametric) + thus ?thesis + using I1 Suc.IH eint_min_ineq by blast + qed + qed + show "val_Zp (f'\(ns n)) = val_Zp (f'\a)" + using newton_seq_fact1 by blast + show "\N.\n. n> N \ (val_Zp (a \ (ns n)) = val_Zp (divide (f\a) (f'\a)))" + proof- + have P: "\m. m > 1 \ (val_Zp (a \ (ns m)) = val_Zp (divide (f\a) (f'\a)))" + proof- + fix n::nat + assume AA: "n >1" + show " (val_Zp (a \ (ns n)) = val_Zp (divide (f\a) (f'\a)))" + proof(cases "(ns 1) = a") + case True + have T0: "\k. \n. n \ k \ ns n = a" + proof- + fix k + show " \n. n \ k \ ns n = a" + proof(induction k) + case 0 + then show ?case + by simp + next + case (Suc k) + show "\n\Suc k. ns n = a" apply auto + proof- + fix n + assume A: "n \Suc k" + show "ns n = a" + proof(cases "n < Suc k") + case True + then show ?thesis using Suc.IH by auto + next + case False thus ?thesis + using A Suc.IH True by auto + qed + qed + qed + qed + show "val_Zp (a \ ns n) = val_Zp (local.divide (f\a) (f'\a))" + by (metis T0 Zp_def Zp_defs(3) f'a_closed f'a_nonzero fa_nonzero + hensel.fa_closed hensel_axioms hensel_hypothesis_weakened le_eq_less_or_eq + newton_seq_fact9 not_nonzero_Qp order_less_le val_of_divide) + next + case False + have F0: "(1::nat) \ n" + using AA by simp + have "(f\a) \ \" + by simp + have "\k. val_Zp (a \ ns (Suc k)) = val_Zp (local.divide (f\a) (f'\a))" + proof- + fix k + show " val_Zp (a \ ns (Suc k)) = val_Zp (local.divide (f\a) (f'\a))" + proof(induction k) + case 0 + have "(a \ ns (Suc 0)) = (local.divide (f\a) (f'\a))" + by (metis R.minus_minus Zp_def hensel.newton_seq_fact7 hensel_axioms + local.a_closed minus_a_inv newton_seq.simps(1) newton_seq.simps(2) newton_seq_fact5 newton_step_closed) + then show ?case by simp + next + case (Suc k) + have I0: "ns (Suc (Suc k)) = ns (Suc k) \ (divide (f\(ns (Suc k))) (f'\(ns (Suc k))))" + by (simp add: newton_step_def) + have I1: "val_Zp (f\(ns (Suc k))) \ val_Zp(f'\(ns (Suc k)))" + using newton_seq_fact3 by blast + have I2: "(divide (f\(ns (Suc k))) (f'\(ns (Suc k)))) \ carrier Zp" + using newton_seq_fact5 by blast + have I3: "ns (Suc (Suc k)) \ ns (Suc k) = \(divide (f\(ns (Suc k))) (f'\(ns (Suc k))))" + using I0 I2 newton_seq_fact7 by blast + then have "val_Zp (ns (Suc (Suc k)) \ ns (Suc k)) = val_Zp (divide (f\(ns (Suc k))) (f'\(ns (Suc k))))" + using I2 val_Zp_of_minus + by presburger + then have "val_Zp (ns (Suc (Suc k)) \ ns (Suc k)) = val_Zp (f\(ns (Suc k))) - val_Zp (f'\(ns (Suc k)))" + by (metis I1 R.zero_closed Zp_def newton_seq_fact6 newton_seq_fact9 padic_integers.val_of_divide padic_integers_axioms) + then have I4: "val_Zp (ns (Suc (Suc k)) \ ns (Suc k)) = val_Zp (f\(ns (Suc k))) - val_Zp ((f'\a))" + using newton_seq_fact1 by presburger + have F3: "val_Zp (a \ ns (Suc k)) = val_Zp (local.divide (f\a) (f'\a))" + using Suc.IH by blast + have F4: "a \ ns (Suc (Suc k)) = (a \ ( ns (Suc k))) \ (ns (Suc k)) \ ns (Suc (Suc k))" + by (metis R.ring_simprules(17) a_minus_def add_comm local.a_closed newton_seq_closed) + have F5: "val_Zp ((ns (Suc k)) \ ns (Suc (Suc k))) > val_Zp (a \ ( ns (Suc k)))" + proof- + have F50: "val_Zp ((ns (Suc k)) \ ns (Suc (Suc k))) = val_Zp (f\(ns (Suc k))) - val_Zp ((f'\a))" + by (metis I4 R.minus_closed minus_a_inv newton_seq_closed val_Zp_of_minus) + + have F51: "val_Zp (f\(ns (Suc k))) > val_Zp ((f\a))" + proof- + have F510: "val_Zp (f\(ns (Suc k))) \ 2*val_Zp (f'\a) + 2^(Suc k)*t " + using newton_seq_fact2 by blast + hence F511: "val_Zp (f\(ns (Suc k))) \ 2*val_Zp (f'\a) + 2*t " + using eint_plus_times[of t "2*val_Zp (f'\a)" "2^(Suc k)" "val_Zp (f\(ns (Suc k)))" 2] t_pos + by (simp add: order_less_le) + have F512: "2*val_Zp (f'\a) + 2*t = 2 *val_Zp (f\a) - 2* val_Zp (f'\a)" + unfolding hensel_factor_def + using eint_minus_distr[of "val_Zp (f\a)" "2 * val_Zp (f'\a)" 2] + eint_minus_comm[of _ _ "eint 2 * (eint 2 * val_Zp (f'\a))"] + by (smt eint_2_minus_1_mult eint_add_cancel_fact eint_minus_comm f'a_not_infinite hensel_hypothesis nat_mult_not_infty order_less_le) + hence "2*val_Zp (f'\a) + 2*t > val_Zp (f\a)" + using hensel_hypothesis + by (smt add_diff_cancel_eint eint_add_cancel_fact eint_add_left_cancel_le + eint_pos_int_times_gt f'a_not_infinite hensel_factor_def nat_mult_not_infty order_less_le t_neq_infty t_pos) + thus ?thesis using F512 + using F511 less_le_trans by blast + qed + thus ?thesis + by (metis F3 F50 Zp_def divide_closed eint_add_cancel_fact eint_minus_ineq + f'a_closed f'a_nonzero f'a_not_infinite fa_closed fa_nonzero hensel.newton_seq_fact7 + hensel_axioms newton_seq.simps(1) newton_seq_fact9 order_less_le val_Zp_of_minus) + qed + have "a \ ns (Suc k) \ (ns (Suc k) \ ns (Suc (Suc k))) = a \ ns (Suc (Suc k))" + by (metis F4 a_minus_def add_assoc) + then show F6: "val_Zp (a \ ns (Suc (Suc k))) = val_Zp (local.divide (f\a) (f'\a))" + using F5 F4 F3 + by (metis R.minus_closed local.a_closed newton_seq_closed order_less_le val_Zp_not_equal_ord_plus_minus val_Zp_ultrametric_eq'') + qed + qed + thus ?thesis + by (metis AA less_imp_add_positive plus_1_eq_Suc) + qed + qed + thus ?thesis + by blast + qed +qed + +lemma hensel_seq_comp_f: + "res_lim ((to_fun f) \ ns) = \" +proof- + have A: "is_Zp_cauchy ((to_fun f) \ ns)" + using f_closed is_Zp_continuous_def newton_seq_is_Zp_cauchy polynomial_is_Zp_continuous + by blast + have "Zp_converges_to ((to_fun f) \ ns) \" + apply(rule Zp_converges_toI) + using A is_Zp_cauchy_def apply blast + apply simp + proof- + fix n + show " \N. \k>N. (((to_fun f) \ ns) k) n = \ n" + proof- + have 0: "\k. (k::nat)>3 \ val_Zp (f\(ns k)) > k" + proof + fix k::nat + assume A: "k >3" + show "val_Zp (f\(ns k)) > k " + proof- + have 0: " val_Zp (f\(ns k)) \ 2*(val_Zp (f'\a)) + (2^k)*t" + using newton_seq_fact2 by blast + have 1: "2*(val_Zp (f'\a)) + (2^k)*t > k " + proof- + have "(2^k)*t \ (2^k) " + apply(cases "t = \") + apply simp + using t_pos eint_mult_mono' + proof - + obtain ii :: "eint \ int" where + f1: "\e. (\ \ e \ (\i. eint i \ e)) \ (eint (ii e) = e \ \ = e)" + by (metis not_infinity_eq) + then have "0 < ii t" + by (metis (no_types) eint_ord_simps(2) t_neq_infty t_pos zero_eint_def) + then show ?thesis + using f1 by (metis eint_pos_int_times_ge eint_mult_mono linorder_not_less + mult.commute order_less_le t_neq_infty t_pos t_times_pow_pos) + qed + hence " 2*(val_Zp (f'\a)) + (2^k)*t \ (2^k) " + by (smt Groups.add_ac(2) add.right_neutral eint_2_minus_1_mult eint_pos_times_is_pos + eint_pow_int_is_pos f'a_nonneg_val ge_plus_pos_imp_gt idiff_0_right linorder_not_less + nat_mult_not_infty order_less_le t_neq_infty) + then have " 2*(val_Zp (f'\a)) + (2^k)*t > k" + using A of_nat_1 of_nat_add of_nat_less_two_power + by (smt eint_ord_simps(1) linorder_not_less order_trans) + then show ?thesis + by metis + qed + thus ?thesis + using 0 less_le_trans by blast + qed + qed + have 1: "\k. (k::nat)>3 \ (f\(ns k)) k = 0" + proof + fix k::nat + assume B: "3(ns k)) k = 0" + proof- + have B0: " val_Zp (f\(ns k)) > k" + using 0 B + by blast + then show ?thesis + by (simp add: f_of_newton_seq_closed zero_below_val_Zp) + qed + qed + have "\k>(max 3 n). (((to_fun f) \ ns) k) n = \ n" + apply auto + proof- + fix k::nat + assume A: "3< k" + assume A': "n < k" + have A0: "(f\(ns k)) k = 0" + using 1[of k] A by auto + then have "(f\(ns k)) n = 0" + using A A' + using above_ord_nonzero[of "(f\(ns k))"] + by (smt UP_cring.to_fun_closed Zp_x_is_UP_cring f_closed le_eq_less_or_eq + newton_seq_closed of_nat_mono residue_of_zero(2) zero_below_ord) + then show A1: "to_fun f (ns k) n = \ n" + by (simp add: residue_of_zero(2)) + qed + then show ?thesis by blast + qed + qed + then show ?thesis + by (metis Zp_converges_to_def unique_limit') +qed + +lemma full_hensels_lemma: + obtains \ where + "f\\ = \" and "\ \ carrier Zp" + "val_Zp (a \ \) > val_Zp (f'\a)" + "(val_Zp (a \ \) = val_Zp (divide (f\a) (f'\a)))" + "val_Zp (f'\\) = val_Zp (f'\a)" +proof(cases "\k. f\(ns k) =\") + case True + obtain k where k_def: "f\(ns k) =\" + using True by blast + obtain N where N_def: "\n. n> N \ (val_Zp (a \ (ns n)) = val_Zp (divide (f\a) (f'\a)))" + using pre_hensel(2) by blast + have Z: "\n. n \k \ f\(ns n) =\" + proof- + fix n + assume A: "n \k" + obtain l where l_def:"n = k + l" + using A le_Suc_ex + by blast + have "\m. f\(ns (k+m)) =\" + proof- + fix m + show "f\(ns (k+m)) =\" + apply(induction m) + apply (simp add: k_def) + using eventually_zero + by simp + qed + then show "f\(ns n) =\" + by (simp add: l_def) + qed + obtain M where M_def: "M = N + k" + by simp + then have M_root: "f\(ns M) =\" + by (simp add: Z) + obtain \ where alpha_def: "\= ns M" + by simp + have T0: "f\\ = \" + using alpha_def M_root + by auto + have T1: "val_Zp (a \ \) > val_Zp (f'\a)" + using alpha_def pre_hensel(1) by blast + have T2: "(val_Zp (a \ \) = val_Zp (divide (f\a) (f'\a)))" + by (metis M_def N_def alpha_def fa_nonzero k_def + less_add_same_cancel1 newton_seq.elims zero_less_Suc) + have T3: "val_Zp (f'\\) = val_Zp (f'\a)" + using alpha_def newton_seq_fact1 by blast + show ?thesis using T0 T1 T2 T3 + using that alpha_def newton_seq_closed + by blast +next + case False + then have Nz: "\k. f\(ns k) \\" + by blast + have ns_cauchy: "is_Zp_cauchy ns" + by (simp add: newton_seq_is_Zp_cauchy) + have fns_cauchy: "is_Zp_cauchy ((to_fun f) \ ns)" + using f_closed is_Zp_continuous_def ns_cauchy polynomial_is_Zp_continuous by blast + have F0: "res_lim ((to_fun f) \ ns) = \" + proof- + show ?thesis + using hensel_seq_comp_f by auto + qed + obtain \ where alpha_def: "\ = res_lim ns" + by simp + have F1: "(f\\)= \" + using F0 alpha_def alt_seq_limit + ns_cauchy polynomial_is_Zp_continuous res_lim_pushforward + res_lim_pushforward' by auto + have F2: "val_Zp (a \ \) > val_Zp (f'\a) \ val_Zp (a \ \) = val_Zp (local.divide (f\a) (f'\a))" + proof- + have 0: "Zp_converges_to ns \" + by (simp add: alpha_def is_Zp_cauchy_imp_has_limit ns_cauchy) + have "val_Zp (a \ \) < \" + using "0" F1 R.r_right_minus_eq Zp_converges_to_def Zp_def hensel.fa_nonzero hensel_axioms local.a_closed val_Zp_def + by auto + hence "1 + max (eint 2 + val_Zp (f'\a)) (val_Zp (\ \ a)) < \" + by (metis "0" R.minus_closed Zp_converges_to_def eint.distinct(2) eint_ord_simps(4) + f'a_not_infinite infinity_ne_i1 local.a_closed max_def minus_a_inv + sum_infinity_imp_summand_infinity val_Zp_of_minus) + then obtain l where l_def: "eint l = 1 + max (eint 2 + val_Zp (f'\a)) (val_Zp (\ \ a))" + by auto + then obtain N where N_def: "(\m>N. 1 + max (2 + val_Zp (f'\a)) (val_Zp (\ \ a)) < val_Zp_dist (ns m) \)" + using 0 l_def Zp_converges_to_def[of ns \] unfolding val_Zp_dist_def + by metis + obtain N' where N'_def: "\n>N'. val_Zp (a \ ns n) = val_Zp (local.divide (f\a) (f'\a))" + using pre_hensel(2) by blast + obtain K where K_def: "K = Suc (max N N')" + by simp + then have F21: "(1+ (max (2 + val_Zp (f'\a)) (val_Zp (\ \ a)))) < val_Zp_dist (ns K) \" + by (metis N_def lessI linorder_not_less max_def order_trans) + have F22: "a \ ns K" + by (smt False K_def N'_def Zp_def cring_def eint.distinct(2) hensel_factor_id lessI + less_le_trans linorder_not_less max_def mult_comm mult_zero_l newton_seq_closed + order_less_le padic_int_is_cring padic_integers.prime padic_integers_axioms ring.r_right_minus_eq + val_Zp_def) + show ?thesis + proof(cases "ns K = \") + case True + then show ?thesis + using pre_hensel F1 False by blast + next + case False + assume "ns K \ \" + show ?thesis + proof- + have P0: " (a \ \) \ nonzero Zp" + by (metis (mono_tags, hide_lams) F1 not_eq_diff_nonzero + \Zp_converges_to ns \\ a_closed Zp_converges_to_def fa_nonzero) + have P1: "(\ \ (ns K)) \ nonzero Zp" + using False not_eq_diff_nonzero \Zp_converges_to ns \\ + Zp_converges_to_def newton_seq_closed + by (metis (mono_tags, hide_lams)) + have P2: "a \ (ns K) \ nonzero Zp" + using F22 not_eq_diff_nonzero + a_closed newton_seq_closed + by blast + have P3: "(a \ \) = a \ (ns K) \ ((ns K) \ \)" + by (metis R.plus_diff_simp \Zp_converges_to ns \\ add_comm Zp_converges_to_def local.a_closed newton_seq_closed) + have P4: "val_Zp (a \ \) \ min (val_Zp (a \ (ns K))) (val_Zp ((ns K) \ \))" + using "0" P3 Zp_converges_to_def newton_seq_closed val_Zp_ultrametric + by auto + have P5: "val_Zp (a \ (ns K)) > val_Zp (f'\a)" + using pre_hensel(1)[of "K"] + by metis + have "1 + max (eint 2 + val_Zp (f'\a)) (val_Zp (\ \ a)) > val_Zp (f'\a)" + proof- + have "1 + max (eint 2 + val_Zp (f'\a)) (val_Zp (\ \ a)) > (eint 2 + val_Zp (f'\a))" + proof - + obtain ii :: int where + f1: "eint ii = 1 + max (eint 2 + val_Zp (f'\a)) (val_Zp (\ \ a))" + by (meson l_def) + then have "1 + (eint 2 + val_Zp (f'\a)) \ eint ii" + by simp + then show ?thesis + using f1 by (metis Groups.add_ac(2) iless_Suc_eq linorder_not_less) + qed + thus ?thesis + by (smt Groups.add_ac(2) eint_pow_int_is_pos f'a_not_infinite ge_plus_pos_imp_gt order_less_le) + qed + hence P6: "val_Zp ((ns K) \ \) > val_Zp (f'\a)" + using F21 unfolding val_Zp_dist_def + by auto + have P7: "val_Zp (a \ \) > val_Zp (f'\a)" + using P4 P5 P6 eint_min_ineq by blast + have P8: "val_Zp (a \ \) = val_Zp (local.divide (f\a) (f'\a))" + proof- + have " 1 + max (2 + val_Zp (f'\a)) (val_Zp_dist \ a) \ val_Zp_dist (ns K) \" + using False F21 + by (simp add: val_Zp_dist_def) + then have "val_Zp(\ \ (ns K)) > max (2 + val_Zp (f'\a)) (val_Zp_dist \ a)" + by (metis "0" Groups.add_ac(2) P1 Zp_converges_to_def eSuc_mono iless_Suc_eq l_def + minus_a_inv newton_seq_closed nonzero_closed val_Zp_dist_def val_Zp_of_minus) + then have "val_Zp(\ \ (ns K)) > val_Zp (a \ \) " + using \Zp_converges_to ns \\ Zp_converges_to_def val_Zp_dist_def val_Zp_dist_sym + by auto + then have P80: "val_Zp (a \ \) = val_Zp (a \ (ns K))" + using P0 P1 Zp_def val_Zp_ultrametric_eq[of "\ \ ns K" "a \ \"] 0 R.plus_diff_simp + Zp_converges_to_def local.a_closed newton_seq_closed nonzero_closed by auto + have P81: "val_Zp (a \ ns K) = val_Zp (local.divide (f\a) (f'\a))" + using K_def N'_def + by (metis (no_types, lifting) lessI linorder_not_less max_def order_less_le order_trans) + then show ?thesis + by (simp add: P80) + qed + thus ?thesis + using P7 by blast + qed + qed + qed + have F3: "val_Zp (f' \ \) = val_Zp (f'\a)" + proof- + have F31: " (f' \ \) = res_lim ((to_fun f') \ ns)" + using alpha_def alt_seq_limit ns_cauchy polynomial_is_Zp_continuous res_lim_pushforward + res_lim_pushforward' f'_closed + by auto + obtain N where N_def: "val_Zp (f'\\ \ f'\(ns N)) > val_Zp ((f'\a))" + by (smt F2 False R.minus_closed Suc_ile_eq Zp_def alpha_def f'_closed f'a_nonzero + local.a_closed minus_a_inv newton_seq.simps(1) newton_seq_is_Zp_cauchy_0 order_trans + padic_integers.poly_diff_val padic_integers_axioms res_lim_in_Zp val_Zp_def val_Zp_of_minus) + show ?thesis + by (metis False N_def alpha_def equal_val_Zp f'_closed newton_seq_closed newton_seq_is_Zp_cauchy_0 newton_seq_fact1 res_lim_in_Zp to_fun_closed) + qed + show ?thesis + using F1 F2 F3 that alpha_def ns_cauchy res_lim_in_Zp + by blast +qed + + +end + +(**************************************************************************************************) +(**************************************************************************************************) +section\Removing Hensel's Lemma from the Hensel Locale\ +(**************************************************************************************************) +(**************************************************************************************************) + +context padic_integers +begin + + +lemma hensels_lemma: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "(pderiv f)\a \ \" + assumes "f\a \\" + assumes "val_Zp (f\a) > 2* val_Zp ((pderiv f)\a)" + obtains \ where + "f\\ = \" and "\ \ carrier Zp" + "val_Zp (a \ \) > val_Zp ((pderiv f)\a)" + "val_Zp (a \ \) = val_Zp (divide (f\a) ((pderiv f)\a))" + "val_Zp ((pderiv f)\\) = val_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 Zp_def that + by blast +qed + +text\Uniqueness of the root found in Hensel's lemma \ + +lemma hensels_lemma_unique_root: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "(pderiv f)\a \ \" + assumes "f\a \\" + assumes "(val_Zp (f\a) > 2* val_Zp ((pderiv f)\a))" + assumes "f\\ = \" + assumes "\ \ carrier Zp" + assumes "val_Zp (a \ \) > val_Zp ((pderiv f)\a)" + assumes "f\\ = \" + assumes "\ \ carrier Zp" + assumes "val_Zp (a \ \) > val_Zp ((pderiv f)\a)" + assumes "val_Zp ((pderiv f)\\) = val_Zp ((pderiv f)\a)" + shows "\ = \" +proof- + have "\ \ a" + using assms(4) assms(6) by auto + have "\ \ a" + using assms(4) assms(9) by auto + have 0: "val_Zp (\ \ \) > val_Zp ((pderiv f)\a)" + proof- + have "\ \ \ = \ ((a \ \) \ (a \ \))" + by (metis R.minus_eq R.plus_diff_simp assms(10) assms(2) assms(7) minus_a_inv) + hence "val_Zp (\ \ \) = val_Zp ((a \ \) \ (a \ \))" + using R.minus_closed assms(10) assms(2) assms(7) val_Zp_of_minus by presburger + thus ?thesis using val_Zp_ultrametric_diff[of "a \ \" "a \ \"] + by (smt R.minus_closed assms(10) assms(11) assms(2) assms(7) assms(8) min.absorb2 min_less_iff_conj) + qed + obtain h where h_def: "h = \ \ \" + by blast + then have h_fact: "h \ carrier Zp \ \ = \ \ h" + by (metis R.l_neg R.minus_closed R.minus_eq R.r_zero add_assoc add_comm assms(10) assms(7)) + then have 1: "f\(\ \ h) = \" + using assms + by blast + obtain c where c_def: "c \ carrier Zp \ f\(\ \ h) = (f \ \) \ (deriv f \)\h \ c \(h[^](2::nat))" + using taylor_deg_1_eval'[of f \ h _ "f \ \" "deriv f \" ] + by (meson taylor_closed assms(1) assms(7) to_fun_closed h_fact shift_closed) + then have "(f \ \) \ (deriv f \)\h \ c \(h[^](2::nat)) = \" + by (simp add: "1") + then have 2: "(deriv f \)\h \ c \(h[^](2::nat)) = \" + by (simp add: assms(1) assms(6) assms(7) deriv_closed h_fact) + have 3: "((deriv f \) \ c \h)\h = \" + proof- + have "((deriv f \) \ c \h)\h = ((deriv f \)\h \ (c \h)\h)" + by (simp add: R.r_distr UP_cring.deriv_closed Zp_x_is_UP_cring assms(1) assms(7) c_def h_fact mult_comm) + then have "((deriv f \) \ c \h)\h = (deriv f \)\h \ (c \(h\h))" + by (simp add: mult_assoc) + then have "((deriv f \) \ c \h)\h = (deriv f \)\h \ (c \(h[^](2::nat)))" + using nat_pow_def[of Zp h "2"] + by (simp add: h_fact) + then show ?thesis + using 2 + by simp + qed + have "h = \" + proof(rule ccontr) + assume "h \ \" + then have "(deriv f \) \ c \h = \" + using 2 3 + by (meson R.m_closed assms(1) assms(7) c_def deriv_closed h_fact local.integral sum_closed) + then have "(deriv f \) = \ c \h" + by (simp add: R.l_minus R.sum_zero_eq_neg UP_cring.deriv_closed Zp_x_is_UP_cring assms(1) assms(7) c_def h_fact) + then have "val_Zp (deriv f \) = val_Zp (c \ h)" + by (meson R.m_closed \deriv f \ \ c \ h = \\ assms(1) assms(7) c_def deriv_closed h_fact val_Zp_not_equal_imp_notequal(3)) + then have P: "val_Zp (deriv f \) = val_Zp h + val_Zp c" + using val_Zp_mult c_def h_fact by force + hence "val_Zp (deriv f \) \ val_Zp h " + using val_pos[of c] + by (simp add: c_def) + then have "val_Zp (deriv f \) \ val_Zp (\ \ \) " + using h_def by blast + then have "val_Zp (deriv f \) > val_Zp ((pderiv f)\a)" + using "0" by auto + then show False using pderiv_eval_deriv[of f \] + using assms(1) assms(12) assms(7) by auto + qed + then show "\ = \" + using assms(10) assms(7) h_def + by auto +qed + +lemma hensels_lemma': + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "val_Zp (f\a) > 2*val_Zp ((pderiv f)\a)" + shows "\!\ \ carrier Zp. f\\ = \ \ val_Zp (a \ \) > val_Zp ((pderiv f)\a)" +proof(cases "f\a = \") + case True + have T0: "pderiv f \ a \ \" + apply(rule ccontr) using assms(3) + unfolding val_Zp_def by simp + then have T1: "a \ carrier Zp \ f\a = \ \ val_Zp (a \ a) > val_Zp ((pderiv f)\a)" + using assms True + by(simp add: val_Zp_def) + have T2: "\b. b \ carrier Zp \ f\b = \ \ val_Zp (a \ b) > val_Zp ((pderiv f)\a) \ a = b" + proof- fix b assume A: "b \ carrier Zp \ f\b = \ \ val_Zp (a \ b) > val_Zp ((pderiv f)\a)" + obtain h where h_def: "h = b \ a" + by blast + then have h_fact: "h \ carrier Zp \ b = a \ h" + by (metis A R.l_neg R.minus_closed R.minus_eq R.r_zero add_assoc add_comm assms(2)) + then have 1: "f\(a \ h) = \" + using assms A by blast + obtain c where c_def: "c \ carrier Zp \ f\(a \ h) = (f \ a) \ (deriv f a)\h \ c \(h[^](2::nat))" + using taylor_deg_1_eval'[of f a h _ "f \ a" "deriv f a" ] + by (meson taylor_closed assms(1) assms(2) to_fun_closed h_fact shift_closed) + then have "(f \ a) \ (deriv f a)\h \ c \(h[^](2::nat)) = \" + by (simp add: "1") + then have 2: "(deriv f a)\h \ c \(h[^](2::nat)) = \" + by (simp add: True assms(1) assms(2) deriv_closed h_fact) + hence 3: "((deriv f a) \ c \h)\h = \" + proof- + have "((deriv f a) \ c \h)\h = ((deriv f a)\h \ (c \h)\h)" + by (simp add: R.l_distr assms(1) assms(2) c_def deriv_closed h_fact) + then have "((deriv f a) \ c \h)\h = (deriv f a)\h \ (c \(h\h))" + by (simp add: mult_assoc) + then have "((deriv f a) \ c \h)\h = (deriv f a)\h \ (c \(h[^](2::nat)))" + using nat_pow_def[of Zp h "2"] + by (simp add: h_fact) + then show ?thesis + using 2 + by simp + qed + have "h = \" + proof(rule ccontr) + assume "h \ \" + then have "(deriv f a) \ c \h = \" + using 2 3 + by (meson R.m_closed UP_cring.deriv_closed Zp_x_is_UP_cring assms(1) assms(2) c_def h_fact local.integral sum_closed) + then have "(deriv f a) = \ c \h" + using R.l_minus R.minus_equality assms(1) assms(2) c_def deriv_closed h_fact by auto + then have "val_Zp (deriv f a) = val_Zp (c \ h)" + by (meson R.m_closed \deriv f a \ c \ h = \\ assms(1) assms(2) c_def deriv_closed h_fact val_Zp_not_equal_imp_notequal(3)) + then have P: "val_Zp (deriv f a) = val_Zp h + val_Zp c" + by (simp add: c_def h_fact val_Zp_mult) + have "val_Zp (deriv f a) \ val_Zp h " + using P val_pos[of c] c_def + by simp + then have "val_Zp (deriv f a) \ val_Zp (b \ a) " + using h_def by blast + then have "val_Zp (deriv f a) > val_Zp ((pderiv f)\a)" + by (metis (no_types, lifting) A assms(2) h_def h_fact minus_a_inv not_less order_trans val_Zp_of_minus) + then have P0:"val_Zp (deriv f a) > val_Zp (deriv f a)" + by (metis UP_cring.pderiv_eval_deriv Zp_x_is_UP_cring assms(1) assms(2)) + thus False by auto + qed + then show "a = b" + by (simp add: assms(2) h_fact) + qed + show ?thesis + using T1 T2 + by blast +next + case False + have F0: "pderiv f \ a \ \" + apply(rule ccontr) using assms(3) + unfolding val_Zp_def by simp + obtain \ where alpha_def: + "f\\ = \" "\ \ carrier Zp" + "val_Zp (a \ \) > val_Zp ((pderiv f)\a)" + "(val_Zp (a \ \) = val_Zp (divide (f\a) ((pderiv f)\a)))" + "val_Zp ((pderiv f)\\) = val_Zp ((pderiv f)\a)" + using assms hensels_lemma F0 False by blast + have 0: "\x. x \ carrier Zp \ f \ x = \ \ val_Zp (a \ x) > val_Zp (pderiv f \ a) \ val_Zp (pderiv f \ a) \ val_Zp (a \ x) \ x= \" + using alpha_def assms hensels_lemma_unique_root[of f a \] F0 False by blast + have 1: "\ \ carrier Zp \ f \ \ = \ \ val_Zp (a \ \) > val_Zp (pderiv f \ a) \ val_Zp (pderiv f \ a) \ val_Zp (a \ \)" + using alpha_def order_less_le by blast + thus ?thesis + using 0 + by (metis (no_types, hide_lams) R.minus_closed alpha_def(1-3) assms(2) equal_val_Zp val_Zp_ultrametric_eq') +qed + +(**************************************************************************************************) +(**************************************************************************************************) +section\Some Applications of Hensel's Lemma to Root Finding for Polynomials over $\mathbb{Z}_p$\ +(**************************************************************************************************) +(**************************************************************************************************) + +lemma Zp_square_root_criterion: + assumes "p \ 2" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "val_Zp b \ val_Zp a" + assumes "a \ \" + assumes "b \ \" + shows "\y \ carrier Zp. a[^](2::nat) \ \

\b[^](2::nat) = (y [^]\<^bsub>Zp\<^esub> (2::nat))" +proof- + have bounds: "val_Zp a < \" "val_Zp a \ 0" "val_Zp b < \" "val_Zp b \ 0" + using assms(2) assms(3) assms(6) assms(5) val_Zp_def val_pos[of b] val_pos[of a] + by auto + obtain f where f_def: "f = monom Zp_x \ 2 \\<^bsub>Zp_x\<^esub> to_polynomial Zp (\ (a[^](2::nat)\ \

\b[^](2::nat)))" + by simp + have "\ \. f\\ = \ \ \ \ carrier Zp" + proof- + have 0: "f \ carrier Zp_x" + using f_def + by (simp add: X_closed assms(2) assms(3) to_poly_closed) + have 1: "(pderiv f)\a = [(2::nat)] \ \ \ a" + proof- + have "pderiv f = pderiv (monom Zp_x \ 2)" + using assms f_def pderiv_add[of "monom Zp_x \ 2"] to_poly_closed R.nat_pow_closed + pderiv_deg_0 + unfolding to_polynomial_def + using P.nat_pow_closed P.r_zero R.add.inv_closed X_closed Zp_int_inc_closed deg_const monom_term_car pderiv_closed sum_closed + by (metis (no_types, lifting) R.one_closed monom_closed) + then have 20: "pderiv f = monom (Zp_x) ([(2::nat) ] \ \) (1::nat)" + using pderiv_monom[of \ 2] + by simp + have 21: "[(2::nat)] \ \ \ \" + using Zp_char_0'[of 2] by simp + have 22: "(pderiv f)\a = [(2::nat)] \ \ \ (a[^]((1::nat)))" + using 20 + by (simp add: Zp_nat_inc_closed assms(2) to_fun_monom) + then show ?thesis + using assms(2) + by (simp add: cring.cring_simprules(12)) + qed + have 2: "(pderiv f)\a \ \" + using 1 assms + by (metis Zp_char_0' Zp_nat_inc_closed local.integral zero_less_numeral) + have 3: "f\a = \ (\

\b[^](2::nat))" + proof- + have 3: "f\a = + monom (UP Zp) \ 2 \ a \ + to_polynomial Zp (\ (a [^] (2::nat) \ [p] \ \ \ b [^] (2::nat)))\a" + unfolding f_def apply(rule to_fun_plus) + apply (simp add: assms(2) assms(3) to_poly_closed) + apply simp + by (simp add: assms(2)) + have 30: "f\a = a[^](2::nat) \ (a[^](2::nat) \ \

\b[^](2::nat))" + unfolding 3 by (simp add: R.minus_eq assms(2) assms(3) to_fun_monic_monom to_fun_to_poly) + have 31: "f\a = a[^](2::nat) \ a[^](2::nat) \ (\

\b[^](2::nat))" + proof- + have 310: "a[^](2::nat) \ carrier Zp" + using assms(2) pow_closed + by blast + have 311: "\

\(b[^](2::nat)) \ carrier Zp" + by (simp add: assms(3) monom_term_car) + have "\ (a [^] (2::nat)\(\

\ b [^] (2::nat))) = \ (a [^] (2::nat)) \ \ (\

\ (b [^] (2::nat)))" + using 310 311 R.minus_add by blast + then show ?thesis + by (simp add: "30" R.minus_eq add_assoc) + qed + have 32: "f\a = (a[^](2::nat) \ a[^](2::nat)) \ (\

\b[^](2::nat))" + using 31 unfolding a_minus_def + by blast + have 33: "\

\b[^](2::nat) \ carrier Zp" + by (simp add: Zp_nat_inc_closed assms(3) monom_term_car) + have 34: "a[^](2::nat) \ carrier Zp" + using assms(2) pow_closed by blast + then have 34: "(a[^](2::nat) \ a[^](2::nat)) = \ " + by simp + have 35: "f\a = \ \ (\

\b[^](2::nat))" + by (simp add: "32" "34") + then show ?thesis + using 33 unfolding a_minus_def + by (simp add: cring.cring_simprules(3)) + qed + have 4: "f\a \\" + using 3 assms + by (metis R.add.inv_eq_1_iff R.m_closed R.nat_pow_closed Zp.integral Zp_int_inc_closed + mult_zero_r nonzero_pow_nonzero p_natpow_prod_Suc(1) p_pow_nonzero(2)) + have 5: "val_Zp (f\a) = 1 + 2*val_Zp b" + proof- + have "val_Zp (f\a) = val_Zp (\

\b[^](2::nat))" + using 3 Zp_int_inc_closed assms(3) monom_term_car val_Zp_of_minus by presburger + then have "val_Zp (\

\b[^](2::nat)) = 1 + val_Zp (b[^](2::nat))" + by (simp add: assms(3) val_Zp_mult val_Zp_p) + then show ?thesis + using assms(3) assms(6) + using Zp_def \val_Zp (to_fun f a) = val_Zp ([p] \ \ \ b [^] 2)\ not_nonzero_Zp + padic_integers_axioms val_Zp_pow' by fastforce + qed + have 6: "val_Zp ((pderiv f)\a) = val_Zp a" + proof- + have 60: "val_Zp ([(2::nat)] \ \ \ a) = val_Zp ([(2::nat)] \ \) + val_Zp a" + by (simp add: Zp_char_0' assms(2) assms(5) val_Zp_mult ord_of_nonzero(2) ord_pos) + have "val_Zp ([(2::nat)] \ \) = 0" + proof- + have "(2::nat) < p" + using prime assms prime_ge_2_int by auto + then have "(2::nat) mod p = (2::nat)" + by simp + then show ?thesis + by (simp add: val_Zp_p_nat_unit) + qed + then show ?thesis + by (simp add: "1" "60") + qed + then have 7: "val_Zp (f\a) > 2* val_Zp ((pderiv f)\a)" + using bounds 5 assms(4) + by (simp add: assms(5) assms(6) one_eint_def val_Zp_def) + obtain \ where + A0: "f\\ = \" "\ \ carrier Zp" + using hensels_lemma[of f a] "0" "2" "4" "7" assms(2) + by blast + show ?thesis + using A0 by blast + qed + then obtain \ where \_def: "f\\ = \ \ \ \ carrier Zp" + by blast + have "f\\ = \ [^](2::nat) \ (a[^](2::nat)\ \

\b[^](2::nat))" + proof- + have 0: "f\\ = + monom (UP Zp) \ 2 \ \ \ + to_polynomial Zp (\ (a [^] (2::nat) \ [p] \ \ \ b [^] (2::nat)))\\" + unfolding f_def apply(rule to_fun_plus) + apply (simp add: assms(2) assms(3) to_poly_closed) + apply simp + by (simp add: \_def) + thus ?thesis + by (simp add: R.minus_eq \_def assms(2) assms(3) to_fun_monic_monom to_fun_to_poly) + qed + then show ?thesis + by (metis R.r_right_minus_eq Zp_int_inc_closed \_def assms(2) assms(3) monom_term_car pow_closed sum_closed) +qed + +lemma Zp_semialg_eq: + assumes "a \ nonzero Zp" + shows "\y \ carrier Zp. \ \ (\

[^] (3::nat))\ (a [^] (4::nat)) = (y [^] (2::nat))" +proof- + obtain f where f_def: "f = monom Zp_x \ 2 \\<^bsub>Zp_x\<^esub> to_poly (\ (\ \ (\

[^] (3::nat))\ (a [^] (4::nat))))" + by simp + have a_car: "a \ carrier Zp" + by (simp add: nonzero_memE assms) + have "f \ carrier Zp_x" + using f_def + by (simp add: a_car to_poly_closed) + hence 0:"f\\ = \ \ (\ \ (\

[^] (3::nat))\ (a [^] (4::nat)))" + using f_def + by (simp add: R.minus_eq assms nat_pow_nonzero nonzero_mult_in_car p_pow_nonzero' to_fun_monom_plus to_fun_to_poly to_poly_closed) + then have 1: "f\\ = \ (\

[^] (3::nat))\ (a [^] (4::nat))" + unfolding a_minus_def + by (smt R.add.inv_closed R.l_minus R.minus_add R.minus_minus R.nat_pow_closed R.one_closed R.r_neg1 a_car monom_term_car p_pow_nonzero(1)) + then have "val_Zp (f\\) = 3 + val_Zp (a [^] (4::nat))" + using assms val_Zp_mult[of "\

[^] (3::nat)" "(a [^] (4::nat))" ] + val_Zp_p_pow p_pow_nonzero[of "3::nat"] val_Zp_of_minus + by (metis R.l_minus R.nat_pow_closed a_car monom_term_car of_nat_numeral) + then have 2: "val_Zp (f\\) = 3 + 4* val_Zp a" + using assms val_Zp_pow' by auto + have "pderiv f = pderiv (monom Zp_x \ 2)" + using assms f_def pderiv_add[of "monom Zp_x \ 2"] to_poly_closed R.nat_pow_closed pderiv_deg_0 + unfolding to_polynomial_def + by (metis (no_types, lifting) P.r_zero R.add.inv_closed R.add.m_closed R.one_closed + UP_zero_closed a_car deg_const deg_nzero_nzero monom_closed monom_term_car p_pow_nonzero(1)) + then have 3: "pderiv f = [(2::nat)] \ \ \\<^bsub>Zp_x\<^esub> X " + by (metis P.nat_pow_eone R.one_closed Suc_1 X_closed diff_Suc_1 monom_rep_X_pow pderiv_monom') + hence 4: "val_Zp ((pderiv f)\\) = val_Zp ([(2::nat)] \ \ )" + by (metis R.add.nat_pow_eone R.nat_inc_prod R.nat_inc_prod' R.nat_pow_one R.one_closed + Zp_nat_inc_closed \pderiv f = pderiv (monom Zp_x \ 2)\ pderiv_monom to_fun_monom) + have "(2::int) = (int (2::nat))" + by simp + then have 5: "[(2::nat)] \ \ = ([(int (2::nat))] \ \ )" + using add_pow_def int_pow_int + by metis + have 6: "val_Zp ((pderiv f)\\) \ 1" + apply(cases "p = 2") + using "4" "5" val_Zp_p apply auto[1] + proof- + assume "p \ 2" + then have 60: "coprime 2 p" + using prime prime_int_numeral_eq primes_coprime two_is_prime_nat by blast + have 61: "2 < p" + using 60 prime + by (smt \p \ 2\ prime_gt_1_int) + then show ?thesis + by (smt "4" "5" \2 = int 2\ mod_pos_pos_trivial nonzero_closed p_nonzero val_Zp_p val_Zp_p_int_unit val_pos) + qed + have 7: "val_Zp (f\\) \ 3" + proof- + have "eint 4 * val_Zp a \ 0" + using 2 val_pos[of a] + by (metis R.nat_pow_closed a_car assms of_nat_numeral val_Zp_pow' val_pos) + thus ?thesis + using "2" by auto + qed + have "2*val_Zp ((pderiv f)\\) \ 2*1" + using 6 one_eint_def eint_mult_mono' + by (smt \2 = int 2\ eint.distinct(2) eint_ile eint_ord_simps(1) eint_ord_simps(2) mult.commute + ord_Zp_p ord_Zp_p_pow ord_Zp_pow p_nonzero p_pow_nonzero(1) times_eint_simps(1) val_Zp_p val_Zp_pow' val_pos) + hence 8: "2 * val_Zp ((pderiv f)\ \) < val_Zp (f\\)" + using 7 le_less_trans[of "2 * val_Zp ((pderiv f)\ \)" "2::eint" 3] + less_le_trans[of "2 * val_Zp ((pderiv f)\ \)" 3 "val_Zp (f\\)"] one_eint_def + by auto + obtain \ where \_def: "f\\ = \" and \_def' :"\ \ carrier Zp" + using 2 6 7 hensels_lemma' 8 \f \ carrier Zp_x\ by blast + have 0: "(monom Zp_x \ 2) \ \ = \ [^] (2::nat)" + by (simp add: \_def' to_fun_monic_monom) + have 1: "to_poly (\ (\ \ (\

[^] (3::nat))\ (a [^] (4::nat)))) \ \ =\( \ \ (\

[^] (3::nat))\ (a [^] (4::nat)))" + by (simp add: \_def' a_car to_fun_to_poly) + then have "\ [^] (2::nat) \ (\ \ (\

[^] (3::nat))\ (a [^] (4::nat))) = \" + using \_def \_def' + by (simp add: R.minus_eq a_car f_def to_fun_monom_plus to_poly_closed) + then show ?thesis + by (metis R.add.m_closed R.nat_pow_closed R.one_closed R.r_right_minus_eq \_def' a_car monom_term_car p_pow_nonzero(1)) +qed + +lemma Zp_nth_root_lemma: + assumes "a \ carrier Zp" + assumes "a \ \" + assumes "n > 1" + assumes "val_Zp (\ \ a) > 2*val_Zp ([(n::nat)]\ \)" + shows "\ b \ carrier Zp. b[^]n = a" +proof- + obtain f where f_def: "f = monom Zp_x \ n \\<^bsub>Zp_x\<^esub> monom Zp_x (\a) 0" + by simp + have "f \ carrier Zp_x" + using f_def monom_closed assms + by simp + have 0: "pderiv f = monom Zp_x ([n]\ \) (n-1)" + by (simp add: assms(1) f_def pderiv_add pderiv_monom) + have 1: "f \ \ = \ \ a" + using f_def + by (metis R.add.inv_closed R.minus_eq R.nat_pow_one R.one_closed assms(1) to_fun_const to_fun_monom to_fun_monom_plus monom_closed) + have 2: "(pderiv f) \ \ = ([n]\ \)" + using 0 to_fun_monom assms + by simp + have 3: "val_Zp (f \ \) > 2* val_Zp ((pderiv f) \ \)" + using 1 2 assms + by (simp add: val_Zp_def) + have 4: "f \ \ \ \" + using 1 assms(1) assms(2) by auto + have 5: "(pderiv f) \ \ \ \" + using "2" Zp_char_0' assms(3) by auto + obtain \ where beta_def: "\ \ carrier Zp \ f \ \ = \" + using hensels_lemma[of f \] + by (metis "3" "5" R.one_closed \f \ carrier Zp_x\) + then have "(\ [^] n) \ a = \" + using f_def R.add.inv_closed assms(1) to_fun_const[of "\ a"] to_fun_monic_monom[of \ n] to_fun_plus monom_closed + unfolding a_minus_def + by (simp add: beta_def) + then have "\ \ carrier Zp \ \ [^] n = a" + using beta_def nonzero_memE not_eq_diff_nonzero assms(1) pow_closed + by blast + then show ?thesis by blast +qed + +end +end diff --git a/thys/Padic_Ints/Padic_Construction.thy b/thys/Padic_Ints/Padic_Construction.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Padic_Construction.thy @@ -0,0 +1,1560 @@ +theory Padic_Construction +imports "HOL-Number_Theory.Residues" "HOL-Algebra.RingHom" "HOL-Algebra.IntRing" +begin + +type_synonym padic_int = "nat \ int" + +section \Inverse Limit Construction of the $p$-adic Integers\ + +text\ + This section formalizes the standard construction of the $p$-adic integers as the inverse + limit of the finite rings $\mathbb{Z} / p^n \mathbb{Z}$ along the residue maps + $\mathbb{Z} / p^n \mathbb{Z} \mapsto \mathbb{Z} / p^n \mathbb{Z} $ defined by + $x \mapsto x \mod p^m$ when $n \geq m$. This is exposited, for example, in section 7.6 of + \cite{dummit2004abstract}. The other main route for formalization is to first define the + $p$-adic absolute value $|\cdot|_p$ on the rational numbers, and then define the field + $\mathbb{Q}_p$ of $p$-adic numbers as the completion of the rationals under this absolute + value. One can then define the ring of $p$-adic integers $\mathbb{Z}_p$ as the unit ball in + $\mathbb{Q}_p$ using the unique extension of $|\cdot|_p$. There exist advantages and + disadvantages to both approaches. The primary advantage to the absolute value approach is + that the construction can be done generically using existing libraries for completions of + normed fields. There are difficulties associated with performing such a construction in + Isabelle using existing HOL formalizations. The chief issue is that the tools in HOL-Analysis + require that a metric space be a type. If one then wanted to construct the fields + $\mathbb{Q}_p$ as metric spaces, one would have to circumvent the apparent dependence on + the parameter $p$, as Isabelle does not support dependent types. A workaround to this proposed + by José Manuel Rodríguez Caballero on the Isabelle mailing list is to define a typeclass for + fields $\mathbb{Q}_p$ as the completions of the rational numbers with a non-Archimedean absolute + value. By Ostrowski's Theorem, any such absolute value must be a $p$-adic absolute value. We can + recover the parameter $p$ from a completion under one of these absolute values as the cardinality + of the residue field. + + Our approach uses HOL-Algebra, where algebraic structures are constructed as records which carry + the data of the underlying carrier set plus other algebraic operations, and assumptions about + these structures can be organized into locales. This approach is practical for abstract + algebraic reasoning where definitions of structures which are dependent on object-level + parameters are ubiquitous. Using this approach, we define $\mathbb{Z}_p$ directly as an + inverse limit of rings, from which $\mathbb{Q}_p$ can later be defined as the field of fractions. +\ + +subsection\Canonical Projection Maps Between Residue Rings\ + +definition residue :: "int \ int \ int" where +"residue n m = m mod n" + +lemma residue_is_hom_0: + assumes "n > 1" + shows "residue n \ ring_hom \ (residue_ring n)" +proof(rule ring_hom_memI) + have R: "residues n" + by (simp add: assms residues_def) + show "\x. x \ carrier \ \ residue n x \ carrier (residue_ring n)" + using assms residue_def residues.mod_in_carrier residues_def by auto + show " \x y. x \ carrier \ \ y \ carrier \ \ + residue n (x \\<^bsub>\\<^esub> y) = residue n x \\<^bsub>residue_ring n\<^esub> residue n y" + by (simp add: R residue_def residues.mult_cong) + show "\x y. x \ carrier \ \ + y \ carrier \ \ + residue n (x \\<^bsub>\\<^esub> y) = residue n x \\<^bsub>residue_ring n\<^esub> residue n y" + by (simp add: R residue_def residues.res_to_cong_simps(1)) + show "residue n \\<^bsub>\\<^esub> = \\<^bsub>residue_ring n\<^esub>" + by (simp add: R residue_def residues.res_to_cong_simps(4)) +qed + +text\The residue map is a ring homomorphism from $\mathbb{Z}/m\mathbb{Z} \to \mathbb{Z}/n\mathbb{Z}$ when n divides m\ + +lemma residue_is_hom_1: + assumes "n > 1" + assumes "m > 1" + assumes "n dvd m" + shows "residue n \ ring_hom (residue_ring m) (residue_ring n)" +proof(rule ring_hom_memI) + have 0: "residues n" + by (simp add: assms(1) residues_def) + have 1: "residues m" + by (simp add: assms(2) residues_def) + show "\x. x \ carrier (residue_ring m) \ residue n x \ carrier (residue_ring n)" + using assms(1) residue_def residue_ring_def by auto + show "\x y. x \ carrier (residue_ring m) \ + y \ carrier (residue_ring m) \ + residue n (x \\<^bsub>residue_ring m\<^esub> y) = residue n x \\<^bsub>residue_ring n\<^esub> residue n y" + using 0 1 assms by (metis mod_mod_cancel residue_def residues.mult_cong residues.res_mult_eq) + show "\x y. x \ carrier (residue_ring m) + \ y \ carrier (residue_ring m) + \ residue n (x \\<^bsub>residue_ring m\<^esub> y) = residue n x \\<^bsub>residue_ring n\<^esub> residue n y" + using 0 1 assms by (metis mod_mod_cancel residue_def residues.add_cong residues.res_add_eq) + show "residue n \\<^bsub>residue_ring m\<^esub> = \\<^bsub>residue_ring n\<^esub>" + by (simp add: assms(1) residue_def residue_ring_def) +qed + +lemma residue_id: + assumes "x \ carrier (residue_ring n)" + assumes "n \0" + shows "residue n x = x" +proof(cases "n=0") + case True + then show ?thesis + by (simp add: residue_def) +next + case False + have 0: "x \0" + using assms(1) by (simp add: residue_ring_def) + have 1: "x < n" + using assms(1) residue_ring_def by auto + have "x mod n = x" + using 0 1 by simp + then show ?thesis + using residue_def by auto +qed + +text\ + The residue map is a ring homomorphism from + $\mathbb{Z}/p^n\mathbb{Z} \to \mathbb{Z}/p^m\mathbb{Z}$ when $n \geq m$: +\ + +lemma residue_hom_p: + assumes "(n::nat) \ m" + assumes "m >0" + assumes "prime (p::int)" + shows "residue (p^m) \ ring_hom (residue_ring (p^n)) (residue_ring (p^m))" +proof(rule residue_is_hom_1) + show " 1 < p^n" using assms + using prime_gt_1_int by auto + show "1 < p^m" + by (simp add: assms(2) assms(3) prime_gt_1_int) + show "p ^ m dvd p ^ n" using assms(1) + by (simp add: dvd_power_le) +qed + +subsection\Defining the Set of $p$-adic Integers\ + +text\ + The set of $p$-adic integers is the set of all maps $f: \mathbb{N} \to \mathbb{Z}$ which maps + $n \to \{0,...,p^n -1\}$ such that $f m \mod p^{n} = f n$ when $m \geq n$. A p-adic integer $x$ + consists of the data of a residue map $x \mapsto x\mod p^n$ which commutes with further reduction + $\mod p^m$. This formalization is specialized to just the $p$-adics, but this definition would + work essentially as-is for any family of rings and residue maps indexed by a partially + ordered type. +\ + +definition padic_set :: "int \ padic_int set" where +"padic_set p = {f::nat \ int .(\ m::nat. (f m) \ carrier (residue_ring (p^m))) + + \(\(n::nat) (m::nat). n > m \ residue (p^m) (f n) = (f m)) }" + + +lemma padic_set_res_closed: + assumes "f \ padic_set p" + shows "(f m) \ (carrier (residue_ring (p^m)))" + using assms padic_set_def by auto + +lemma padic_set_res_coherent: + assumes "f \ padic_set p" + assumes "n \ m" + assumes "prime p" + shows "residue (p^m) (f n) = (f m)" +proof(cases "n=m") + case True + have "(f m) \ carrier (residue_ring (p^m))" + using assms padic_set_res_closed by blast + then have "residue (p^m) (f m) = (f m)" + by (simp add: residue_def residue_ring_def) + then show ?thesis + using True by blast +next + case False + then show ?thesis + using assms(1) assms(2) padic_set_def by auto +qed + +text\ + A consequence of this formalization is that each $p$-adic number is trivially + defined to take a value of $0$ at $0$: +\ + +lemma padic_set_zero_res: + assumes "prime p" + assumes "f \ (padic_set p)" + shows "f 0 = 0" +proof- + have "f 0 \ carrier (residue_ring 1)" + using assms(1) padic_set_res_closed + by (metis assms(2) power_0) + then show ?thesis + using residue_ring_def by simp +qed + +lemma padic_set_memI: + fixes f :: "padic_int" + assumes "\m. (f m) \ (carrier (residue_ring (p^m)))" + assumes "(\(m::nat) n. (n > m \ (residue (p^m) (f n) = (f m))))" + shows "f \ padic_set (p::int)" + by (simp add: assms(1) assms(2) padic_set_def) + +lemma padic_set_memI': + fixes f :: "padic_int" + assumes "\m. (f m) \ {0..(m::nat) n. n > m \ (f n) mod p^m = (f m)" + shows "f \ padic_set (p::int)" + apply(rule padic_set_memI) + using assms(1) residue_ring_def apply auto[1] + by (simp add: assms(2) residue_def) + + +section\The standard operations on the $p$-adic integers\ + + (**********************************************************************************************) + subsection\Addition\ + (**********************************************************************************************) + +text\Addition and multiplication are defined componentwise on residue rings:\ + +definition padic_add :: "int \ padic_int \ padic_int \ padic_int " + where "padic_add p f g \ (\ n. (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n))" + +lemma padic_add_res: +"(padic_add p f g) n = (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n)" + by (simp add: padic_add_def) + +text\Definition of the $p$-adic additive unit:\ + +definition padic_zero :: "int \ padic_int" where +"padic_zero p \ (\n. 0)" + +lemma padic_zero_simp: +"padic_zero p n = \\<^bsub>residue_ring (p^n)\<^esub>" +"padic_zero p n = 0" + apply (simp add: padic_zero_def residue_ring_def) + using padic_zero_def by auto + +lemma padic_zero_in_padic_set: + assumes "p > 0" + shows "padic_zero p \ padic_set p" + apply(rule padic_set_memI) + by(auto simp: assms padic_zero_def residue_def residue_ring_def) + +text\$p$-adic additive inverses:\ + +definition padic_a_inv :: "int \ padic_int \ padic_int" where +"padic_a_inv p f \ \ n. \\<^bsub>residue_ring (p^n)\<^esub> (f n)" + +lemma padic_a_inv_simp: +"padic_a_inv p f n\ \\<^bsub>residue_ring (p^n)\<^esub> (f n)" + by (simp add: padic_a_inv_def) + +lemma padic_a_inv_simp': + assumes "prime p" + assumes "f \ padic_set p" + assumes "n >0" + shows "padic_a_inv p f n = (if n=0 then 0 else (- (f n)) mod (p^n))" +proof- + have "residues (p^n)" + by (simp add: assms(1) assms(3) prime_gt_1_int residues.intro) + then show ?thesis + using residue_ring_def padic_a_inv_def residues.res_neg_eq + by auto +qed + +text\ + We show that \<^const>\padic_set\ is closed under additive inverses. Note that we have to treat the + case of residues at $0$ separately. +\ + +lemma residue_1_prop: +"\\<^bsub>residue_ring 1\<^esub> \\<^bsub>residue_ring 1\<^esub> = \\<^bsub>residue_ring 1\<^esub>" +proof- + let ?x = "\\<^bsub>residue_ring 1\<^esub>" + let ?y = "\\<^bsub>residue_ring 1\<^esub> \\<^bsub>residue_ring 1\<^esub>" + let ?G = "add_monoid (residue_ring 1)" + have P0:" ?x \\<^bsub>residue_ring 1\<^esub> ?x = ?x" + by (simp add: residue_ring_def) + have P1: "?x \ carrier (residue_ring 1)" + by (simp add: residue_ring_def) + have "?x \ carrier ?G \ ?x \\<^bsub>?G\<^esub> ?x = \\<^bsub>?G\<^esub> \ ?x \\<^bsub>?G\<^esub> ?x = \\<^bsub>?G\<^esub>" + using P0 P1 by auto + then show ?thesis + by (simp add: m_inv_def a_inv_def residue_ring_def) +qed + +lemma residue_1_zero: + "residue 1 n = 0" + by (simp add: residue_def) + +lemma padic_a_inv_in_padic_set: + assumes "f \ padic_set p" + assumes "prime (p::int)" + shows "(padic_a_inv p f) \ padic_set p" +proof(rule padic_set_memI) + show "\m. padic_a_inv p f m \ carrier (residue_ring (p ^ m))" + proof- + fix m + show "padic_a_inv p f m \ carrier (residue_ring (p ^ m))" + proof- + have P0: "padic_a_inv p f m = \\<^bsub>residue_ring (p^m)\<^esub> (f m)" + using padic_a_inv_def by simp + then show ?thesis + by (metis (no_types, lifting) assms(1) assms(2) cring.cring_simprules(3) neq0_conv + one_less_power padic_set_res_closed padic_set_zero_res power_0 prime_gt_1_int residue_1_prop + residue_ring_def residues.cring residues.intro ring.simps(1)) + qed + qed + show "\m n. m < n \ residue (p ^ m) (padic_a_inv p f n) = padic_a_inv p f m" + proof- + fix m n::nat + assume "m < n" + show "residue (p ^ m) (padic_a_inv p f n) = padic_a_inv p f m" + proof(cases "m=0") + case True + then have 0: "residue (p ^ m) (padic_a_inv p f n) = 0" using residue_1_zero + by simp + have "f m = 0" + using assms True padic_set_def residue_ring_def padic_set_zero_res + by auto + then have 1: "padic_a_inv p f m = 0" using residue_1_prop assms + by (simp add: True padic_a_inv_def residue_ring_def) + then show ?thesis using 0 1 + by simp + next + case False + have 0: "f n \ carrier (residue_ring (p^n)) " + using assms(1) padic_set_res_closed by auto + have 1: "padic_a_inv p f n = \\<^bsub>residue_ring (p^n)\<^esub> (f n)" using padic_a_inv_def + by simp + have 2: "padic_a_inv p f m = \\<^bsub>residue_ring (p^m)\<^esub> (f m)" using False padic_a_inv_def + by simp + have 3: "residue (p ^ m) \ ring_hom (residue_ring (p ^ n)) (residue_ring (p ^ m))" + using residue_hom_p False \m < n\ assms(2) by auto + have 4: " cring (residue_ring (p ^ n))" + using \m < n\ assms(2) prime_gt_1_int residues.cring residues.intro by auto + have 5: " cring (residue_ring (p ^ m))" + using False assms(2) prime_gt_1_int residues.cring residues.intro by auto + have "ring_hom_cring (residue_ring (p ^ n)) (residue_ring (p ^ m)) (residue (p ^ m))" + using 3 4 5 UnivPoly.ring_hom_cringI by blast + then show ?thesis using 0 1 2 ring_hom_cring.hom_a_inv + by (metis \m < n\ assms(1) assms(2) less_imp_le_nat padic_set_res_coherent) + qed + qed + qed + + (**********************************************************************************************) + subsection\Multiplication\ + (**********************************************************************************************) + +definition padic_mult :: "int \ padic_int \ padic_int \ padic_int" + where "padic_mult p f g \ (\ n. (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n))" + +lemma padic_mult_res: +"(padic_mult p f g) n = (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n)" + by (simp add: padic_mult_def) + +text\Definition of the $p$-adic multiplicative unit:\ + +definition padic_one :: "int \ padic_int" where +"padic_one p \ (\n.(if n=0 then 0 else 1))" + +lemma padic_one_simp: + assumes "n >0" + shows "padic_one p n = \\<^bsub>residue_ring (p^n)\<^esub>" + "padic_one p n = 1" + apply (simp add: assms padic_one_def residue_ring_def) + using assms padic_one_def by auto + +lemma padic_one_in_padic_set: + assumes "prime p" + shows "padic_one p \ padic_set p" + apply(rule padic_set_memI) + by(auto simp : assms padic_one_def prime_gt_1_int residue_def residue_ring_def) + +lemma padic_simps: +"padic_zero p n = \\<^bsub>residue_ring (p^n)\<^esub>" +"padic_a_inv p f n \ \\<^bsub>residue_ring (p^n)\<^esub> (f n)" +"(padic_mult p f g) n = (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n)" +"(padic_add p f g) n = (f n) \\<^bsub>(residue_ring (p^n))\<^esub> (g n)" +"n>0 \padic_one p n = \\<^bsub>residue_ring (p^n)\<^esub>" + apply (simp add: padic_zero_simp) + apply (simp add: padic_a_inv_simp) + apply (simp add: padic_mult_def) + apply (simp add: padic_add_res) + using padic_one_simp by auto + +lemma residue_1_mult: + assumes "x \ carrier (residue_ring 1)" + assumes "y \ carrier (residue_ring 1)" + shows "x \\<^bsub>residue_ring 1\<^esub> y = 0" + by (simp add: residue_ring_def) + +lemma padic_mult_in_padic_set: + assumes "f \ (padic_set p)" + assumes "g \ (padic_set p)" + assumes "prime p" + shows "(padic_mult p f g)\ (padic_set p)" +proof(rule padic_set_memI') + show "\m. padic_mult p f g m \ {0..

m n. m < n \ padic_mult p f g n mod p ^ m = padic_mult p f g m" + proof- + fix m n::nat + assume A: "m < n" + then show "padic_mult p f g n mod p ^ m = padic_mult p f g m" + proof(cases "m=0") + case True + then show ?thesis + by (metis assms(1) assms(2) mod_by_1 padic_mult_def padic_set_res_closed power_0 residue_1_mult) + next + case False + have 0:"residue (p ^ m) \ ring_hom (residue_ring (p^n)) (residue_ring (p^m))" + using A residue_hom_p assms False by auto + have 1:"f n \ carrier (residue_ring (p^n))" + using assms(1) padic_set_res_closed by auto + have 2:"g n \ carrier (residue_ring (p^n))" + using assms(2) padic_set_res_closed by auto + have 3: "residue (p^m) (f n \\<^bsub>residue_ring (p^n)\<^esub> g n) + = f m \\<^bsub>residue_ring (p^m)\<^esub> g m" + using "0" "1" "2" A assms(1) assms(2) assms(3) less_imp_le of_nat_power padic_set_res_coherent + by (simp add: assms(2) ring_hom_mult) + then show ?thesis + using ring_hom_mult padic_simps[simp] residue_def + by auto + qed + qed +qed + +section\The $p$-adic Valuation\ + +text\This section defines the integer-valued $p$-adic valuation. Maps $0$ to $-1$ for now, otherwise is correct. We want the valuation to be integer-valued, but in practice we know it will always be positive. When we extend the valuation from the $p$-adic integers to the $p$-adic field we will have elements of negative valuation. \ + +definition padic_val :: "int \ padic_int \ int" where +"padic_val p f \ if (f = padic_zero p) then -1 else int (LEAST k::nat. (f (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + +text\Characterization of $padic\_val$ on nonzero elements\ + +lemma val_of_nonzero: + assumes "f \ padic_set p" + assumes "f \ padic_zero p" + assumes "prime p" + shows "f (nat (padic_val p f) + 1) \ \\<^bsub>residue_ring (p^((nat (padic_val p f) + 1)))\<^esub>" + "f (nat (padic_val p f)) = \\<^bsub>residue_ring (p^((nat (padic_val p f))))\<^esub>" + "f (nat (padic_val p f) + 1) \ 0" + "f (nat (padic_val p f)) = 0" +proof- + let ?vf = "padic_val p f" + have 0: "?vf =int (LEAST k::nat. (f (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + using assms(2) padic_val_def by auto + have 1: "(\ k::nat. (f (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + proof- + obtain k where 1: "(f k) \ (padic_zero p k)" + using assms(2) by (meson ext) + have 2: "k \ 0" + proof + assume "k=0" + then have "f k = 0" + using assms padic_set_zero_res by blast + then show False + using padic_zero_def 1 by simp + qed + then obtain m where "k = Suc m" + by (meson lessI less_Suc_eq_0_disj) + then have "(f (Suc m)) \ \\<^bsub>residue_ring (p^(Suc m))\<^esub>" + using "1" padic_zero_simp by simp + then show ?thesis + by auto + qed + then have "(f (Suc (nat ?vf))) \ \\<^bsub>residue_ring (p^(Suc (nat ?vf)))\<^esub>" + using 0 by (metis (mono_tags, lifting) LeastI_ex nat_int) + then show C0: "f (nat (padic_val p f) + 1) \ \\<^bsub>residue_ring (p^((nat (padic_val p f) + 1)))\<^esub>" + using 0 1 by simp + show C1: "f (nat (padic_val p f)) = \\<^bsub>residue_ring (p^((nat (padic_val p f))))\<^esub>" + proof(cases "(padic_val p f) = 0") + case True + then show ?thesis + using assms(1) assms(3) padic_set_zero_res residue_ring_def by auto + next + case False + have "\ f (nat (padic_val p f)) \ \\<^bsub>residue_ring (p ^ nat (padic_val p f))\<^esub>" + proof + assume "f (nat (padic_val p f)) \ \\<^bsub>residue_ring (p ^ nat (padic_val p f))\<^esub>" + obtain k where " (Suc k) = (nat (padic_val p f))" using False + using "0" gr0_conv_Suc by auto + then have "?vf \ int (LEAST k::nat. (f (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + using False by (metis (mono_tags, lifting) Least_le + \f (nat (padic_val p f)) \ \\<^bsub>residue_ring (p ^ nat (padic_val p f))\<^esub>\ + add_le_same_cancel2 nat_int not_one_le_zero plus_1_eq_Suc) + then show False using "0" by blast + qed + then show "f (nat (padic_val p f)) = \\<^bsub>residue_ring (p ^ nat (padic_val p f))\<^esub>" by auto + qed + show "f (nat (padic_val p f) + 1) \ 0" + using C0 residue_ring_def + by auto + show "f (nat (padic_val p f)) = 0" + by (simp add: C1 residue_ring_def) +qed + +text\If $x \mod p^{n+1} \neq 0$, then $n \geq val x$.\ + +lemma below_val_zero: + assumes "prime p" + assumes "x \ (padic_set p)" + assumes "x (Suc n) \ \\<^bsub>residue_ring (p^(Suc n))\<^esub>" + shows "int n \ (padic_val p x )" +proof(cases "x = padic_zero p") + case True + then show ?thesis + using assms(3) padic_zero_simp by blast +next + case False + then have "padic_val p x = int (LEAST k::nat. x (Suc k) \ \\<^bsub>residue_ring (p ^ Suc k)\<^esub>)" + using padic_val_def by auto + then show "of_nat n \ (padic_val p x )" + by (metis (mono_tags, lifting) Least_le assms(3) nat_int nat_le_iff) +qed + +text\If $n < val x$ then $x \mod p^n = 0$:\ + +lemma zero_below_val: + assumes "prime p" + assumes "x \ padic_set p" + assumes "n \ padic_val p x" + shows "x n = \\<^bsub>residue_ring (p^n)\<^esub>" + "x n = 0" +proof- + show "x n = \\<^bsub>residue_ring (p ^ n)\<^esub>" + proof(cases "n=0") + case True + then have "x 0 \carrier (residue_ring (p^0))" + using assms(2) padic_set_res_closed by blast + then show ?thesis + by (simp add: True residue_ring_def) + next + case False + show ?thesis + proof(cases "x = padic_zero p") + case True + then show ?thesis + by (simp add: padic_zero_simp) + next + case F: False + then have A: "padic_val p x = int (LEAST k::nat. (x (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + using padic_val_def by auto + have "\ (x n) \ \\<^bsub>residue_ring (p^n)\<^esub>" + proof + assume "(x n) \ \\<^bsub>residue_ring (p^n)\<^esub>" + obtain k where "n = Suc k" + using False old.nat.exhaust by auto + then have "k \ padic_val p x" using A + using \x n \ \\<^bsub>residue_ring (p ^ n)\<^esub>\ assms(1) assms(2) below_val_zero by blast + then have "n > padic_val p x" + using \n = Suc k\ by linarith + then show False using assms(3) + by linarith + qed + then show ?thesis + by simp + qed + qed + show "x n = 0" + by (simp add: \x n = \\<^bsub>residue_ring (p ^ n)\<^esub>\ residue_ring_def) +qed + +text\Zero is the only element with valuation equal to $-1$:\ + +lemma val_zero: + assumes P: "f \ (padic_set p)" + shows "padic_val p f = -1 \ (f = (padic_zero p))" +proof + show "padic_val p f = -1 \ (f = (padic_zero p))" + proof + assume A:"padic_val p f = -1" + fix k + show "f k = padic_zero p k" + proof- + have "f k \ padic_zero p k \ False" + proof- + assume A0: " f k \ padic_zero p k" + have False + proof- + have "f 0 \ carrier (residue_ring 1)" using P padic_set_def + by (metis (no_types, lifting) mem_Collect_eq power_0) + then have "f 0 = \\<^bsub>residue_ring (p^0)\<^esub>" + by (simp add: residue_ring_def) + then have "k>0" + using A0 gr0I padic_zero_def + by (metis padic_zero_simp) + then have "(LEAST k. 0 < k \ f (Suc k) \ padic_zero p (Suc k)) \0 " + by simp + then have "padic_val p f \0" + using A0 padic_val_def by auto + then show ?thesis using A0 by (simp add: A) + qed + then show ?thesis by blast + qed + then show ?thesis + by blast + qed + qed + assume B: "f = padic_zero p" + then show "padic_val p f = -1" + using padic_val_def by simp +qed + +text\ + The valuation turns multiplication into integer addition on nonzero elements. Note that this i + the first instance where we need to explicity use the fact that $p$ is a prime. +\ + +lemma val_prod: + assumes "prime p" + assumes "f \ (padic_set p)" + assumes "g \ (padic_set p)" + assumes "f \ padic_zero p" + assumes "g \ padic_zero p" + shows "padic_val p (padic_mult p f g) = padic_val p f + padic_val p g" +proof- + let ?vp = "padic_val p (padic_mult p f g)" + let ?vf = "padic_val p f" + let ?vg = "padic_val p g" + have 0: "f (nat ?vf + 1) \ \\<^bsub>residue_ring (p^(nat ?vf + 1))\<^esub>" + using assms(2) assms(4) val_of_nonzero assms(1) by blast + have 1: "g (nat ?vg + 1) \ \\<^bsub>residue_ring (p^(nat ?vg + 1))\<^esub>" + using assms(3) assms(5) val_of_nonzero assms(1) by blast + have 2: "f (nat ?vf) = \\<^bsub>residue_ring (p^(nat ?vf))\<^esub>" + using assms(1) assms(2) assms(4) val_of_nonzero(2) by blast + have 3: "g (nat ?vg) = \\<^bsub>residue_ring (p^(nat ?vg))\<^esub>" + using assms(1) assms(3) assms(5) val_of_nonzero(2) by blast + let ?nm = "((padic_mult p f g) (Suc (nat (?vf + ?vg))))" + let ?n = "(f (Suc (nat (?vf + ?vg))))" + let ?m = "(g (Suc (nat (?vf + ?vg))))" + have A: "?nm = ?n \\<^bsub>residue_ring (p^((Suc (nat (?vf + ?vg))))) \<^esub> ?m" + using padic_mult_def by simp + have 5: "f (nat ?vf + 1) = residue (p^(nat ?vf + 1)) ?n" + proof- + have "(Suc (nat (?vf + ?vg))) \ (nat ?vf + 1)" + by (simp add: assms(5) padic_val_def) + then have "f (nat ?vf + 1) = residue (p^(nat ?vf + 1)) (f (Suc (nat (?vf + ?vg))))" + using assms(1) assms(2) padic_set_res_coherent by presburger + then show ?thesis by auto + qed + have 6: "f (nat ?vf) = residue (p^(nat ?vf)) ?n" + using add.commute assms(1) assms(2) assms(5) int_nat_eq nat_int + nat_le_iff not_less_eq_eq padic_set_res_coherent padic_val_def plus_1_eq_Suc by auto + have 7: "g (nat ?vg + 1) = residue (p^(nat ?vg + 1)) ?m" + proof- + have "(Suc (nat (?vf + ?vg))) \ (nat ?vg + 1)" + by (simp add: assms(4) padic_val_def) + then have "g (nat ?vg + 1) = residue (p^(nat ?vg + 1)) (g (Suc (nat (?vf + ?vg))))" + using assms(1) assms(3) padic_set_res_coherent by presburger + then show ?thesis by auto + qed + have 8: "g (nat ?vg) = residue (p^(nat ?vg)) ?m" + proof- + have "(Suc (nat (?vf + ?vg))) \ (nat ?vg)" + by (simp add: assms(4) padic_val_def) + then have "g (nat ?vg) = residue (p^(nat ?vg)) (g (Suc (nat (?vf + ?vg))))" + using assms(1) assms(3) padic_set_res_coherent by presburger + then show ?thesis by auto + qed + have 9: "f (nat ?vf) = 0" + by (simp add: "2" residue_ring_def) + have 10: "g (nat ?vg) = 0" + by (simp add: "3" residue_ring_def) + have 11: "f (nat ?vf + 1) \ 0" + using "0" residue_ring_def by auto + have 12: "g (nat ?vg + 1) \0" + using "1" residue_ring_def by auto + have 13:"\i. ?n = i*p^(nat ?vf) \ \ p dvd (nat i)" + proof- + have "residue (p^(nat ?vf)) (?n) = f (nat ?vf)" + by (simp add: "6") + then have P0: "residue (p^(nat ?vf)) (?n) = 0" + using "9" by linarith + have "residue (p^(nat ?vf + 1)) (?n) = f (nat ?vf + 1)" + using "5" by linarith + then have P1: "residue (p^(nat ?vf + 1)) (?n) \ 0" + using "11" by linarith + have P2: "?n mod (p^(nat ?vf)) = 0" + using P0 residue_def by auto + have P3: "?n mod (p^(nat ?vf + 1)) \ 0" + using P1 residue_def by auto + have "p^(nat ?vf) dvd ?n" + using P2 by auto + then obtain i where A0:"?n = i*(p^(nat ?vf))" + by fastforce + have "?n \ carrier (residue_ring (p^(Suc (nat (?vf + ?vg)))))" + using assms(2) padic_set_res_closed by blast + then have "?n \0" + by (simp add: residue_ring_def) + then have NN:"i \ 0" + proof- + have S0:"?n \0" + using \0 \ f (Suc (nat (padic_val p f + padic_val p g)))\ by blast + have S1:"(p^(nat ?vf)) > 0" + using assms(1) prime_gt_0_int zero_less_power by blast + have "\ i<0" + proof + assume "i < 0" + then have "?n < 0" + using S1 A0 by (metis mult.commute times_int_code(1) zmult_zless_mono2) + then show False + using S0 by linarith + qed + then show ?thesis by auto + qed + have A1: "\ p dvd (nat i)" + proof + assume "p dvd nat i" + then obtain j where "nat i = j*p" + by fastforce + then have "?n = j*p*(p^(nat ?vf))" using A0 NN + by simp + then show False + using P3 by auto + qed + then show ?thesis + using A0 by blast + qed + have 14:"\ i. ?m = i*p^(nat ?vg) \ \ p dvd (nat i)" + proof- + have "residue (p^(nat ?vg)) (?m) = g (nat ?vg)" + by (simp add: "8") + then have P0: "residue (p^(nat ?vg)) (?m) = 0" + using "10" by linarith + have "residue (p^(nat ?vg + 1)) (?m) = g (nat ?vg + 1)" + using "7" by auto + then have P1: "residue (p^(nat ?vg + 1)) (?m) \ 0" + using "12" by linarith + have P2: "?m mod (p^(nat ?vg)) = 0" + using P0 residue_def by auto + have P3: "?m mod (p^(nat ?vg + 1)) \ 0" + using P1 residue_def by auto + have "p^(nat ?vg) dvd ?m" + using P2 by auto + then obtain i where A0:"?m = i*(p^(nat ?vg))" + by fastforce + have "?m \ carrier (residue_ring (p^(Suc (nat (?vf + ?vg)))))" + using assms(3) padic_set_res_closed by blast + then have S0: "?m \0" + by (simp add: residue_ring_def) + then have NN:"i \ 0" + using 0 assms(1) prime_gt_0_int[of p] zero_le_mult_iff zero_less_power[of p] + by (metis A0 linorder_not_less) + have A1: "\ p dvd (nat i)" + proof + assume "p dvd nat i" + then obtain j where "nat i = j*p" + by fastforce + then have "?m = j*p*(p^(nat ?vg))" using A0 NN + by (metis int_nat_eq ) + then show False + using P3 by auto + qed + then show ?thesis + by (metis (no_types, lifting) A0) + qed + obtain i where I:"?n = i*p^(nat ?vf) \ \ p dvd (nat i)" + using "13" by blast + obtain j where J:"?m = j*p^(nat ?vg) \ \ p dvd (nat j)" + using "14" by blast + let ?i = "(p^(Suc (nat (?vf + ?vg))))" + have P:"?nm mod ?i = ?n*?m mod ?i" + proof- + have P1:"?nm = (?n \\<^bsub>residue_ring ?i \<^esub> ?m)" + using A by simp + have P2:"(?n \\<^bsub>residue_ring ?i \<^esub> ?m) = (residue ?i (?n)) \\<^bsub>residue_ring ?i\<^esub> (residue ?i (?m))" + using assms(1) assms(2) assms(3) padic_set_res_closed prime_ge_0_int residue_id by presburger + then have P3:"(?n \\<^bsub>residue_ring ?i \<^esub> ?m) = (residue ?i (?n*?m))" + by (metis monoid.simps(1) residue_def residue_ring_def) + then show ?thesis + by (simp add: P1 residue_def) + qed + then have 15: "?nm mod ?i = i*j*p^((nat ?vf) +(nat ?vg)) mod ?i" + by (simp add: I J mult.assoc mult.left_commute power_add) + have 16: "\ p dvd (i*j)" using 13 14 + using I J assms(1) prime_dvd_mult_iff + by (metis dvd_0_right int_nat_eq) + have 17: "((nat ?vf) +(nat ?vg)) < (Suc (nat (?vf + ?vg)))" + by (simp add: assms(4) assms(5) nat_add_distrib padic_val_def) + have 18:"?nm mod ?i \0" + proof- + have A0:"\ p^((Suc (nat (?vf + ?vg)))) dvd p^((nat ?vf) +(nat ?vg)) " + using 17 + by (metis "16" assms(1) dvd_power_iff dvd_trans less_int_code(1) linorder_not_less one_dvd prime_gt_0_int) + then have A1: "p^((nat ?vf) +(nat ?vg)) mod ?i \ 0" + using dvd_eq_mod_eq_0 + by auto + have "\ p^((Suc (nat (?vf + ?vg)))) dvd i*j*p^((nat ?vf) +(nat ?vg)) " + using 16 A0 assms(1) assms(4) assms(5) nat_int_add padic_val_def by auto + then show ?thesis + using "15" by force + qed + have 19: "(?nm mod ?i ) mod (p^(nat ?vf + nat ?vg)) = i*j*p^((nat ?vf) +(nat ?vg)) mod (p^(nat ?vf + nat ?vg))" + using 15 by (simp add: assms(4) assms(5) nat_add_distrib padic_val_def) + have 20: "?nm mod (p^(nat ?vf + nat ?vg)) = 0" + proof- + have "(?nm mod ?i ) mod (p^(nat ?vf + nat ?vg)) = 0" + using 19 + by simp + then show ?thesis + using "17" assms(1) int_nat_eq mod_mod_cancel[of "p^(nat ?vf + nat ?vg)" ?i] + mod_pos_pos_trivial + by (metis le_imp_power_dvd less_imp_le_nat) + qed + have 21: "(padic_mult p f g) \ padic_zero p" + proof + assume "(padic_mult p f g) = padic_zero p" + then have "(padic_mult p f g) (Suc (nat (padic_val p f + padic_val p g))) = padic_zero p (Suc (nat (padic_val p f + padic_val p g)))" + by simp + then have "?nm = (padic_zero p (Suc (nat (padic_val p f + padic_val p g))))" + by blast + then have "?nm = 0" + by (simp add: padic_zero_def) + then show False + using "18" by auto + qed + have 22: "(padic_mult p f g)\ (padic_set p)" + using assms(1) assms(2) assms(3) padic_mult_in_padic_set by blast + have 23: "\ j. j < Suc (nat (padic_val p f + padic_val p g)) \ (padic_mult p f g) j = \\<^bsub>residue_ring (p^j)\<^esub>" + proof- + fix k + let ?j = "Suc (nat (padic_val p f + padic_val p g))" + assume P: "k < ?j" + show "(padic_mult p f g) k = \\<^bsub>residue_ring (p^k)\<^esub>" + proof- + have P0: "(padic_mult p f g) (nat ?vf + nat ?vg) = \\<^bsub>residue_ring (p^(nat ?vf + nat ?vg))\<^esub>" + proof- + let ?k = "(nat ?vf + nat ?vg)" + have "((padic_mult p f g) ?k) = residue (p^?k) ((padic_mult p f g) ?k) " + using P 22 padic_set_res_coherent by (simp add: assms(1) prime_gt_0_nat) + then have "((padic_mult p f g) ?k) = residue (p^?k) ?nm" + using "17" "22" assms(1) padic_set_res_coherent by fastforce + then have "((padic_mult p f g) ?k) = residue (p^?k) ?nm" + by (simp add: residue_def) + then have "((padic_mult p f g) ?k) = residue (p^?k) 0" + using "20" residue_def by auto + then show ?thesis + by (simp add: residue_def residue_ring_def) + qed + then show ?thesis + proof(cases "k = (nat ?vf + nat ?vg)") + case True then show ?thesis + using P0 by blast + next + case B: False + then show ?thesis + proof(cases "k=0") + case True + then show ?thesis + using "22" assms(1) padic_set_zero_res residue_ring_def by auto + next + case C: False + then have "((padic_mult p f g) k) = residue (p^k) ((padic_mult p f g) (nat ?vf + nat ?vg)) " + using B P 22 padic_set_res_coherent by (simp add: assms(1) assms(4) assms(5) padic_val_def prime_gt_0_nat) + then have S: "((padic_mult p f g) k) = residue (p^k) \\<^bsub>residue_ring (p^((nat ?vf + nat ?vg)))\<^esub>" + by (simp add: P0) + have "residue (p^k) \ ring_hom (residue_ring (p^((nat ?vf + nat ?vg)))) (residue_ring (p^k))" + using B P C residue_hom_p + using assms(1) assms(4) assms(5) less_Suc0 nat_int not_less_eq of_nat_power padic_val_def prime_nat_int_transfer by auto + then show ?thesis using S + using P0 padic_zero_def padic_zero_simp residue_def by auto + qed + qed + qed +qed + have 24: "(padic_mult p f g) (Suc (nat ?vf + nat ?vg)) \ \\<^bsub>residue_ring ((p ^ Suc (nat (padic_val p f + padic_val p g))))\<^esub>" + by (metis (no_types, lifting) "18" A P assms(4) assms(5) monoid.simps(1) nat_int nat_int_add padic_val_def residue_ring_def ring.simps(1)) + have 25: "padic_val p (padic_mult p f g) = int (LEAST k::nat. ((padic_mult p f g) (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + using padic_val_def 21 by auto + have 26:"(nat (padic_val p f + padic_val p g)) \ {k::nat. ((padic_mult p f g) (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>}" using 24 + using "18" assms(1) prime_gt_0_nat + by (metis (mono_tags, lifting) mem_Collect_eq mod_0 residue_ring_def ring.simps(1)) + have 27: "\ j. j < (nat (padic_val p f + padic_val p g)) \ + j \ {k::nat. ((padic_mult p f g) (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>}" + by (simp add: "23") + have "(nat (padic_val p f + padic_val p g)) = (LEAST k::nat. ((padic_mult p f g) (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>) " + proof- + obtain P where C0: "P= (\ k. ((padic_mult p f g) (Suc k)) \ \\<^bsub>residue_ring (p^(Suc k))\<^esub>)" + by simp + obtain x where C1: "x = (nat (padic_val p f + padic_val p g))" + by blast + have C2: "P x" + using "26" C0 C1 by blast + have C3:"\ j. j< x \ \ P j" + using C0 C1 by (simp add: "23") + have C4: "\ j. P j \ x \j" + using C3 le_less_linear by blast + have "x = (LEAST k. P k)" + using C2 C4 Least_equality by auto + then show ?thesis using C0 C1 by auto + qed + then have "padic_val p (padic_mult p f g) = (nat (padic_val p f + padic_val p g))" + using "25" by linarith + then show ?thesis + by (simp add: assms(4) assms(5) padic_val_def) + +qed + +section\Defining the Ring of $p$-adic Integers:\ + +definition padic_int :: "int \ padic_int ring" + where "padic_int p \ \carrier = (padic_set p), + Group.monoid.mult = (padic_mult p), one = (padic_one p), + zero = (padic_zero p), add = (padic_add p)\" + +lemma padic_int_simps: + "\\<^bsub>padic_int p\<^esub> = padic_one p" + "\\<^bsub>padic_int p\<^esub> = padic_zero p" + "(\\<^bsub>padic_int p\<^esub>) = padic_add p" + "(\\<^bsub>padic_int p\<^esub>) = padic_mult p" + "carrier (padic_int p) = padic_set p" + unfolding padic_int_def by auto + +lemma residues_n: + assumes "n \ 0" + assumes "prime p" + shows "residues (p^n)" +proof + have "p > 1" using assms(2) + using prime_gt_1_int by auto + then show " 1 < p ^ n " + using assms(1) by auto +qed + +text\$p$-adic multiplication is associative\ + +lemma padic_mult_assoc: +assumes "prime p" +shows "\x y z. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + z \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" +proof- + fix x y z + assume Ax: " x \ carrier (padic_int p)" + assume Ay: " y \ carrier (padic_int p)" + assume Az: " z \ carrier (padic_int p)" + show "x \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" + proof + fix n + show "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n = (x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)) n" + proof(cases "n=0") + case True + then show ?thesis using padic_int_simps + by (metis Ax Ay Az assms padic_mult_in_padic_set padic_set_zero_res) + next + case False + then have "residues (p^n)" + by (simp add: assms residues_n) + then show ?thesis + using residues.cring padic_set_res_closed padic_mult_in_padic_set Ax Ay Az padic_mult_res + by (simp add: cring.cring_simprules(11) padic_int_def) + qed + qed +qed + +text\The $p$-adic integers are closed under addition:\ + +lemma padic_add_closed: + assumes "prime p" + shows "\x y. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>(padic_int p)\<^esub> y \ carrier (padic_int p)" +proof + fix x::"padic_int" + fix y::"padic_int" + assume Px: "x \carrier (padic_int p) " + assume Py: "y \carrier (padic_int p)" + show "x \\<^bsub>(padic_int p)\<^esub> y \ carrier (padic_int p)" + proof- + let ?f = "x \\<^bsub>(padic_int p)\<^esub> y" + have 0: "(\(m::nat). (?f m) \ (carrier (residue_ring (p^m))))" + proof fix m + have A1 : "?f m = (x m) \\<^bsub>(residue_ring (p^m))\<^esub> (y m)" + by (simp add: padic_int_def padic_add_def) + have A2: "(x m) \(carrier (residue_ring (p^m)))" + using Px by (simp add: padic_int_def padic_set_def) + have A3: "(y m) \(carrier (residue_ring (p^m)))" + using Py by (simp add: padic_int_def padic_set_def) + then show "(?f m) \ (carrier (residue_ring (p^m)))" + using A1 assms of_nat_0_less_iff prime_gt_0_nat residue_ring_def by force + qed + have 1: "(\(n::nat) (m::nat). (n > m \ (residue (p^m) (?f n) = (?f m))))" + proof + fix n::nat + show "(\(m::nat). (n > m \ (residue (p^m) (?f n) = (?f m))))" + proof + fix m::nat + show "(n > m \ (residue (p^m) (?f n) = (?f m)))" + proof + assume A: "m < n" + show "(residue (p^m) (?f n) = (?f m))" + proof(cases "m = 0") + case True + then have A0: "(residue (p^m) (?f n)) = 0" + by (simp add: residue_1_zero) + have A1: "?f m = 0" using True + by (simp add: padic_add_res padic_int_simps(3) residue_ring_def) + then show ?thesis + using A0 by linarith + next + case False + then have "m \0" using A by linarith + have D: "p^n mod p^m = 0" using A + by (simp add: le_imp_power_dvd) + let ?LHS = "residue (p ^ m) ((x \\<^bsub>padic_int p\<^esub> y) n)" + have A0: "?LHS = residue (p ^ m) ((x n)\\<^bsub>residue_ring (p^n)\<^esub>( y n))" + by (simp add: padic_int_def padic_add_def) + have "residue (p^m) \ ring_hom (residue_ring ((p^n))) (residue_ring ((p^m)))" + using A False assms residue_hom_p by auto + then have "residue (p ^ m) ((x n)\\<^bsub>residue_ring (p^n)\<^esub>( y n)) = (residue (p ^ m) (x n))\\<^bsub>residue_ring (p^m)\<^esub>((residue (p ^ m) (y n)))" + by (metis (no_types, lifting) padic_int_simps(5) Px Py mem_Collect_eq padic_set_def ring_hom_add) + then have "?LHS =(residue (p ^ m) (x n))\\<^bsub>residue_ring (p^m)\<^esub>((residue (p ^ m) (y n)))" + using A0 by force + then show ?thesis + using A Px Py padic_set_def by (simp add: padic_int_def padic_add_def) + qed + qed + qed + qed + then show ?thesis + using "0" padic_set_memI padic_int_simps by auto + qed + then have " x \\<^bsub>padic_int p\<^esub> y \ (padic_set p)" + by(simp add: padic_int_def) + then show "carrier (padic_int p) \ carrier (padic_int p)" + by blast +qed + +text\$p$-adic addition is associative:\ + +lemma padic_add_assoc: +assumes "prime p" +shows " \x y z. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ z \ carrier (padic_int p) + \ x \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" +proof- + fix x y z + assume Ax: "x \ carrier (padic_int p)" + assume Ay: "y \ carrier (padic_int p)" + assume Az: "z \ carrier (padic_int p)" + show " (x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" + proof + fix n + show "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n = (x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)) n " + proof- + have Ex: "(x n) \ carrier (residue_ring (p^n))" + using Ax padic_set_def padic_int_simps by auto + have Ey: "(y n) \ carrier (residue_ring (p^n))" + using Ay padic_set_def padic_int_simps by auto + have Ez: "(z n) \ carrier (residue_ring (p^n))" + using Az padic_set_def padic_int_simps by auto + let ?x = "(x n)" + let ?y = "(y n)" + let ?z = "(z n)" + have P1: "(?x \\<^bsub>residue_ring (p^n)\<^esub> ?y) \\<^bsub>residue_ring (p^n)\<^esub> ?z = (x n) \\<^bsub>residue_ring (p^n)\<^esub> ((y \\<^bsub>padic_int p\<^esub> z) n)" + proof(cases "n = 0") + case True + then show ?thesis + by (simp add: residue_ring_def) + next + case False + then have "residues (p^n)" + by (simp add: assms residues_n) + then show ?thesis + using Ex Ey Ez cring.cring_simprules(7) padic_add_res residues.cring padic_int_simps + by fastforce + qed + have " ((y n)) \\<^bsub>residue_ring (p^n)\<^esub> z n =((y \\<^bsub>padic_int p\<^esub> z) n)" + using padic_add_def padic_int_simps by simp + then have P0: "(x n) \\<^bsub>residue_ring (p^n)\<^esub> ((y \\<^bsub>padic_int p\<^esub> z) n) = ((x n) \\<^bsub>residue_ring (p^n)\<^esub> ((y n) \\<^bsub>residue_ring (p^n)\<^esub> z n))" + using padic_add_def padic_int_simps by simp + have "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n = ((x \\<^bsub>padic_int p\<^esub> y) n) \\<^bsub>residue_ring (p^n)\<^esub> z n" + using padic_add_def padic_int_simps by simp + then have "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n =((x n) \\<^bsub>residue_ring (p^n)\<^esub> (y n)) \\<^bsub>residue_ring (p^n)\<^esub> z n" + using padic_add_def padic_int_simps by simp + then have "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n =((x n) \\<^bsub>residue_ring (p^n)\<^esub> ((y n) \\<^bsub>residue_ring (p^n)\<^esub> z n))" + using Ex Ey Ez P1 P0 by linarith + then have "((x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n = (x n) \\<^bsub>residue_ring (p^n)\<^esub> ((y \\<^bsub>padic_int p\<^esub> z) n)" + using P0 by linarith + then show ?thesis by (simp add: padic_int_def padic_add_def) + qed + qed +qed + +text\$p$-adic addition is commutative:\ + +lemma padic_add_comm: + assumes "prime p" + shows " \x y. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y = y \\<^bsub>padic_int p\<^esub> x" +proof- + fix x y + assume Ax: "x \ carrier (padic_int p)" assume Ay:"y \ carrier (padic_int p)" + show "x \\<^bsub>padic_int p\<^esub> y = y \\<^bsub>padic_int p\<^esub> x" + proof fix n + show "(x \\<^bsub>padic_int p\<^esub> y) n = (y \\<^bsub>padic_int p\<^esub> x) n " + proof(cases "n=0") + case True + then show ?thesis + by (metis Ax Ay assms padic_add_def padic_set_zero_res padic_int_simps(3,5)) + next + case False + have LHS0: "(x \\<^bsub>padic_int p\<^esub> y) n = (x n) \\<^bsub>residue_ring (p^n)\<^esub> (y n)" + by (simp add: padic_int_simps padic_add_res) + have RHS0: "(y \\<^bsub>padic_int p\<^esub> x) n = (y n) \\<^bsub>residue_ring (p^n)\<^esub> (x n)" + by (simp add: padic_int_simps padic_add_res) + have Ex: "(x n) \ carrier (residue_ring (p^n))" + using Ax padic_set_res_closed padic_int_simps by auto + have Ey: "(y n) \ carrier (residue_ring (p^n))" + using Ay padic_set_res_closed padic_int_simps by auto + have LHS1: "(x \\<^bsub>padic_int p\<^esub> y) n = ((x n) +(y n)) mod (p^n)" + using LHS0 residue_ring_def by simp + have RHS1: "(y \\<^bsub>padic_int p\<^esub> x) n = ((y n) +(x n)) mod (p^n)" + using RHS0 residue_ring_def by simp + then show ?thesis using LHS1 RHS1 by presburger + qed + qed +qed + +text\$padic\_zero$ is an additive identity:\ + +lemma padic_add_zero: +assumes "prime p" +shows "\x. x \ carrier (padic_int p) \ \\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x = x" +proof- + fix x + assume Ax: "x \ carrier (padic_int p)" + show " \\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x = x " + proof fix n + have A: "(padic_zero p) n = 0" + by (simp add: padic_zero_def) + have "((padic_zero p) \\<^bsub>padic_int p\<^esub> x) n = x n" + using Ax padic_int_simps(5) padic_set_res_closed residue_ring_def + by(auto simp add: padic_zero_def padic_int_simps padic_add_res residue_ring_def) + then show "(\\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x) n = x n" + by (simp add: padic_int_def) + qed +qed + +text\Closure under additive inverses:\ + +lemma padic_add_inv: +assumes "prime p" +shows "\x. x \ carrier (padic_int p) \ + \y\carrier (padic_int p). y \\<^bsub>padic_int p\<^esub> x = \\<^bsub>padic_int p\<^esub>" +proof- + fix x + assume Ax: " x \ carrier (padic_int p)" + show "\y\carrier (padic_int p). y \\<^bsub>padic_int p\<^esub> x = \\<^bsub>padic_int p\<^esub>" + proof + let ?y = "(padic_a_inv p) x" + show "?y \\<^bsub>padic_int p\<^esub> x = \\<^bsub>padic_int p\<^esub>" + proof + fix n + show "(?y \\<^bsub>padic_int p\<^esub> x) n = \\<^bsub>padic_int p\<^esub> n" + proof(cases "n=0") + case True + then show ?thesis + using Ax assms padic_add_closed padic_set_zero_res + padic_a_inv_in_padic_set padic_zero_def padic_int_simps by auto + next + case False + have C: "(x n) \ carrier (residue_ring (p^n))" + using Ax padic_set_res_closed padic_int_simps by auto + have R: "residues (p^n)" + using False by (simp add: assms residues_n) + have "(?y \\<^bsub>padic_int p\<^esub> x) n = (?y n) \\<^bsub>residue_ring (p^n)\<^esub> x n" + by (simp add: padic_int_def padic_add_res) + then have "(?y \\<^bsub>padic_int p\<^esub> x) n = 0" + using C R residue_ring_def[simp] residues.cring + by (metis (no_types, lifting) cring.cring_simprules(9) padic_a_inv_def residues.zero_cong) + then show ?thesis + by (simp add: padic_int_def padic_zero_def) + qed + qed + then show "padic_a_inv p x \ carrier (padic_int p)" + using padic_a_inv_in_padic_set padic_int_simps + Ax assms prime_gt_0_nat by auto + qed +qed + +text\The ring of padic integers forms an abelian group under addition:\ + +lemma padic_is_abelian_group: +assumes "prime p" +shows "abelian_group (padic_int p)" + proof (rule abelian_groupI) + show 0: "\x y. x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>(padic_int p)\<^esub> y \ carrier (padic_int p)" + using padic_add_closed by (simp add: assms) + show zero: "\\<^bsub>padic_int p\<^esub> \ carrier (padic_int p)" + by (metis "0" assms padic_add_inv padic_int_simps(5) padic_one_in_padic_set) + show add_assoc: " \x y z. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + z \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z + = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" + using assms padic_add_assoc by auto + show comm: " \x y. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y = y \\<^bsub>padic_int p\<^esub> x" + using assms padic_add_comm by blast + show "\x. x \ carrier (padic_int p) \ \\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x = x" + using assms padic_add_zero by blast + show "\x. x \ carrier (padic_int p) \ + \y\carrier (padic_int p). y \\<^bsub>padic_int p\<^esub> x = \\<^bsub>padic_int p\<^esub>" + using assms padic_add_inv by blast + qed + +text\One is a multiplicative identity:\ + +lemma padic_one_id: +assumes "prime p" +assumes "x \ carrier (padic_int p)" +shows "\\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x = x" +proof + fix n + show "(\\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x) n = x n " + proof(cases "n=0") + case True + then show ?thesis + by (metis padic_int_simps(1,4,5) assms(1) assms(2) padic_mult_in_padic_set padic_one_in_padic_set padic_set_zero_res) + next + case False + then have "residues (p^n)" + by (simp add: assms(1) residues_n) + then show ?thesis + using False assms(2) cring.cring_simprules(12) padic_int_simps + padic_mult_res padic_one_simp padic_set_res_closed residues.cring by fastforce + qed +qed + +text\$p$-adic multiplication is commutative:\ + +lemma padic_mult_comm: +assumes "prime p" +assumes "x \ carrier (padic_int p)" +assumes "y \ carrier (padic_int p)" +shows "x \\<^bsub>padic_int p\<^esub> y = y \\<^bsub>padic_int p\<^esub> x" +proof + fix n + have Ax: "(x n) \ carrier (residue_ring (p^n))" + using padic_set_def assms(2) padic_int_simps by auto + have Ay: "(y n) \carrier (residue_ring (p^n))" + using padic_set_def assms(3) padic_set_res_closed padic_int_simps + by blast + show "(x \\<^bsub>padic_int p\<^esub> y) n = (y \\<^bsub>padic_int p\<^esub> x) n" + proof(cases "n=0") + case True + then show ?thesis + by (metis padic_int_simps(4,5) assms(1) assms(2) assms(3) padic_set_zero_res padic_simps(3)) + next + case False + have LHS0: "(x \\<^bsub>padic_int p\<^esub> y) n = (x n) \\<^bsub>residue_ring (p^n)\<^esub> (y n)" + by (simp add: padic_int_def padic_mult_res) + have RHS0: "(y \\<^bsub>padic_int p\<^esub> x) n = (y n) \\<^bsub>residue_ring (p^n)\<^esub> (x n)" + by (simp add: padic_int_def padic_mult_res) + have Ex: "(x n) \ carrier (residue_ring (p^n))" + using Ax padic_set_res_closed by auto + have Ey: "(y n) \ carrier (residue_ring (p^n))" + using Ay padic_set_res_closed by auto + have LHS1: "(x \\<^bsub>padic_int p\<^esub> y) n = ((x n) *(y n)) mod (p^n)" + using LHS0 + by (simp add: residue_ring_def) + have RHS1: "(y \\<^bsub>padic_int p\<^esub> x) n = ((y n) *(x n)) mod (p^n)" + using RHS0 + by (simp add: residue_ring_def) + then show ?thesis using LHS1 RHS1 + by (simp add: mult.commute) + qed +qed + +lemma padic_is_comm_monoid: +assumes "prime p" +shows "Group.comm_monoid (padic_int p)" +proof(rule comm_monoidI) + show "\x y. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y \ carrier (padic_int p)" + by (simp add: padic_int_def assms padic_mult_in_padic_set) + show "\\<^bsub>padic_int p\<^esub> \ carrier (padic_int p)" + by (metis padic_int_simps(1,5) assms padic_one_in_padic_set) + show "\x y z. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + z \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z = x \\<^bsub>padic_int p\<^esub> (y \\<^bsub>padic_int p\<^esub> z)" + using assms padic_mult_assoc by auto + show "\x. x \ carrier (padic_int p) \ \\<^bsub>padic_int p\<^esub> \\<^bsub>padic_int p\<^esub> x = x" + using assms padic_one_id by blast + show "\x y. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + x \\<^bsub>padic_int p\<^esub> y = y \\<^bsub>padic_int p\<^esub> x" + using padic_mult_comm by (simp add: assms) +qed + +lemma padic_int_is_cring: + assumes "prime p" + shows "cring (padic_int p)" +proof (rule cringI) + show "abelian_group (padic_int p)" + by (simp add: assms padic_is_abelian_group) + show "Group.comm_monoid (padic_int p)" + by (simp add: assms padic_is_comm_monoid) + show "\x y z. + x \ carrier (padic_int p) \ + y \ carrier (padic_int p) \ + z \ carrier (padic_int p) \ + (x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z = + x \\<^bsub>padic_int p\<^esub> z \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z " + proof- + fix x y z + assume Ax: " x \ carrier (padic_int p)" + assume Ay: " y \ carrier (padic_int p)" + assume Az: " z \ carrier (padic_int p)" + show "(x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z + = x \\<^bsub>padic_int p\<^esub> z \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z" + proof + fix n + have Ex: " (x n) \ carrier (residue_ring (p^n))" + using Ax padic_set_def padic_int_simps by auto + have Ey: " (y n) \ carrier (residue_ring (p^n))" + using Ay padic_set_def padic_int_simps by auto + have Ez: " (z n) \ carrier (residue_ring (p^n))" + using Az padic_set_def padic_int_simps by auto + show "( (x \\<^bsub>padic_int p\<^esub> y) \\<^bsub>padic_int p\<^esub> z) n + = (x \\<^bsub>padic_int p\<^esub> z \\<^bsub>padic_int p\<^esub> y \\<^bsub>padic_int p\<^esub> z) n " + proof(cases "n=0") + case True + then show ?thesis + by (metis Ax Ay Az assms padic_add_closed padic_int_simps(4) padic_int_simps(5) padic_mult_in_padic_set padic_set_zero_res) + next + case False + then have "residues (p^n)" + by (simp add: assms residues_n) + then show ?thesis + using Ex Ey Ez cring.cring_simprules(13) padic_add_res padic_int_simps + padic_mult_res residues.cring by fastforce + qed + qed + qed +qed + +text\The $p$-adic ring has no nontrivial zero divisors. Note that this argument is short because we have proved that the valuation is multiplicative on nonzero elements, which is where the primality assumption is used.\ + +lemma padic_no_zero_divisors: +assumes "prime p" +assumes "a \ carrier (padic_int p)" +assumes "b \carrier (padic_int p)" +assumes "a \\\<^bsub>padic_int p\<^esub> " +assumes "b \\\<^bsub>padic_int p\<^esub> " +shows "a \\<^bsub>padic_int p\<^esub> b \ \\<^bsub>padic_int p\<^esub> " +proof + assume C: "a \\<^bsub>padic_int p\<^esub> b = \\<^bsub>padic_int p\<^esub>" + show False + proof- + have 0: "a = \\<^bsub>padic_int p\<^esub> \ b = \\<^bsub>padic_int p\<^esub>" + proof(cases "a = \\<^bsub>padic_int p\<^esub>") + case True + then show ?thesis by auto + next + case False + have "\ b \\\<^bsub>padic_int p\<^esub>" + proof + assume "b \ \\<^bsub>padic_int p\<^esub>" + have "padic_val p (a \\<^bsub>padic_int p\<^esub> b) = (padic_val p a) + (padic_val p b)" + using False assms(1) assms(2) assms(3) assms(5) val_prod padic_int_simps by auto + then have "padic_val p (a \\<^bsub>padic_int p\<^esub> b) \ -1" + using False \b \ \\<^bsub>padic_int p\<^esub>\ padic_val_def padic_int_simps by auto + then show False + using C padic_val_def padic_int_simps by auto + qed + then show ?thesis + by blast + qed + show ?thesis + using "0" assms(4) assms(5) by blast + qed +qed + +lemma padic_int_is_domain: + assumes "prime p" + shows "domain (padic_int p)" +proof(rule domainI) + show "cring (padic_int p)" + using padic_int_is_cring assms(1) by auto + show "\\<^bsub>padic_int p\<^esub> \ \\<^bsub>padic_int p\<^esub>" + proof + assume "\\<^bsub>padic_int p\<^esub> = \\<^bsub>padic_int p\<^esub> " + then have "(\\<^bsub>padic_int p\<^esub>) 1 = \\<^bsub>padic_int p\<^esub> 1" by auto + then show False + using padic_int_simps(1,2) + unfolding padic_one_def padic_zero_def by auto + qed + show "\a b. a \\<^bsub>padic_int p\<^esub> b = \\<^bsub>padic_int p\<^esub> \ + a \ carrier (padic_int p) \ + b \ carrier (padic_int p) \ + a = \\<^bsub>padic_int p\<^esub> \ b = \\<^bsub>padic_int p\<^esub>" + using assms padic_no_zero_divisors + by (meson prime_nat_int_transfer) +qed + +section\The Ultrametric Inequality:\ + +lemma padic_val_ultrametric: + assumes "prime p" + assumes "a \ carrier (padic_int p) " + assumes "b \ carrier (padic_int p) " + assumes "a \ \\<^bsub>(padic_int p)\<^esub>" + assumes "b \ \\<^bsub>(padic_int p)\<^esub>" + assumes "a \\<^bsub>(padic_int p)\<^esub> b \ \\<^bsub>(padic_int p)\<^esub>" + shows "padic_val p (a \\<^bsub>(padic_int p)\<^esub> b) \ min (padic_val p a) (padic_val p b)" +proof- + let ?va = " nat (padic_val p a)" + let ?vb = "nat (padic_val p b)" + let ?vab = "nat (padic_val p (a \\<^bsub>(padic_int p)\<^esub> b))" + have P:" \ ?vab < min ?va ?vb" + proof + assume P0: "?vab < min ?va ?vb" + then have "Suc ?vab \ min ?va ?vb" + using Suc_leI by blast + have "(a \\<^bsub>(padic_int p)\<^esub> b) \ carrier (padic_int p) " + using assms(1) assms(2) assms(3) padic_add_closed by simp + then have C: "(a \\<^bsub>(padic_int p)\<^esub> b) (?vab + 1) \ \\<^bsub>residue_ring (p^(?vab + 1))\<^esub>" + using val_of_nonzero(1) assms(6) + by (simp add: padic_int_def val_of_nonzero(1) assms(1)) + have S: "(a \\<^bsub>(padic_int p)\<^esub> b) (?vab + 1) = (a (?vab + 1)) \\<^bsub>residue_ring (p^((?vab + 1)))\<^esub> (b ((?vab + 1)))" + by (simp add: padic_int_def padic_add_def) + have "int (?vab + 1) \ padic_val p a" + using P0 using Suc_le_eq by auto + then have A: "(a (?vab + 1)) = \\<^bsub>residue_ring (p^((?vab + 1)))\<^esub> " + using assms(1) assms(2) zero_below_val padic_int_simps residue_ring_def + by auto + have "int (?vab + 1) \ padic_val p b" + using P0 using Suc_le_eq by auto + then have B: "(b (?vab + 1)) = \\<^bsub>residue_ring (p^((?vab + 1)))\<^esub> " + using assms(1) assms(3) zero_below_val + by (metis A \int (nat (padic_val p (a \\<^bsub>padic_int p\<^esub> b)) + 1) \ padic_val p a\ + assms(2) padic_int_simps(3,5)) + have "p^(?vab + 1) > 1" + using assms(1) by (metis add.commute plus_1_eq_Suc power_gt1 prime_gt_1_int) + then have "residues (p^(?vab + 1))" + using less_imp_of_nat_less residues.intro by fastforce + then have "(a \\<^bsub>(padic_int p)\<^esub> b) (?vab + 1) = \\<^bsub>residue_ring (p^((?vab + 1)))\<^esub> " + using A B by (metis (no_types, lifting) S cring.cring_simprules(2) + cring.cring_simprules(8) residues.cring) + then show False using C by auto + qed + have A0: "(padic_val p a) \ 0" + using assms(4) padic_val_def by(auto simp: padic_int_def) + have A1: "(padic_val p b) \ 0" + using assms(5) padic_val_def by(auto simp: padic_int_def) + have A2: "padic_val p (a \\<^bsub>(padic_int p)\<^esub> b) \ 0" + using assms(6) padic_val_def by(auto simp: padic_int_def) + show ?thesis using P A0 A1 A2 + by linarith +qed + +lemma padic_a_inv: + assumes "prime p" + assumes "a \ carrier (padic_int p)" + shows "\\<^bsub>padic_int p\<^esub> a = (\ n. \\<^bsub>residue_ring (p^n)\<^esub> (a n))" +proof + fix n + show "(\\<^bsub>padic_int p\<^esub> a) n = \\<^bsub>residue_ring (p^n)\<^esub> a n" + proof(cases "n=0") + case True + then show ?thesis + by (metis (no_types, lifting) abelian_group.a_inv_closed assms(1) assms(2) padic_int_simps(5) + padic_is_abelian_group padic_set_zero_res power_0 residue_1_prop residue_ring_def ring.simps(1)) + next + case False + then have R: "residues (p^n)" + by (simp add: assms(1) residues_n) + have "(\\<^bsub>padic_int p\<^esub> a) \\<^bsub>padic_int p\<^esub> a = \\<^bsub>padic_int p\<^esub>" + by (simp add: abelian_group.l_neg assms(1) assms(2) padic_is_abelian_group) + then have P: "(\\<^bsub>padic_int p\<^esub> a) n \\<^bsub>residue_ring (p^n)\<^esub> a n = 0" + by (metis padic_add_res padic_int_simps(2) padic_int_simps(3) padic_zero_def) + have Q: "(a n) \ carrier (residue_ring (p^n))" + using assms(2) padic_set_res_closed by(auto simp: padic_int_def) + show ?thesis using R Q residues.cring + by (metis P abelian_group.a_inv_closed abelian_group.minus_equality assms(1) assms(2) + padic_int_simps(5) padic_is_abelian_group padic_set_res_closed residues.abelian_group + residues.res_zero_eq) + qed +qed + +lemma padic_val_a_inv: + assumes "prime p" + assumes "a \ carrier (padic_int p)" + shows "padic_val p a = padic_val p (\\<^bsub>padic_int p\<^esub> a)" +proof(cases "a = \\<^bsub>padic_int p\<^esub>") + case True + then show ?thesis + by (metis abelian_group.a_inv_closed abelian_group.r_neg abelian_groupE(5) assms(1) assms(2) padic_is_abelian_group) +next + case False + have 0: "\ n. (a n) = \\<^bsub>residue_ring (p^n)\<^esub> \ (\\<^bsub>padic_int p\<^esub> a) n = \\<^bsub>residue_ring (p^n)\<^esub>" + using padic_a_inv + by (metis (no_types, lifting) assms(1) assms(2) cring.cring_simprules(22) power_0 residue_1_prop residues.cring residues_n) + have 1: "\ n. (a n) \ \\<^bsub>residue_ring (p^n)\<^esub> \ (\\<^bsub>padic_int p\<^esub> a) n \ \\<^bsub>residue_ring (p^n)\<^esub>" + using padic_a_inv + by (metis (no_types, lifting) abelian_group.a_inv_closed abelian_group.minus_minus assms(1) + assms(2) cring.cring_simprules(22) padic_int_simps(5) padic_is_abelian_group padic_set_zero_res + residues.cring residues_n) + have A:"padic_val p (\\<^bsub>padic_int p\<^esub> a) \ (padic_val p a)" + proof- + have "\ padic_val p (\\<^bsub>padic_int p\<^esub> a) < (padic_val p a)" + proof + assume "padic_val p (\\<^bsub>padic_int p\<^esub> a) < padic_val p a" + let ?n = "padic_val p (\\<^bsub>padic_int p\<^esub> a)" + let ?m = " padic_val p a" + have "(\\<^bsub>padic_int p\<^esub> a) \ (padic_zero p)" + by (metis False abelian_group.l_neg assms(1) assms(2) padic_add_zero padic_int_simps(2) padic_is_abelian_group) + then have P0: "?n \0" + by (simp add: padic_val_def) + have P1: "?m \0" using False + using \0 \ padic_val p (\\<^bsub>padic_int p\<^esub> a)\ + \padic_val p (\\<^bsub>padic_int p\<^esub> a) < padic_val p a\ by linarith + have "(Suc (nat ?n)) < Suc (nat (padic_val p a))" + using P0 P1 \padic_val p (\\<^bsub>padic_int p\<^esub> a) < padic_val p a\ by linarith + then have "int (Suc (nat ?n)) \ (padic_val p a)" + using of_nat_less_iff by linarith + then have "a (Suc (nat ?n)) = \\<^bsub>residue_ring (p ^ ((Suc (nat ?n))))\<^esub>" + using assms(1) assms(2) zero_below_val residue_ring_def by(auto simp: padic_int_def) + then have "(\\<^bsub>padic_int p\<^esub> a) (Suc (nat ?n)) = \\<^bsub>residue_ring (p ^ ((Suc (nat ?n))))\<^esub>" + using 0 by simp + then show False using below_val_zero assms + by (metis Suc_eq_plus1 \\\<^bsub>padic_int p\<^esub> a \ padic_zero p\ abelian_group.a_inv_closed + padic_int_simps(5) padic_is_abelian_group val_of_nonzero(1)) + qed + then show ?thesis + by linarith + qed + have B: "padic_val p (\\<^bsub>padic_int p\<^esub> a) \ (padic_val p a)" + proof- + let ?n = "nat (padic_val p a)" + have "a (Suc ?n) \ \\<^bsub>residue_ring (p^(Suc ?n))\<^esub> " + using False assms(2) val_of_nonzero(1) + by (metis padic_int_simps(2,5) Suc_eq_plus1 assms(1)) + then have "(\\<^bsub>padic_int p\<^esub> a) (Suc ?n) \ \\<^bsub>residue_ring (p^(Suc ?n))\<^esub> " + using 1 by blast + then have "padic_val p (\\<^bsub>padic_int p\<^esub> a) \ int ?n" using assms(1) assms(2) below_val_zero + by (metis padic_int_simps(5) abelian_group.a_inv_closed padic_is_abelian_group) + then show ?thesis + using False padic_val_def padic_int_simps by auto + qed + then show ?thesis using A B by auto +qed + +end diff --git a/thys/Padic_Ints/Padic_Int_Polynomials.thy b/thys/Padic_Ints/Padic_Int_Polynomials.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Padic_Int_Polynomials.thy @@ -0,0 +1,188 @@ +theory Padic_Int_Polynomials +imports Padic_Int_Topology Cring_Poly +begin + +context padic_integers +begin + +text\ + This theory states and proves basic lemmas connecting the topology on $\mathbb{Z}_p$ with the + functions induced by polynomial evaluation over $\mathbb{Z}_p$. This will imply that polynomial + evaluation applied to a Cauchy Sequence will always produce a cauchy sequence, which is a key + fact in the proof of Hensel's Lemma. +\ + +type_synonym padic_int_poly = "nat \ padic_int" + +lemma monom_term_car: + assumes "c \ carrier Zp" + assumes "x \ carrier Zp" + shows "c \ x[^](n::nat) \ carrier Zp" + using assms R.nat_pow_closed + by (simp add: monoid.nat_pow_closed cring.cring_simprules(5) cring_def ring_def) + +text\Univariate polynomial ring over Zp\ + +abbreviation(input) Zp_x where +"Zp_x \ UP Zp" + +lemma Zp_x_is_UP_cring: +"UP_cring Zp" + using UP_cring.intro domain_axioms domain_def by auto + +lemma Zp_x_is_UP_domain: +"UP_domain Zp" + by (simp add: UP_domain_def domain_axioms) + +lemma Zp_x_domain: +"domain Zp_x " + by (simp add: UP_domain.UP_domain Zp_x_is_UP_domain) + +lemma pow_closed: + assumes "a \ carrier Zp" + shows "a[^](n::nat) \ carrier Zp" + by (meson domain_axioms domain_def cring_def assms monoid.nat_pow_closed ring_def) + +lemma(in ring) pow_zero: + assumes "(n::nat)>0" + shows "\[^] n = \" + by (simp add: assms nat_pow_zero) + +lemma sum_closed: + assumes "f \ carrier Zp" + assumes "g \ carrier Zp" + shows "f \ g \ carrier Zp" + by (simp add: assms(1) assms(2) cring.cring_simprules(1)) + +lemma mult_zero: + assumes "f \ carrier Zp" + shows "f \ \ = \" + "\ \ f = \" + apply (simp add: assms cring.cring_simprules(27)) + by (simp add: assms cring.cring_simprules(26)) + +lemma monom_poly_is_Zp_continuous: + assumes "c \ carrier Zp" + assumes "f = monom Zp_x c n" + shows "is_Zp_continuous (to_fun f)" + using monomial_is_Zp_continuous assms monom_to_monomial by auto + +lemma degree_0_is_Zp_continuous: + assumes "f \ carrier Zp_x" + assumes "degree f = 0" + shows "is_Zp_continuous (to_fun f)" + using const_to_constant[of "lcf f"] assms constant_is_Zp_continuous ltrm_deg_0 + by (simp add: cfs_closed) + +lemma UP_sum_is_Zp_continuous: + assumes "a \ carrier Zp_x" + assumes "b \ carrier Zp_x" + assumes "is_Zp_continuous (to_fun a)" + assumes "is_Zp_continuous (to_fun b)" + shows "is_Zp_continuous (to_fun (a \\<^bsub>Zp_x\<^esub> b))" + using sum_of_cont_is_cont assms + by (simp add: to_fun_Fun_add) + +lemma polynomial_is_Zp_continuous: + assumes "f \ carrier Zp_x" + shows "is_Zp_continuous (to_fun f)" + apply(rule poly_induct3) + apply (simp add: assms) + using UP_sum_is_Zp_continuous apply blast + using monom_poly_is_Zp_continuous by blast +end + +text\Notation for polynomial function application\ + +context padic_integers +begin +notation to_fun (infixl\\\ 70) +text\Evaluating polynomials in the residue rings\ + +lemma res_to_fun_monic_monom: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a k = b k" + shows "(monom Zp_x \ n \ a) k = (monom Zp_x \ n \ b) k" +proof(induction n) + case 0 + then show ?case + using UP_cring.to_fun_X_pow Zp_x_is_UP_domain assms(1) assms(2) nat_pow_0 to_fun_one monom_one + by presburger +next + case (Suc n) + fix n::nat + assume IH: "to_fun (monom Zp_x \ n) a k = to_fun (monom Zp_x \ n) b k" + show "to_fun (monom Zp_x \ (Suc n)) a k = to_fun (monom Zp_x \ (Suc n)) b k" + proof- + have LHS0: "(monom Zp_x \ (Suc n) \ a) k = ((monom Zp_x \ n \ a) \ (X \ a)) k" + by (simp add: UP_cring.to_fun_monic_monom Zp_x_is_UP_cring assms(1)) + then show ?thesis + using assms IH Zp_x_is_UP_domain + Zp_continuous_is_Zp_closed + by (metis (mono_tags, lifting) R.one_closed X_poly_def monom_closed monom_one_Suc + residue_of_prod to_fun_X to_fun_mult) + qed +qed + +lemma res_to_fun_monom: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "c \ carrier Zp" + assumes "a k = b k" + shows "(monom Zp_x c n \ a) k = (monom Zp_x c n \ b) k" + using res_to_fun_monic_monom assms + by (smt to_fun_monic_monom to_fun_monom residue_of_prod) + +lemma to_fun_res_ltrm: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "f \ carrier Zp_x" + assumes "a k = b k" + shows "((ltrm f)\a) k = ((ltrm f)\b) k" + by (simp add: lcf_closed assms(1) assms(2) assms(3) assms(4) res_to_fun_monom) + +text\Polynomial application commutes with taking residues\ +lemma to_fun_res: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a k = b k" + shows "(f\a) k = (f\b) k" + apply(rule poly_induct3[of f]) + apply (simp add: assms(1)) + using assms(2) assms(3) to_fun_plus residue_of_sum apply presburger + using assms(2) assms(3) assms(4) res_to_fun_monom by blast + + +text\If a and b have equal kth residues, then so do f'(a) and f'(b)\ + +lemma deriv_res: + assumes "f \ carrier Zp_x" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a k = b k" + shows "(deriv f a) k = (deriv f b) k" + using assms to_fun_res[of "pderiv f" a b k] + by (simp add: pderiv_closed pderiv_eval_deriv) + +text\Propositions about evaluation:\ + + +lemma to_fun_monom_plus: + assumes "a \ carrier Zp" + assumes "g \ carrier Zp_x" + assumes "c \ carrier Zp" + shows "(monom Zp_x a n \\<^bsub>Zp_x\<^esub> g)\c = a \ c[^]n \ (g \ c)" + by (simp add: assms(1) assms(2) assms(3) to_fun_monom to_fun_plus) + +lemma to_fun_monom_minus: + assumes "a \ carrier Zp" + assumes "g \ carrier Zp_x" + assumes "c \ carrier Zp" + shows "(monom Zp_x a n \\<^bsub>Zp_x\<^esub> g)\c = a \ c[^]n \ (g \ c)" + by (simp add: UP_cring.to_fun_diff Zp_x_is_UP_cring assms(1) assms(2) assms(3) to_fun_monom) + +end + +end diff --git a/thys/Padic_Ints/Padic_Int_Topology.thy b/thys/Padic_Ints/Padic_Int_Topology.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Padic_Int_Topology.thy @@ -0,0 +1,1093 @@ +theory Padic_Int_Topology +imports Padic_Integers Function_Ring +begin + +type_synonym padic_int_seq = "nat \ padic_int" + +type_synonym padic_int_fun = "padic_int \ padic_int" + +sublocale padic_integers < FunZp?: U_function_ring "Zp" + unfolding U_function_ring_def + by (simp add: R.ring_axioms) + +context padic_integers +begin + +(**************************************************************************************************) +(**************************************************************************************************) +(***********************************)section\Sequences over Zp\(***********************************) +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + The $p$-adic valuation can be thought of as equivalent to the $p$-adic absolute value, but with + the notion of size inverted so that small numbers have large valuation, and zero has maximally + large valuation. The $p$-adic distance between two points is just the valuation of the difference + of those points, and is thus equivalent to the metric induced by the $p$-adic absolute value. + For background on valuations and absolute values for $p$-adic rings see \cite{engler2005valued}. + In what follows, we develop the topology of the $p$-adic from a valuative perspective rather than + a metric perspective. Though equivalent to the metric approach in the $p$-adic case, this + approach is more general in that there exist valued rings whose valuations take values in + non-Archimedean ordered ablelian groups which do not embed into the real numbers. +\ + + (**********************************************************************) + (**********************************************************************) + subsection\The Valuative Distance Function on $\mathbb{Z}_p$\ + (**********************************************************************) + (**********************************************************************) + +text\ + The following lemmas establish that the $p$-adic distance function satifies the standard + properties of an ultrametric. It is symmetric, obeys the ultrametric inequality, and only + identical elements are infinitely close. +\ + +definition val_Zp_dist :: "padic_int \ padic_int \ eint" where +"val_Zp_dist a b \ val_Zp (a \ b)" + +lemma val_Zp_dist_sym: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + shows "val_Zp_dist a b = val_Zp_dist b a" +proof- + have 1: "a \ b = \ (b \ a)" using assms(1) assms(2) + using minus_a_inv by blast + then show ?thesis + using R.minus_closed Zp_def assms(1) assms(2) padic_integers.val_Zp_of_minus + padic_integers_axioms val_Zp_dist_def by auto +qed + +lemma val_Zp_dist_ultrametric: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "c \ carrier Zp" + shows "val_Zp_dist b c \ min (val_Zp_dist a c) (val_Zp_dist a b)" +proof- + let ?X = "b \ a" + let ?Y = "a \ c" + let ?Z = "b \ c" + have 0: "?Z = ?X \ ?Y" + using R.add.m_comm assms(1) assms(2) assms(3) R.plus_diff_simp by auto + have 4: "val_Zp ?Z \ min (val_Zp ?X) (val_Zp ?Y)" + using "0" assms(1) assms(2) assms(3) val_Zp_ultrametric by auto + then show ?thesis + using assms val_Zp_dist_sym + unfolding val_Zp_dist_def + by (simp add: min.commute) +qed + +lemma val_Zp_dist_infty: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "val_Zp_dist a b = \" + shows "a = b" + using assms unfolding val_Zp_dist_def + by (metis R.r_right_minus_eq not_eint_eq val_ord_Zp) + +lemma val_Zp_dist_infty': + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "a = b" + shows "val_Zp_dist a b = \" + using assms unfolding val_Zp_dist_def + by (simp add: val_Zp_def) + +text\ + The following property will be useful in the proof of Hensel's Lemma: two $p$-adic integers are + close together if and only if their residues are equal at high orders. +\ + +lemma val_Zp_dist_res_eq: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "val_Zp_dist a b > k" + shows "(a k) = (b k)" + using assms(1) assms(2) assms(3) val_Zp_dist_def + by (simp add: Zp_residue_eq) + +lemma val_Zp_dist_res_eq2: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "(a k) = (b k)" + shows "val_Zp_dist a b \ k" + using assms(1) assms(2) assms(3) Zp_residue_eq2 + unfolding val_Zp_dist_def + by (simp add: val_Zp_def) + +lemma val_Zp_dist_triangle_eqs: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "c \ carrier Zp" + assumes "val_Zp_dist a b > n" + assumes "val_Zp_dist a c > n" + assumes "(k::nat) < n" + shows "a k = b k" + "a k = c k" + "b k = c k" + unfolding val_Zp_dist_def +proof- + show 0: "a k = b k" + using assms(1) assms(2) assms(4) assms(6) val_Zp_dist_res_eq + by (metis less_imp_le_nat p_residue_padic_int) + show 1: "a k = c k" + using assms(1) assms(3) assms(5) assms(6) val_Zp_dist_res_eq + by (meson eint_ord_simps(1) le_less_trans less_imp_triv not_less of_nat_le_iff) + show "b k = c k" + using 0 1 by auto +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Cauchy Sequences\ + (**********************************************************************) + (**********************************************************************) + +text\ + The definition of Cauchy sequence here is equivalent to standard the metric notion, and is + identical to the one found on page 50 of \cite{engler2005valued}. +\ + +lemma closed_seqs_diff_closed: + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + shows "s m \ a \ carrier Zp" + using assms + by (simp add: closed_seqs_memE) + +definition is_Zp_cauchy :: "padic_int_seq \ bool" where +"is_Zp_cauchy s = ((s \ closed_seqs Zp) \ (\ (n::int). \ (N::nat). \ m k::nat. + + (m>N \ k>N \ (val_Zp_dist (s m) (s k)) > eint n)))" + +text\Relation for a sequence which converges to a point:\ + +definition Zp_converges_to :: "padic_int_seq \ padic_int \ bool" where + "Zp_converges_to s a = ((a \ carrier Zp \ s \ closed_seqs Zp) + \ (\(n::int). (\(k:: nat). (\( m::nat). + (m > k \ (val_Zp ((s m) \ a)) > eint n) ))))" + +lemma is_Zp_cauchy_imp_closed: + assumes "is_Zp_cauchy s" + shows "s \ closed_seqs Zp" + using assms unfolding is_Zp_cauchy_def by blast + +text\ + Analogous to the lemmas about residues and $p$-adic distances, we can characterize Cauchy + sequences without reference to a distance function: a sequence is Cauchy if and only if for + every natural number $k$, the $k^{th}$ residues of the elements in the sequence are eventually + all equal. +\ + +lemma is_Zp_cauchy_imp_res_eventually_const_0: + assumes "is_Zp_cauchy s" + fixes n::nat + obtains N where "\ n0 n1. n0 > N \ n1 > N \ (s n0) n = (s n1) n" +proof- + have "\ (N::nat). \ m k::nat. (m>N \ k>N \ (val_Zp_dist (s m) (s k)) > (int n))" + using assms is_Zp_cauchy_def by blast + then obtain N where P0: " \ m k::nat. (m>N \ k>N \ (val_Zp_dist (s m) (s k)) > (int n))" + by blast + have P1: "\ n0 n1. n0 > N \ n1 > N \ (s n0) n = (s n1) n" + proof- + fix n0 n1 + assume A: "n0 > N \ n1 > N" + have "(n0>N \ n1>N \ (val_Zp_dist (s n0) (s n1)) > (int n))" + using P0 by blast + then have C0: "(val_Zp_dist (s n0) (s n1)) > (int n)" + using A by blast + show "(s n0) n = (s n1) n" + proof- + have A0: "(val_Zp_dist (s n0) (s n1)) > (int n)" + using C0 by blast + have A1: "s n0 \ carrier Zp" + using is_Zp_cauchy_imp_closed[of s] assms + by (simp add: closed_seqs_memE) + have A2: "s n1 \ carrier Zp" + using is_Zp_cauchy_def assms closed_seqs_memE[of _ Zp] + by blast + show ?thesis + using A0 val_Zp_dist_res_eq A1 A2 by metis + qed + qed + then show ?thesis + using that by blast +qed + +lemma is_Zp_cauchyI: + assumes "s \ closed_seqs Zp" + assumes "\ n. (\N. (\ n0 n1. n0 > N \ n1 > N \ (s n0) n = (s n1) n))" + shows "is_Zp_cauchy s" +proof- + have "(\ (n::int). \ (N::nat). \ m k::nat. (m>N \ k>N \ (val_Zp_dist (s m) (s k)) > n))" + proof + fix n + show "\ (N::nat). \ m k::nat. (m>N \ k>N \ (val_Zp_dist (s m) (s k)) > eint n)" + proof- + have "(\N. (\ n0 n1. n0 > N \ n1 > N \ (s n0) (Suc (nat n)) = (s n1) (Suc (nat n))))" + using assms(2) by blast + then obtain N where N_def: "(\ n0 n1. n0 > N \ n1 > N \ (s n0) (Suc (nat n)) = (s n1) (Suc (nat n)))" + by blast + have 0: "n \ eint (int (nat n))" + by simp + have "\m k. N < m \ N < k \ (nat n) < val_Zp_dist (s m) (s k)" + proof + fix m + show "\k. N < m \ N < k \ int (nat n) < val_Zp_dist (s m) (s k)" + proof + fix k + show "N < m \ N < k \ int(nat n) < val_Zp_dist (s m) (s k)" + proof + assume A: "N < m \ N < k" + then have E: "(s m) (Suc(nat n)) = (s k) (Suc(nat n))" + using N_def by blast + then show " int (nat n) < val_Zp_dist (s m) (s k)" + proof- + have 0: "(s m) \ carrier Zp" + by (simp add: assms(1) closed_seqs_memE) + have 1: "(s k) \ carrier Zp" + using Zp_def assms(1) closed_seqs_memE[of _ Zp] padic_integers_axioms by blast + have "int (Suc (nat n)) \ val_Zp_dist (s m) (s k)" + using E 0 1 val_Zp_dist_res_eq[of "(s m)" "(s k)" "Suc (nat n)"] val_Zp_dist_res_eq2 + by blast + then have "int (nat n) < val_Zp_dist (s m) (s k)" + by (metis Suc_ile_eq add.commute of_nat_Suc) + then show ?thesis + by blast + qed + qed + qed + qed + hence "\m k. N < m \ N < k \ n < val_Zp_dist (s m) (s k)" + using 0 + by (simp add: order_le_less_subst2) + thus ?thesis by blast + qed + qed + then show ?thesis + using is_Zp_cauchy_def assms by blast +qed + +lemma is_Zp_cauchy_imp_res_eventually_const: + assumes "is_Zp_cauchy s" + fixes n::nat + obtains N r where "r \ carrier (Zp_res_ring n)" and "\ m. m > N \ (s m) n = r" +proof- + obtain N where A0: "\ n0 n1. n0 > N \ n1 > N \ (s n0) n = (s n1) n" + using assms is_Zp_cauchy_imp_res_eventually_const_0 padic_integers_axioms by blast + obtain r where A1: "s (Suc N) n = r" + by simp + have 0: "r \ carrier (Zp_res_ring n)" + using Zp_def \s (Suc N) n = r\ assms closed_seqs_memE[of _ Zp] + is_Zp_cauchy_def padic_integers_axioms residues_closed + by blast + have 1: "\ m. m > N \ (s m) n = r" + proof- + fix m + assume " m > N" + then show "(s m) n = r" + using A0 A1 by blast + qed + then show ?thesis + using 0 1 that by blast +qed + +text\ + This function identifies the eventual residues of the elements of a cauchy sequence. + Since a $p$-adic integer is defined to be the map which identifies its residues, this map + will turn out to be precisely the limit of a cauchy sequence. +\ +definition res_lim :: "padic_int_seq \ padic_int" where +"res_lim s = (\ k. (THE r. (\N. (\ m. m > N \ (s m) k = r))))" + +lemma res_lim_Zp_cauchy_0: + assumes "is_Zp_cauchy s" + assumes "f = (res_lim s) k" + shows "(\N. (\ m. (m > N \ (s m) k = f)))" + "f \ carrier (Zp_res_ring k)" +proof- + obtain N r where P0: "r \ carrier (Zp_res_ring k)" and P1: "\ m. m > N \ (s m) k = r" + by (meson assms(1) is_Zp_cauchy_imp_res_eventually_const) + obtain P where P_def: "P = (\ x. (\N. (\ m. m > N \ (s m) k = x)))" + by simp + have 0: "P r" + using P1 P_def by auto + have 1: "(\x. P x \ x = r)" + proof- + fix x + assume A_x: "P x" + obtain N0 where "(\ m. m > N0 \ (s m) k = x)" + using A_x P_def by blast + let ?M = "max N0 N" + have C0: "s (Suc ?M) k = x" + by (simp add: \\m>N0. s m k = x\) + have C1: "s (Suc ?M) k = r" + by (simp add: P1) + show "x = r" + using C0 C1 by auto + qed + have R: "r = (THE x. P x)" + using the_equality 0 1 by metis + have "(res_lim s) k = (THE r. \N. \m>N. s m k = r)" + using res_lim_def by simp + then have "f = (THE r. \N. \m>N. s m k = r)" + using assms by auto + then have "f = (THE r. P r)" + using P_def by auto + then have "r = f" + using R by auto + then show "(\N. (\ m. (m > N \ (s m) k = f)))" using 0 P_def + by blast + show "f \ carrier (Zp_res_ring k)" + using P0 \r = f\ by auto +qed + +lemma res_lim_Zp_cauchy: + assumes "is_Zp_cauchy s" + obtains N where "\ m. (m > N \ (s m) k = (res_lim s) k)" + using res_lim_Zp_cauchy_0 assms by presburger + +lemma res_lim_in_Zp: + assumes "is_Zp_cauchy s" + shows "res_lim s \ carrier Zp" +proof- + have "res_lim s \ padic_set p" + proof(rule padic_set_memI) + show "\m. res_lim s m \ carrier (residue_ring (p^ m))" + using res_lim_Zp_cauchy_0 by (simp add: assms) + show "\m n. m < n \ residue (p^ m) (res_lim s n) = res_lim s m" + proof- + fix m n + obtain N where N0: "\ l. (l > N \ (s l) m = (res_lim s) m)" + using assms res_lim_Zp_cauchy by blast + obtain M where M0: "\ l. (l > M \ (s l) n = (res_lim s) n)" + using assms prod.simps(2) res_lim_Zp_cauchy by auto + obtain K where K_def: "K = max M N" + by simp + have Km: "\ l. (l > K \ (s l) m = (res_lim s) m)" + using K_def N0 by simp + have Kn: "\ l. (l > K \ (s l) n = (res_lim s) n)" + using K_def M0 by simp + assume "m < n" + show "residue (p^ m) (res_lim s n) = res_lim s m" + proof- + obtain l where l_def: "l = Suc K" + by simp + have ln: "(res_lim s n) = (s l) n" + by (simp add: Kn l_def) + have lm: "(res_lim s m) = (s l) m" + by (simp add: Km l_def) + have s_car: "s l \ carrier Zp" + using assms is_Zp_cauchy_def closed_seqs_memE[of _ Zp] by blast + then show ?thesis + using s_car lm ln \m < n\ p_residue_def p_residue_padic_int by auto + qed + qed + qed + then show ?thesis + by (simp add: Zp_def padic_int_simps(5)) +qed + + (**********************************************************************) + (**********************************************************************) + subsection\Completeness of $\mathbb{Z}_p$\ + (**********************************************************************) + (**********************************************************************) + +text\ + We can use the developments above to show that a sequence of $p$-adic integers is convergent + if and only if it is cauchy, and that limits of convergent sequences are always unique. +\ + +lemma is_Zp_cauchy_imp_has_limit: + assumes "is_Zp_cauchy s" + assumes "a = res_lim s" + shows "Zp_converges_to s a" +proof- + have 0: "(a \ carrier Zp \ s \ closed_seqs Zp)" + using assms(1) assms(2) is_Zp_cauchy_def res_lim_in_Zp by blast + have 1: "(\(n::int). (\(k:: nat). (\( m::nat). + (m > k \ (val_Zp ((s m) \ a)) \ n))))" + proof + fix n + show "\k. \m>k. eint n \ val_Zp (s m \ a)" + proof- + obtain K where K_def: "\m. (m > K \ (s m) (nat n) = (res_lim s) (nat n))" + using assms(1) res_lim_Zp_cauchy + by blast + have "\m>K. n \ val_Zp_dist (s m) a" + proof + fix m + show "K < m \ n \ val_Zp_dist (s m) a" + proof + assume A: "K < m" + show " n \ val_Zp_dist (s m) a" + proof- + have X: "(s m) \ carrier Zp" + using "0" closed_seqs_memE[of _ Zp] + by blast + have "(s m) (nat n) = (res_lim s) (nat n)" + using A K_def by blast + then have "(s m) (nat n) = a (nat n)" + using assms(2) by blast + then have "int (nat n) \ val_Zp_dist (s m) a" + using X val_Zp_dist_res_eq2 "0" by blast + then show ?thesis + by (metis eint_ord_simps(1) int_ops(1) less_not_sym nat_eq_iff2 not_le order_trans zero_eint_def) + qed + qed + qed + then show ?thesis + using val_Zp_dist_def by auto + qed + qed + then show ?thesis using + "0" Zp_converges_to_def + by (metis Suc_ile_eq val_Zp_dist_def) +qed + +lemma convergent_imp_Zp_cauchy: + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + assumes "Zp_converges_to s a" + shows "is_Zp_cauchy s" + apply(rule is_Zp_cauchyI) + using assms apply simp +proof- + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n " + proof- + obtain k where k_def:"\m>k. n < val_Zp_dist (s m) a" + using assms val_Zp_dist_def + unfolding Zp_converges_to_def + by presburger + have A0: "\n0 n1. k < n0 \ k < n1 \ s n0 n = s n1 n " + proof- have "\n0 n1. k < n0 \ k < n1 \ s n0 n = s n1 n" + proof + fix n0 n1 + assume A: " k < n0 \ k < n1" + show " s n0 n = s n1 n " + proof- + have 0: "n < val_Zp_dist (s n0) a" + using k_def using A + by blast + have 1: "n < val_Zp_dist (s n1) a" + using k_def using A + by blast + hence 2: "(s n0) n = a n" + using 0 assms val_Zp_dist_res_eq[of "s n0" a n] closed_seqs_memE + by blast + hence 3: "(s n1) n = a n" + using 1 assms val_Zp_dist_res_eq[of "s n1" a n] closed_seqs_memE + by blast + show ?thesis + using 2 3 + by auto + qed + qed + thus ?thesis by blast + qed + show ?thesis + using A0 + by blast + qed +qed + +lemma unique_limit: + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "Zp_converges_to s a" + assumes "Zp_converges_to s b" + shows "a = b" +proof- + have "\k. a k = b k" + proof- + fix k::nat + obtain k0 where k0_def:"\m>k0. k < val_Zp_dist (s m) a" + using assms unfolding val_Zp_dist_def Zp_converges_to_def + by blast + obtain k1 where k1_def:"\m>k1. k < val_Zp_dist (s m) b" + using assms unfolding val_Zp_dist_def Zp_converges_to_def + by blast + have k0_prop: "\m. m> k0 \ (s m) k = a k" proof- fix m assume A: "m > k0" + then show "(s m) k = a k" + using k0_def assms closed_seqs_memE[of s Zp] val_Zp_dist_res_eq[of _ a k] of_nat_Suc + by blast + qed + have k1_prop: "\m. m> k1 \ (s m) k = b k" + using k1_def assms closed_seqs_memE[of s Zp] + by (simp add: val_Zp_dist_res_eq) + have "\ m. m > (max k0 k1) \ a k = b k" + using k0_prop k1_prop + by force + then show "a k = b k" + by blast + qed + then show ?thesis + by blast +qed + +lemma unique_limit': + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + assumes "Zp_converges_to s a" + shows "a = res_lim s" + using unique_limit[of s a "res_lim s"] assms + convergent_imp_Zp_cauchy is_Zp_cauchy_imp_has_limit res_lim_in_Zp + by blast + +lemma Zp_converges_toE: + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + assumes "Zp_converges_to s a" + shows "\N. \k > N. s k n = a n" + by (metis assms(1) assms(2) assms(3) + convergent_imp_Zp_cauchy + res_lim_Zp_cauchy unique_limit') + +lemma Zp_converges_toI: + assumes "s \ closed_seqs Zp" + assumes "a \ carrier Zp" + assumes "\n. \N. \k > N. s k n = a n" + shows "Zp_converges_to s a" +proof- + have 0: "(a \ carrier Zp \ s \ closed_seqs Zp)" + using assms + by auto + have 1: "(\n::int. \k. \m>k. n < val_Zp_dist (s m) a) " + proof + fix n::int + show "\k. \m>k. n < val_Zp_dist (s m) a " + proof(cases "n < 0") + case True + have "\m>0. n < val_Zp_dist (s m) a " + proof + fix m ::nat + show "0 < m \ n < val_Zp_dist (s m) a" + proof + have 0: "eint n < 0" + by (simp add: True zero_eint_def) + have 1: " s m \ a \ carrier Zp" + using assms + by (simp add: closed_seqs_diff_closed) + thus " n < val_Zp_dist (s m) a" + using 0 True val_pos[of "s m \ a"] + unfolding val_Zp_dist_def + by auto + qed + qed + then show ?thesis + by blast + next + case False + then have P0: "n \ 0" + by auto + obtain N where N_def: "\k > N. s k (Suc (nat n)) = a (Suc (nat n))" + using assms by blast + have "\m>N. n < val_Zp_dist (s m) a " + proof + fix m + show " N < m \ n < val_Zp_dist (s m) a" + proof + assume A: "N < m" + then have A0: "s m (Suc (nat n)) = a (Suc (nat n))" + using N_def by blast + have "(Suc (nat n)) \ val_Zp_dist (s m) a" + using assms A0 val_Zp_dist_res_eq2[of "s m" a "Suc (nat n)"] closed_seqs_memE + by blast + thus "n < val_Zp_dist (s m) a" + using False + by (meson P0 eint_ord_simps(2) less_Suc_eq less_le_trans nat_less_iff) + qed + qed + then show ?thesis + by blast + qed + qed + show ?thesis + unfolding Zp_converges_to_def + using 0 1 + by (simp add: val_Zp_dist_def) +qed + +text\Sums and products of cauchy sequences are cauchy:\ + +lemma sum_of_Zp_cauchy_is_Zp_cauchy: + assumes "is_Zp_cauchy s" + assumes "is_Zp_cauchy t" + shows "is_Zp_cauchy (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t)" +proof(rule is_Zp_cauchyI) + show "(s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) \ closed_seqs Zp" + using assms seq_plus_closed is_Zp_cauchy_def by blast + show "\n. \N. \n0 n1. N < n0 \ N < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof- + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof- + obtain N1 where N1_def: "\n0 n1. N1 < n0 \ N1 < n1 \ s n0 n = s n1 n" + using assms(1) is_Zp_cauchy_imp_res_eventually_const_0 padic_integers_axioms by blast + obtain N2 where N2_def: "\n0 n1. N2 < n0 \ N2 < n1 \ t n0 n = t n1 n" + using assms(2) is_Zp_cauchy_imp_res_eventually_const_0 padic_integers_axioms by blast + obtain M where M_def: "M = max N1 N2" + by simp + have "\n0 n1. M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + fix n0 + show "\n1. M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + fix n1 + show " M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + assume A: "M < n0 \ M < n1 " + have Fs: " s n0 n = s n1 n" using A M_def N1_def by auto + have Ft: " t n0 n = t n1 n" using A M_def N2_def by auto + then show "(s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + using seq_plus_simp[of s t] assms + by (simp add: Fs is_Zp_cauchy_imp_closed residue_of_sum) + qed + qed + qed + then show ?thesis + by blast + qed + qed +qed + +lemma prod_of_Zp_cauchy_is_Zp_cauchy: + assumes "is_Zp_cauchy s" + assumes "is_Zp_cauchy t" + shows "is_Zp_cauchy (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t)" +proof(rule is_Zp_cauchyI) + show "(s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) \ closed_seqs Zp" + using assms(1) assms(2) is_Zp_cauchy_def seq_mult_closed by auto + show "\n. \N. \n0 n1. N < n0 \ N < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof- + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof- + obtain N1 where N1_def: "\n0 n1. N1 < n0 \ N1 < n1 \ s n0 n = s n1 n" + using assms(1) is_Zp_cauchy_imp_res_eventually_const_0 padic_integers_axioms by blast + obtain N2 where N2_def: "\n0 n1. N2 < n0 \ N2 < n1 \ t n0 n = t n1 n" + using assms(2) is_Zp_cauchy_imp_res_eventually_const_0 padic_integers_axioms by blast + obtain M where M_def: "M = max N1 N2" + by simp + have "\n0 n1. M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + fix n0 + show "\n1. M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + fix n1 + show " M < n0 \ M < n1 \ (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + proof + assume A: "M < n0 \ M < n1 " + have Fs: " s n0 n = s n1 n" using A M_def N1_def by auto + have Ft: " t n0 n = t n1 n" using A M_def N2_def by auto + then show "(s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n0 n = (s \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> t) n1 n" + using seq_mult_simp[of s t] is_Zp_cauchy_imp_closed assms + by (simp add: Fs residue_of_prod) + qed + qed + qed + then show ?thesis + by blast + qed + qed +qed + +text\Constant sequences are cauchy:\ + +lemma constant_is_Zp_cauchy: + assumes "is_constant_seq Zp s" + shows "is_Zp_cauchy s" +proof(rule is_Zp_cauchyI) + show "s \ closed_seqs Zp" + proof(rule closed_seqs_memI) + fix k + show "s k \ carrier Zp" + using assms is_constant_seq_imp_closed + by (simp add: is_constant_seq_imp_closed closed_seqs_memE) + qed + obtain x where "\ k. s k = x" + using assms + by (meson is_constant_seqE) + then show "\n. \N. \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n" + by simp +qed + +text\Scalar multiplies of cauchy sequences are cauchy:\ + +lemma smult_is_Zp_cauchy: + assumes "is_Zp_cauchy s" + assumes "a \ carrier Zp" + shows "is_Zp_cauchy (a \\<^bsub>Zp\<^bsup>\\<^esup>\<^esub> s)" + apply(rule is_Zp_cauchyI) + apply (meson U_function_ring.ring_seq_smult_closed U_function_ring_axioms assms(1) assms(2) is_Zp_cauchy_def) + using assms ring_seq_smult_eval[of s a] is_Zp_cauchy_imp_closed[of s] + by (metis res_lim_Zp_cauchy residue_of_prod) + +lemma Zp_cauchy_imp_approaches_res_lim: + assumes "is_Zp_cauchy s" + assumes "a = res_lim s" + obtains N where "\n. n> N \ val_Zp (a \ (s n)) > eint k" +proof- + have " (\n::int. \k. \m>k. n < val_Zp_dist (s m) a)" + using Zp_converges_to_def[of s a] assms is_Zp_cauchy_imp_has_limit[of s a] val_Zp_dist_def + by simp + then have "\N. \m>N. k < val_Zp_dist (s m) a" + by blast + then obtain N where N_def: "\m>N. k < val_Zp_dist (s m) a" + by blast + have "\n. n> N \ val_Zp (a \ (s n)) > k" + proof- + fix m + assume "m > N" + then have 0: "k < val_Zp_dist (s m) a" + using N_def + by (simp add: val_Zp_def) + show "k < val_Zp (a \ s m)" + using "0" assms(1) assms(2) is_Zp_cauchy_def closed_seqs_memE[of _ Zp] val_Zp_dist_def val_Zp_dist_sym res_lim_in_Zp by auto + qed + then show ?thesis + using that + by blast +qed + +(**************************************************************************************************) +(**************************************************************************************************) +(****************************) section\Continuous Functions\ (***********************************) +(**************************************************************************************************) +(**************************************************************************************************) + +text\ + For convenience, we will use a slightly unorthodox definition of continuity here. + Since $\mathbb{Z}_p$ is complete, a function is continuous if and only if its compositions + with cauchy sequences are cauchy sequences. Thus we can define a continuous function on + $\mathbb{Z}_p$ as a function which carries cauchy sequences to cauchy sequences under + composition. Note that this does not generalize to a definition of continuity for functions + defined on incomplete subsets of $\mathbb{Z}_p$. For example, the function $1/x$ defined on + $\mathbb{Z}_p - \{0\}$ clearly does not have this property but is continuous. However, towards + a proof of Hensel's Lemma we will only need to consider polynomial functions and so this + definition suffices for our purposes. +\ + +subsection\Defining Continuous Functions and Basic Examples\ +abbreviation Zp_constant_function ("\\<^bsub>Zp\<^esub>") where +"\\<^bsub>Zp\<^esub> a \ constant_function (carrier Zp) a" + +definition is_Zp_continuous ::"padic_int_fun \ bool" where +"is_Zp_continuous f = (f \ carrier (Fun Zp) \(\ s. is_Zp_cauchy s \ is_Zp_cauchy(f \ s)))" + +lemma Zp_continuous_is_Zp_closed: + assumes "is_Zp_continuous f" + shows "f \ carrier (Fun Zp)" + using assms is_Zp_continuous_def by blast + +lemma is_Zp_continuousI: + assumes "f \ carrier (Fun Zp)" + assumes "\s. is_Zp_cauchy s \ is_Zp_cauchy (f \ s)" + shows "is_Zp_continuous f" +proof- + have "(\ s. is_Zp_cauchy s \ is_Zp_cauchy(f \ s))" + proof + fix s + show "is_Zp_cauchy s \ is_Zp_cauchy (f \ s) " + by (simp add: assms(2)) + qed + then show ?thesis + using assms(1) is_Zp_continuous_def by blast +qed + +lemma Zp_continuous_is_closed: + assumes "is_Zp_continuous f" + shows "f \ carrier (Fun Zp)" + using assms unfolding is_Zp_continuous_def by blast + +lemma constant_is_Zp_continuous: + assumes "a \ carrier Zp" + shows "is_Zp_continuous (const a)" +proof(rule is_Zp_continuousI) + show "\\<^bsub>Zp\<^esub> a \ carrier (function_ring (carrier Zp) Zp)" + by (simp add: assms constant_function_closed) + show "\s. is_Zp_cauchy s \ is_Zp_cauchy (\\<^bsub>Zp\<^esub> a \ s) " + proof- + fix s + assume A: "is_Zp_cauchy s" + have "is_constant_seq Zp (\\<^bsub>Zp\<^esub> a \ s)" + using constant_function_comp_is_constant_seq[of a s] A assms + is_Zp_cauchy_imp_closed by blast + then show "is_Zp_cauchy (\\<^bsub>Zp\<^esub> a \ s)" + using A assms constant_is_Zp_cauchy + by blast + qed +qed + +lemma sum_of_cont_is_cont: + assumes "is_Zp_continuous f" + assumes "is_Zp_continuous g" + shows "is_Zp_continuous (f \\<^bsub>Fun Zp\<^esub> g)" + apply(rule is_Zp_continuousI) + using assms Zp_continuous_is_closed assms function_sum_comp_is_seq_sum[of _ f g] + apply (simp add: fun_add_closed) + using assms(1) assms(2) function_sum_comp_is_seq_sum is_Zp_cauchy_def is_Zp_continuous_def sum_of_Zp_cauchy_is_Zp_cauchy by auto + +lemma prod_of_cont_is_cont: + assumes "is_Zp_continuous f" + assumes "is_Zp_continuous g" + shows "is_Zp_continuous (f \\<^bsub>Fun Zp\<^esub> g)" + apply(rule is_Zp_continuousI) + using assms Zp_continuous_is_closed assms + apply (simp add: fun_mult_closed) + using function_mult_comp_is_seq_mult[of _ f g] assms(1) assms(2) is_Zp_cauchy_imp_closed + is_Zp_continuous_def prod_of_Zp_cauchy_is_Zp_cauchy by auto + +lemma smult_is_continuous: + assumes "is_Zp_continuous f" + assumes "a \ carrier Zp" + shows "is_Zp_continuous (a \\<^bsub>Fun Zp\<^esub> f)" + apply(rule is_Zp_continuousI) + using assms apply (simp add: assms function_smult_closed is_Zp_continuous_def) + using ring_seq_smult_comp_assoc assms + by (simp add: is_Zp_cauchy_imp_closed is_Zp_continuous_def smult_is_Zp_cauchy) + +lemma id_Zp_is_Zp_continuous: +"is_Zp_continuous ring_id" + apply(rule is_Zp_continuousI) + by (auto simp add: is_Zp_cauchy_imp_closed ring_id_seq_comp) + +lemma nat_pow_is_Zp_continuous: + assumes "is_Zp_continuous f" + shows "is_Zp_continuous (f[^]\<^bsub>Fun Zp\<^esub>(n::nat))" + apply(induction n) + using constant_is_Zp_continuous function_one_is_constant apply force + by (simp add: assms prod_of_cont_is_cont) + +lemma ring_id_pow_closed: +"(ring_id)[^]\<^bsub>Fun Zp\<^esub> (n::nat) \ carrier (Fun Zp)" + by (simp add: function_ring_is_monoid monoid.nat_pow_closed) + +lemma monomial_equation: + assumes "c \ carrier Zp" + shows "monomial c n = c \\<^bsub>Fun Zp\<^esub> (ring_id)[^]\<^bsub>Fun Zp\<^esub>n" + apply(rule function_ring_car_eqI) + apply (simp add: assms monomial_functions) + using assms function_smult_closed[of c "ring_id [^]\<^bsub>Fun Zp\<^esub> n"] ring_id_pow_closed apply blast + unfolding monomial_function_def + using assms function_smult_eval[of c "(ring_id)[^]\<^bsub>Fun Zp\<^esub> (n::nat)"] + function_nat_pow_eval[of ring_id _ n] + by (simp add: ring_id_eval ring_id_pow_closed) + +lemma monomial_is_Zp_continuous: + assumes "c \ carrier Zp" + shows "is_Zp_continuous (monomial c n)" + using monomial_equation[of c n] nat_pow_is_Zp_continuous + by (simp add: assms id_Zp_is_Zp_continuous smult_is_continuous) + +subsection\Composition by a Continuous Function Commutes with Taking Limits of Sequences\ + +text \ + Due to our choice of definition for continuity, a little bit of care is required to show that + taking the limit of a cauchy sequence commutes with composition by a continuous function. + For a sequence $(s_n)_{n \in \mathbb{N}}$ converging to a point $t$, we can consider the + alternating sequence $(s_0, t, s_1, t, s_2, t, \dots)$ which is also cauchy. Clearly + composition with $f$ yields $(f(s_0), f(t), f(s_1), f(t), f(s_2), f(t), \dots)$ from which + we can see that the limit must be $f(t)$. +\ +definition alt_seq where +"alt_seq s = (\k. (if (even k) then (s k) else (res_lim s)))" + +lemma alt_seq_Zp_cauchy: + assumes "is_Zp_cauchy s" + shows "is_Zp_cauchy (alt_seq s)" +proof(rule is_Zp_cauchyI) + show "(alt_seq s) \ closed_seqs Zp" + unfolding alt_seq_def using assms is_Zp_cauchy_imp_closed + by (simp add: closed_seqs_memE closed_seqs_memI res_lim_in_Zp) + fix n + show "\N. \n0 n1. N < n0 \ N < n1 \ alt_seq s n0 n = alt_seq s n1 n " + proof- + obtain N where N_def: " \n0 n1. N < n0 \ N < n1 \ s n0 n = s n1 n " + using assms is_Zp_cauchy_imp_res_eventually_const_0 + padic_integers_axioms + by blast + have "\n0 n1. N < n0 \ N < n1 \ alt_seq s n0 n = alt_seq s n1 n " + apply auto + proof- + fix n0 n1 + assume A: "N < n0" "N < n1" + show "alt_seq s n0 n = alt_seq s n1 n" + using N_def + unfolding alt_seq_def + by (smt A(1) A(2) assms lessI max_less_iff_conj + res_lim_Zp_cauchy padic_integers_axioms) + qed + then show ?thesis + by blast + qed +qed + +lemma alt_seq_limit: + assumes "is_Zp_cauchy s" + shows "res_lim(alt_seq s) = res_lim s" +proof- + have "\k. res_lim(alt_seq s) k = res_lim s k" + proof- + fix k + obtain N where N_def: "\ m. m> N \ s m k = res_lim s k" + using assms res_lim_Zp_cauchy + by blast + obtain N' where N'_def: "\ m. m> N' \ (alt_seq s) m k = res_lim (alt_seq s) k" + using assms res_lim_Zp_cauchy + alt_seq_Zp_cauchy + by blast + have "\m. m > (max N N') \ s m k = res_lim (alt_seq s) k" + proof- + fix m + assume A0: "m > (max N N')" + have A1: "s m k = res_lim s k" + using A0 N_def + by simp + have A2: "(alt_seq s) m k = res_lim (alt_seq s) k" + using A0 N'_def + by simp + have A3: "(alt_seq s) m k = res_lim s k" + using alt_seq_def A1 A2 + by presburger + show "s m k = res_lim (alt_seq s) k" + using A1 A2 A3 + by auto + qed + then have P:"\m. m > (max N N') \ (res_lim s k) = res_lim (alt_seq s) k" + using N_def + by auto + show "res_lim(alt_seq s) k = res_lim s k" + using P[of "Suc (max N N')"] + by auto + qed + then show ?thesis + by (simp add: ext) +qed + +lemma res_lim_pushforward: + assumes "is_Zp_continuous f" + assumes "is_Zp_cauchy s" + assumes "t = alt_seq s" + shows "res_lim (f \ t) = f (res_lim t)" +proof- + have 0: "Zp_converges_to (f \ t) (res_lim (f \ t))" + using assms alt_seq_Zp_cauchy is_Zp_cauchy_imp_has_limit + is_Zp_continuous_def + by blast + have "\k. res_lim (f \ t) k = f (res_lim t) k" + proof- + fix k + show "res_lim (f \ t) k = f (res_lim t) k" + proof- + obtain N where N_def: "\m. m> N \ (f \ t) m k = (res_lim (f \ t)) k" + using 0 + by (meson convergent_imp_Zp_cauchy Zp_converges_to_def res_lim_Zp_cauchy) + obtain M where M_def: "M = 2*(Suc N) + 1" + by simp + have 0: "t M = res_lim s" + using assms + unfolding alt_seq_def + by (simp add: M_def) + have 1: "(f \ t) M k = (res_lim (f \ t)) k" + using N_def M_def + by auto + have 2: "(f \ t) M k = f (t M) k" + by simp + have 3: "(f \ t) M k = f (res_lim s) k" + using 0 2 by simp + have 4: "(f \ t) M k = f (res_lim t) k" + using 3 assms alt_seq_limit[of s] + by auto + show ?thesis + using 4 1 by auto + qed + qed + then show ?thesis by(simp add: ext) +qed + +lemma res_lim_pushforward': + assumes "is_Zp_continuous f" + assumes "is_Zp_cauchy s" + assumes "t = alt_seq s" + shows "res_lim (f \ s) = res_lim (f \ t)" +proof- + obtain a where a_def: "a = res_lim (f \ s)" + by simp + obtain b where b_def: "b = res_lim (f \ t)" + by simp + have "\k. a k = b k" + proof- + fix k + obtain Na where Na_def: "\m. m > Na \ (f \ s) m k = a k" + using a_def assms is_Zp_continuous_def + padic_integers_axioms res_lim_Zp_cauchy + by blast + obtain Nb where Nb_def: "\m. m > Nb \ (f \ t) m k = b k" + using b_def assms is_Zp_continuous_def + padic_integers_axioms res_lim_Zp_cauchy + alt_seq_Zp_cauchy + by blast + obtain M where M_def: "M = 2*(max Na Nb) + 1" + by simp + have M0: "odd M" + by (simp add: M_def) + have M1: "M > Na" + using M_def + by auto + have M2: "M > Nb" + using M_def + by auto + have M3: "t M = res_lim s" + using assms alt_seq_def M0 + by auto + have M4: "((f \ t) M) = f (res_lim s)" + using M3 + by auto + have M5: "((f \ t) M) k = b k" + using M2 Nb_def by auto + have M6: "f (res_lim s) = f (res_lim t)" + using assms alt_seq_limit[of s] + by auto + have M7: "f (res_lim t) k = b k" + using M4 M5 M6 by auto + have M8: "(f \ s) M k = (f \ s) (Suc M) k" + using M1 Na_def by auto + have M9: "(f \ s) (Suc M) = (f \ t) (Suc M)" + using assms unfolding alt_seq_def + using M_def + by auto + have M10: "(f \ t) M k = (f \ t) (Suc M) k" + using M2 Nb_def by auto + have M11: "(f \ t) M k = (f \ s) M k" + using M10 M8 M9 by auto + show "a k = b k" + using M1 M11 M5 Na_def by auto + qed + then show ?thesis using a_def b_def ext[of a b] by auto +qed + +lemma continuous_limit: + assumes "is_Zp_continuous f" + assumes "is_Zp_cauchy s" + shows "Zp_converges_to (f \ s) (f (res_lim s))" +proof- + obtain t where t_def: "t = alt_seq s" + by simp + have 0: "Zp_converges_to (f \ s) (res_lim (f \ s))" + using assms(1) assms(2) is_Zp_continuous_def + is_Zp_cauchy_imp_has_limit padic_integers_axioms by blast + have 1: "Zp_converges_to (f \ s) (res_lim (f \ t))" + using "0" assms(1) assms(2) res_lim_pushforward' t_def by auto + have 2: "Zp_converges_to (f \ s) (f (res_lim t))" + using "1" assms(1) assms(2) res_lim_pushforward padic_integers_axioms t_def by auto + then show "Zp_converges_to (f \ s) (f (res_lim s))" + by (simp add: alt_seq_limit assms(2) t_def) +qed + + +end +end diff --git a/thys/Padic_Ints/Padic_Integers.thy b/thys/Padic_Ints/Padic_Integers.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Padic_Integers.thy @@ -0,0 +1,2505 @@ +theory Padic_Integers + imports Padic_Construction + Extended_Int + Supplementary_Ring_Facts + "HOL-Algebra.Subrings" + "HOL-Number_Theory.Residues" + "HOL-Algebra.Multiplicative_Group" + +begin + +text\ + In what follows we establish a locale for reasoning about the ring of $p$-adic integers for a + fixed prime $p$. We will elaborate on the basic metric properties of $\mathbb{Z}_p$ and construct + the angular component maps to the residue rings. +\ + +section\A Locale for $p$-adic Integer Rings\ +locale padic_integers = +fixes Zp:: "_ ring" (structure) +fixes p +defines "Zp \ padic_int p" +assumes prime: "prime p" + +sublocale padic_integers < UPZ?: UP Zp "UP Zp" + by simp + +sublocale padic_integers < Zp?:UP_cring Zp "UP Zp" + unfolding UP_cring_def + by(auto simp add: Zp_def padic_int_is_cring prime) + +sublocale padic_integers < Zp?:ring Zp + using Zp_def cring.axioms(1) padic_int_is_cring prime + by blast + +sublocale padic_integers < Zp?:cring Zp + by (simp add: Zp_def padic_int_is_cring prime) + +sublocale padic_integers < Zp?:domain Zp + by (simp add: Zp_def padic_int_is_domain padic_integers.prime padic_integers_axioms) + + +context padic_integers +begin + +lemma Zp_defs: +"\ = padic_one p" +"\ = padic_zero p" +"carrier Zp = padic_set p" +"(\) = padic_mult p" +"(\) = padic_add p" + unfolding Zp_def using padic_int_simps by auto + +end + +(*************************************************************************************************) +(*************************************************************************************************) +(***********************)section \Residue Rings\(*************************************) +(*************************************************************************************************) +(*************************************************************************************************) + +lemma(in field) field_inv: + assumes "a \ carrier R" + assumes "a \\" + shows "inv\<^bsub>R\<^esub> a \ a = \" + "a \ inv\<^bsub>R\<^esub> a = \" + "inv \<^bsub>R\<^esub> a \ carrier R" +proof- + have "a \ Units R" + using assms by (simp add: local.field_Units) + then show "inv\<^bsub>R\<^esub> a \ a = \" + by simp + show "a \ inv a = \" + using \a \ Units R\ by auto + show "inv \<^bsub>R\<^esub> a \ carrier R" + by (simp add: \a \ Units R\) +qed + +text\$p_residue$ defines the standard projection maps between residue rings:\ + +definition(in padic_integers) p_residue:: "nat \ int \ _" where +"p_residue m n \ residue (p^m) n" + +lemma(in padic_integers) p_residue_alt_def: +"p_residue m n = n mod (p^m)" + using residue_def + by (simp add: p_residue_def) + +lemma(in padic_integers) p_residue_range: +"p_residue m n \ {0.. k" + shows "p_residue k (p_residue m n) = p_residue k n" + using assms + unfolding p_residue_def residue_def + by (simp add: le_imp_power_dvd mod_mod_cancel) + +text\Compatibility of p\_residue with elements of $\mathbb{Z}_p$:\ + +lemma(in padic_integers) p_residue_padic_int: + assumes "x \ carrier Zp" + assumes "m \ k" + shows "p_residue k (x m) = x k" + using Zp_def assms(1) assms(2) padic_set_res_coherent prime + by (simp add: p_residue_def padic_int_simps(5)) + +text\Definition of residue rings:\ + +abbreviation(in padic_integers) (input) Zp_res_ring:: "nat \ _ ring" where +"(Zp_res_ring n) \ residue_ring (p^n)" + +lemma (in padic_integers) p_res_ring_zero: +"\\<^bsub>Zp_res_ring k\<^esub> = 0" + by (simp add: residue_ring_def) + +lemma (in padic_integers) p_res_ring_one: + assumes "k > 0" + shows "\\<^bsub>Zp_res_ring k\<^esub> = 1" + using assms + by (simp add: residue_ring_def) + +lemma (in padic_integers) p_res_ring_car: +"carrier (Zp_res_ring k) = {0.. carrier (Zp_res_ring m)" + using p_residue_range residue_ring_def prime prime_gt_0_nat p_residue_def + by fastforce + +text\First residue ring is a field:\ + +lemma(in padic_integers) p_res_ring_1_field: +"field (Zp_res_ring 1)" + by (metis int_nat_eq power_one_right prime prime_ge_0_int prime_nat_int_transfer residues_prime.intro residues_prime.is_field) + +text\$0^{th}$ residue ring is the zero ring:\ + +lemma(in padic_integers) p_res_ring_0: +"carrier (Zp_res_ring 0) = {0}" + by (simp add: residue_ring_def) + +lemma(in padic_integers) p_res_ring_0': + assumes "x \ carrier (Zp_res_ring 0)" + shows "x = 0" + using p_res_ring_0 assms by blast + +text\If $m>0$ then $Zp\_res\_ring m$ is an instance of the residues locale:\ + +lemma(in padic_integers) p_residues: + assumes "m >0" + shows "residues (p^m)" +proof- + have "p^m > 1" + using assms + by (simp add: prime prime_gt_1_int) + then show "residues (p^m)" + using less_imp_of_nat_less residues.intro by fastforce +qed + +text\If $m>0$ then $Zp\_res\_ring m$ is a commutative ring:\ + +lemma(in padic_integers) R_cring: + assumes "m >0" + shows "cring (Zp_res_ring m)" + using p_residues assms residues.cring by auto + +lemma(in padic_integers) R_comm_monoid: + assumes "m >0" + shows "comm_monoid (Zp_res_ring m)" + by (simp add: assms p_residues residues.comm_monoid) + +lemma(in padic_integers) zero_rep: +"\ = (\m. (p_residue m 0))" + unfolding p_residue_def + using Zp_defs(2) padic_zero_simp(1) residue_def residue_ring_def by auto + +text\The operations on residue rings are just the standard operations on the integers $\mod p^n$. This means that the basic closure properties and algebraic properties of operations on these rings hold for all integers, not just elements of the ring carrier:\ + +lemma(in padic_integers) residue_add: + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x + y) mod p^k" + unfolding residue_ring_def + by simp + +lemma(in padic_integers) residue_add_closed: + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" + using p_residue_def p_residue_range residue_def residue_ring_def by auto + +lemma(in padic_integers) residue_add_closed': + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ {0..\<^bsub>Zp_res_ring k\<^esub> y) = (x * y) mod p^k" + unfolding residue_ring_def + by simp + +lemma(in padic_integers) residue_mult_closed: + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" + using p_residue_def p_residue_range residue_def residue_ring_def by auto + +lemma(in padic_integers) residue_mult_closed': + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ {0..\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> z = x \\<^bsub>Zp_res_ring k\<^esub> (y \\<^bsub>Zp_res_ring k\<^esub> z)" + using residue_add + by (simp add: add.commute add.left_commute mod_add_right_eq) + +lemma(in padic_integers) residue_mult_comm: + shows "x \\<^bsub>Zp_res_ring k\<^esub> y = y \\<^bsub>Zp_res_ring k\<^esub> x" + using residue_mult + by (simp add: mult.commute) + +lemma(in padic_integers) residue_mult_assoc: + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> z = x \\<^bsub>Zp_res_ring k\<^esub> (y \\<^bsub>Zp_res_ring k\<^esub> z)" + using residue_mult + by (simp add: mod_mult_left_eq mod_mult_right_eq mult.assoc) + +lemma(in padic_integers) residue_add_comm: + shows "x \\<^bsub>Zp_res_ring k\<^esub> y = y \\<^bsub>Zp_res_ring k\<^esub> x" + using residue_add + by presburger + +lemma(in padic_integers) residue_minus_car: + assumes "y \ carrier (Zp_res_ring k)" + shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x - y) mod p^k" +proof(cases "k = 0") + case True + then show ?thesis + using residue_ring_def a_minus_def + by(simp add: a_minus_def residue_ring_def) +next + case False + have "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = x \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> y \\<^bsub>Zp_res_ring k\<^esub> y)" + by (simp add: a_minus_def residue_add_assoc) + then have 0: "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = x mod p^k" + using assms False + by (smt cring.cring_simprules(9) prime residue_add residues.cring residues.res_zero_eq residues_n) + have 1: "x mod p ^ k = ((x - y) mod p ^ k + y) mod p ^ k" + proof - + have f1: "x - y = x + - 1 * y" + by auto + have "y + (x + - 1 * y) = x" + by simp + then show ?thesis + using f1 by presburger + qed + have "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = (x - y) mod p^k \\<^bsub>Zp_res_ring k\<^esub> y" + using residue_add[of k "(x - y) mod p^k" y] 0 1 + by linarith + then show ?thesis using assms residue_add_closed + by (metis False a_minus_def cring.cring_simprules(10) cring.cring_simprules(19) + prime residues.cring residues.mod_in_carrier residues_n) +qed + +lemma(in padic_integers) residue_a_inv: + shows "\\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" +proof- + have "y \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) = (y mod p^k) \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) " + using residue_minus_car[of "\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" k y] residue_add + by (simp add: mod_add_left_eq) + then have 0: "y \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) = \\<^bsub>Zp_res_ring k\<^esub>" + by (metis cring.cring_simprules(17) p_res_ring_zero padic_integers.p_res_ring_0' + padic_integers_axioms prime residue_add_closed residues.cring residues.mod_in_carrier residues_n) + have 1: "(\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) \\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub>" + using residue_add_comm 0 by auto + have 2: "\x. x \ carrier (Zp_res_ring k) \ x \\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub> \ y \\<^bsub>Zp_res_ring k\<^esub> x = \\<^bsub>Zp_res_ring k\<^esub> \ x = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" + using 0 1 + by (metis cring.cring_simprules(3) cring.cring_simprules(8) mod_by_1 padic_integers.p_res_ring_0' + padic_integers.p_res_ring_zero padic_integers_axioms power_0 prime residue_1_prop + residue_add_assoc residues.cring residues.mod_in_carrier residues_n) + have 3: "carrier (add_monoid (residue_ring (p ^ k))) = carrier (Zp_res_ring k)" + by simp + have 4: "(\\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub>) = (\\<^bsub>Zp_res_ring k\<^esub>)" + by simp + have 5: "\x. x \ carrier (add_monoid (residue_ring (p ^ k))) \ + x \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> y = \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> \ + y \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> x = \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> + \ x = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" + using 0 1 2 3 4 + by simp + show ?thesis + unfolding a_inv_def m_inv_def + apply(rule the_equality) + using 1 2 3 4 5 unfolding a_inv_def m_inv_def + apply (metis (no_types, lifting) "0" "1" cring.cring_simprules(3) mod_by_1 + monoid.select_convs(2) padic_integers.p_res_ring_zero padic_integers_axioms power_0 prime + residue_1_prop residue_add_closed residues.cring residues.mod_in_carrier residues_n) + using 1 2 3 4 5 unfolding a_inv_def m_inv_def + by blast +qed + +lemma(in padic_integers) residue_a_inv_closed: +"\\<^bsub>Zp_res_ring k\<^esub> y \ carrier (Zp_res_ring k)" + apply(cases "k = 0") + apply (metis add.comm_neutral add.commute + atLeastLessThanPlusOne_atLeastAtMost_int + insert_iff mod_by_1 p_res_ring_car p_res_ring_zero padic_integers.p_res_ring_0 + padic_integers_axioms power_0 residue_1_prop residue_a_inv) + by (simp add: prime residues.mod_in_carrier residues.res_neg_eq residues_n) + +lemma(in padic_integers) residue_minus: +"(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x - y) mod p^k" + using residue_minus_car[of "y mod p^k" k x] residue_a_inv[of k y] unfolding a_minus_def + by (metis a_minus_def mod_diff_right_eq p_residue_alt_def p_residue_range') + +lemma(in padic_integers) residue_minus_closed: +"(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" + by (simp add: a_minus_def residue_add_closed) + +lemma (in padic_integers) residue_plus_zero_r: +"0 \\<^bsub>Zp_res_ring k\<^esub> y = y mod p^k" + by (simp add: residue_add) + +lemma (in padic_integers) residue_plus_zero_l: +"y \\<^bsub>Zp_res_ring k\<^esub> 0 = y mod p^k" + by (simp add: residue_add) + +lemma (in padic_integers) residue_times_zero_r: +"0 \\<^bsub>Zp_res_ring k\<^esub> y = 0" + by (simp add: residue_mult) + +lemma (in padic_integers) residue_times_zero_l: +"y \\<^bsub>Zp_res_ring k\<^esub> 0 = 0" + by (simp add: residue_mult) + +lemma (in padic_integers) residue_times_one_r: +"1 \\<^bsub>Zp_res_ring k\<^esub> y = y mod p^k" + by (simp add: residue_mult) + +lemma (in padic_integers) residue_times_one_l: +"y \\<^bsub>Zp_res_ring k\<^esub> 1 = y mod p^k" + by (simp add: residue_mult_comm residue_times_one_r) + +text\Similarly to the previous lemmas, many identities about taking residues of $p$-adic integers hold even for elements which lie outside the carrier of $\mathbb{Z}_p$:\ + +lemma (in padic_integers) residue_of_sum: +"(a \ b) k = (a k) \\<^bsub>Zp_res_ring k\<^esub> (b k)" + using Zp_def residue_ring_def[of "p^k"] Zp_defs(5) padic_add_res + by auto + +lemma (in padic_integers) residue_of_sum': + "(a \ b) k = ((a k) + (b k)) mod p^k" + using residue_add residue_of_sum by auto + +lemma (in padic_integers) residue_closed[simp]: + assumes "b \ carrier Zp" + shows "b k \ carrier (Zp_res_ring k)" + using Zp_def assms padic_integers.Zp_defs(3) padic_integers_axioms padic_set_res_closed + by auto + +lemma (in padic_integers) residue_of_diff: + assumes "b \ carrier Zp" + shows "(a \ b) k = (a k) \\<^bsub>Zp_res_ring k\<^esub> (b k)" + unfolding a_minus_def + using Zp_def add.inv_closed assms(1) padic_a_inv prime residue_of_sum by auto + +lemma (in padic_integers) residue_of_prod: +"(a \ b) k = (a k) \ \<^bsub>Zp_res_ring k\<^esub> (b k)" + by (simp add: Zp_defs(4) padic_mult_def) + +lemma (in padic_integers) residue_of_prod': +"(a \ b) k = ((a k) * (b k)) mod (p^k)" + by (simp add: residue_mult residue_of_prod) + +lemma (in padic_integers) residue_of_one: + assumes "k > 0" + shows "\ k = \\<^bsub>Zp_res_ring k\<^esub>" + "\ k = 1" + apply (simp add: Zp_defs(1) assms padic_one_simp(1)) + by (simp add: Zp_def assms padic_int_simps(1) padic_one_simp(1) residue_ring_def) + +lemma (in padic_integers) residue_of_zero: + shows "\ k = \\<^bsub>Zp_res_ring k\<^esub>" + "\ k = 0" + apply (simp add: Zp_defs(2) padic_zero_simp(1)) + by (simp add: p_residue_alt_def zero_rep) + +lemma(in padic_integers) Zp_residue_mult_zero: + assumes "a k = 0" + shows "(a \ b) k = 0" "(b \ a) k = 0" + using assms residue_of_prod' + by auto + +lemma(in padic_integers) Zp_residue_add_zero: + assumes "b \ carrier Zp" + assumes "(a:: padic_int) k = 0" + shows "(a \ b) k = b k" "(b \ a) k = b k" + apply (metis Zp_def assms(1) assms(2) cring.cring_simprules(8) mod_by_1 padic_int_is_cring power.simps(1) + prime residue_add_closed residue_of_sum residue_of_sum' residues.cring residues.res_zero_eq residues_n) + by (metis Zp_def assms(1) assms(2) cring.cring_simprules(16) mod_by_1 padic_int_is_cring + power.simps(1) prime residue_add_closed residue_of_sum residue_of_sum' residues.cring + residues.res_zero_eq residues_n) + +text\P-adic addition and multiplication are globally additive and associative:\ + +lemma padic_add_assoc0: +assumes "prime p" +shows "padic_add p (padic_add p x y) z = padic_add p x (padic_add p y z)" + using assms unfolding padic_add_def + by (simp add: padic_integers.residue_add_assoc padic_integers_def) + +lemma(in padic_integers) add_assoc: +"a \ b \ c = a \ (b \ c)" + using padic_add_assoc0[of p a b c] prime Zp_defs by auto + +lemma padic_add_comm0: +assumes "prime p" +shows "(padic_add p x y)= (padic_add p y x)" + using assms unfolding padic_add_def + using padic_integers.residue_add_comm[of p] + by (simp add: padic_integers_def) + +lemma(in padic_integers) add_comm: +"a \ b = b \ a" + using padic_add_comm0[of p a b] prime Zp_defs by auto + +lemma padic_mult_assoc0: +assumes "prime p" +shows "padic_mult p (padic_mult p x y) z = padic_mult p x (padic_mult p y z)" + using assms unfolding padic_mult_def + by (simp add: padic_integers.residue_mult_assoc padic_integers_def) + +lemma(in padic_integers) mult_assoc: +"a \ b \ c = a \ (b \ c)" + using padic_mult_assoc0[of p a b c] prime Zp_defs by auto + +lemma padic_mult_comm0: +assumes "prime p" +shows "(padic_mult p x y)= (padic_mult p y x)" + using assms unfolding padic_mult_def + using padic_integers.residue_mult_comm[of p] + by (simp add: padic_integers_def) + +lemma(in padic_integers) mult_comm: +"a \ b = b \ a" + using padic_mult_comm0[of p a b] prime Zp_defs by auto + +lemma(in padic_integers) mult_zero_l: +"a \ \ = \" +proof fix x show "(a \ \) x = \ x" + using Zp_residue_mult_zero[of \ x a] + by (simp add: residue_of_zero(2)) +qed + +lemma(in padic_integers) mult_zero_r: +"\ \ a = \" + using mult_zero_l mult_comm by auto + +lemma (in padic_integers) p_residue_ring_car_memI: + assumes "(m::int) \0" + assumes "m < p^k" + shows "m \ carrier (Zp_res_ring k)" + using residue_ring_def[of "p^k"] assms(1) assms(2) + by auto + +lemma (in padic_integers) p_residue_ring_car_memE: + assumes "m \ carrier (Zp_res_ring k)" + shows "m < p^k" "m \ 0" + using assms residue_ring_def by auto + +lemma (in padic_integers) residues_closed: + assumes "a \ carrier Zp" + shows "a k \ carrier (Zp_res_ring k)" + using Zp_def assms padic_integers.Zp_defs(3) padic_integers_axioms padic_set_res_closed by blast + +lemma (in padic_integers) mod_in_carrier: + "a mod (p^n) \ carrier (Zp_res_ring n)" + using p_residue_alt_def p_residue_range' by auto + +lemma (in padic_integers) Zp_residue_a_inv: + assumes "a \ carrier Zp" + shows "(\ a) k = \\<^bsub>Zp_res_ring k\<^esub> (a k)" + "(\ a) k = (- (a k)) mod (p^k)" + using Zp_def assms padic_a_inv prime apply auto[1] + using residue_a_inv + by (metis Zp_def assms mod_by_1 p_res_ring_zero padic_a_inv padic_integers.prime + padic_integers_axioms power_0 residue_1_prop residues.res_neg_eq residues_n) + +lemma (in padic_integers) residue_of_diff': + assumes "b \ carrier Zp" + shows "(a \ b) k = ((a k) - (b k)) mod (p^k)" + by (simp add: assms residue_minus_car residue_of_diff residues_closed) + +lemma (in padic_integers) residue_UnitsI: + assumes "n \ 1" + assumes "(k::int) > 0" + assumes "k < p^n" + assumes "coprime k p" + shows "k \ Units (Zp_res_ring n)" + using residues.res_units_eq assms + by (metis (mono_tags, lifting) coprime_power_right_iff mem_Collect_eq not_one_le_zero prime residues_n) + +lemma (in padic_integers) residue_UnitsE: + assumes "n \ 1" + assumes "k \ Units (Zp_res_ring n)" + shows "coprime k p" + using residues.res_units_eq assms + by (simp add: p_residues) + +lemma(in padic_integers) residue_units_nilpotent: + assumes "m > 0" + assumes "k = card (Units (Zp_res_ring m))" + assumes "x \ (Units (Zp_res_ring m))" + shows "x[^]\<^bsub>Zp_res_ring m\<^esub> k = 1" +proof- + have 0: "group (units_of (Zp_res_ring m))" + using assms(1) cring_def monoid.units_group padic_integers.R_cring + padic_integers_axioms ring_def + by blast + have 1: "finite (Units (Zp_res_ring m))" + using p_residues assms(1) residues.finite_Units + by auto + have 2: "x[^]\<^bsub>units_of (Zp_res_ring m)\<^esub> (order (units_of (Zp_res_ring m))) = \\<^bsub>units_of (Zp_res_ring m)\<^esub>" + by (metis "0" assms(3) group.pow_order_eq_1 units_of_carrier) + then show ?thesis + by (metis "1" assms(1) assms(2) assms(3) cring.units_power_order_eq_one + padic_integers.R_cring padic_integers.p_residues padic_integers_axioms residues.res_one_eq) +qed + +lemma (in padic_integers) residue_1_unit: + assumes "m > 0" + shows "1 \ Units (Zp_res_ring m)" + "\\<^bsub>Zp_res_ring m\<^esub> \ Units (Zp_res_ring m)" +proof- + have 0: "group (units_of (Zp_res_ring m))" + using assms(1) cring_def monoid.units_group padic_integers.R_cring + padic_integers_axioms ring_def + by blast + have 1: "1 = \\<^bsub>units_of (Zp_res_ring m)\<^esub>" + by (simp add: residue_ring_def units_of_def) + have "\\<^bsub>units_of (Zp_res_ring m)\<^esub> \ carrier (units_of (Zp_res_ring m))" + using 0 Group.monoid.intro[of "units_of (Zp_res_ring m)"] + by (simp add: group.is_monoid) + then show "1 \ Units (Zp_res_ring m)" + using 1 by (simp add: units_of_carrier) + then show " \\<^bsub>Zp_res_ring m\<^esub> \ Units (Zp_res_ring m) " + by (simp add: "1" units_of_one) +qed + +lemma (in padic_integers) zero_not_in_residue_units: + assumes "n \ 1" + shows "(0::int) \ Units (Zp_res_ring n)" + using assms p_residues residues.res_units_eq by auto + +text\Cardinality bounds on the units of residue rings:\ + +lemma (in padic_integers) residue_units_card_geq_2: + assumes "n \2" + shows "card (Units (Zp_res_ring n)) \ 2" +proof(cases "p = 2") + case True + then have "(3::int) \ Units (Zp_res_ring n)" + proof- + have "p \2" + by (simp add: True) + then have "p^n \ 2^n" + using assms + by (simp add: True) + then have "p^n \ 4" + using assms power_increasing[of 2 n 2] + by (simp add: True) + then have "(3::int) < p^n" + by linarith + then have 0: "(3::int) \ carrier (Zp_res_ring n)" + by (simp add: residue_ring_def) + have 1: "coprime 3 p" + by (simp add: True numeral_3_eq_3) + show ?thesis using residue_UnitsI[of n "3::int"] + using "1" \3 < p ^ n\ assms by linarith + qed + then have 0: "{(1::int), 3} \ Units (Zp_res_ring n)" + using assms padic_integers.residue_1_unit padic_integers_axioms by auto + have 1: "finite (Units (Zp_res_ring n))" + using assms padic_integers.p_residues padic_integers_axioms residues.finite_Units by auto + have 2: "{(1::int),3}\Units (Zp_res_ring n)" + using "0" by auto + have 3: "card {(1::int),3} = 2" + by simp + show ?thesis + using 2 1 + by (metis "3" card_mono) +next + case False + have "1 \ Units (Zp_res_ring n)" + using assms padic_integers.residue_1_unit padic_integers_axioms by auto + have "2 \ Units (Zp_res_ring n)" + apply(rule residue_UnitsI) + using assms apply linarith + apply simp + proof- + show "2 < p^n" + proof- + have "p^n > p" + by (metis One_nat_def assms le_simps(3) numerals(2) power_one_right + power_strict_increasing_iff prime prime_gt_1_int) + then show ?thesis using False prime prime_gt_1_int[of p] + by auto + qed + show "coprime 2 p" + using False + by (metis of_nat_numeral prime prime_nat_int_transfer primes_coprime two_is_prime_nat) + qed + then have 0: "{(1::int), 2} \ Units (Zp_res_ring n)" + using \1 \ Units (Zp_res_ring n)\ by blast + have 1: "card {(1::int),2} = 2" + by simp + then show ?thesis + using residues.finite_Units 0 + by (metis One_nat_def assms card_mono dual_order.trans + le_simps(3) one_le_numeral padic_integers.p_residues padic_integers_axioms) +qed + +lemma (in padic_integers) residue_ring_card: +"finite (carrier (Zp_res_ring n)) \ card (carrier (Zp_res_ring n)) = nat (p^n)" + using p_res_ring_car[of n] + by simp + +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 padic_integers) residue_of_Units: + assumes "k > 0" + assumes "a \ Units Zp" + shows "a k \ Units (Zp_res_ring k)" +proof- + have 0: "a k \\<^bsub>Zp_res_ring k\<^esub> (inv \<^bsub>Zp\<^esub> a) k = 1" + by (metis R.Units_r_inv assms(1) assms(2) residue_of_one(2) residue_of_prod) + have 1: "a k \ carrier (Zp_res_ring k)" + by (simp add: R.Units_closed assms(2) residues_closed) + have 2: "(inv \<^bsub>Zp\<^esub> a) k \ carrier (Zp_res_ring k)" + by (simp add: assms(2) residues_closed) + show ?thesis using 0 1 2 comm_monoid.UnitsI[of "Zp_res_ring k"] + using assms(1) p_residues residues.comm_monoid residues.res_one_eq + by presburger +qed + +(*************************************************************************************************) +(*************************************************************************************************) +(**********************)section\$int$ and $nat$ inclusions in $\mathbb{Z}_p$.\ (****************************) +(*************************************************************************************************) +(*************************************************************************************************) + +lemma(in ring) int_inc_zero: +"[(0::int)]\ \ = \" + by (simp add: add.int_pow_eq_id) + +lemma(in ring) int_inc_zero': + assumes "x \ carrier R" + shows "[(0::int)] \ x = \" + by (simp add: add.int_pow_eq_id assms) + +lemma(in ring) nat_inc_zero: +"[(0::nat)]\ \ = \" + by auto + +lemma(in ring) nat_mult_zero: +"[(0::nat)]\ x = \" + by simp + +lemma(in ring) nat_inc_closed: + fixes n::nat + shows "[n] \ \ \ carrier R" + by simp + +lemma(in ring) nat_mult_closed: + fixes n::nat + assumes "x \ carrier R" + shows "[n] \ x \ carrier R" + by (simp add: assms) + +lemma(in ring) int_inc_closed: + fixes n::int + shows "[n] \ \ \ carrier R" + by simp + +lemma(in ring) int_mult_closed: + fixes n::int + assumes "x \ carrier R" + shows "[n] \ x \ carrier R" + by (simp add: assms) + +lemma(in ring) nat_inc_prod: + fixes n::nat + fixes m::nat + shows "[m]\([n] \ \) = [(m*n)]\\" + by (simp add: add.nat_pow_pow mult.commute) + +lemma(in ring) nat_inc_prod': + fixes n::nat + fixes m::nat + shows "[(m*n)]\\ = [m]\ \ \ ([n] \ \)" + by (simp add: add.nat_pow_pow add_pow_rdistr) + +lemma(in padic_integers) Zp_nat_inc_zero: + shows "[(0::nat)] \ x = \" + by simp + +lemma(in padic_integers) Zp_int_inc_zero': + shows "[(0::int)] \ x = \" + using Zp_nat_inc_zero[of x] + unfolding add_pow_def int_pow_def by auto + +lemma(in padic_integers) Zp_nat_inc_closed: + fixes n::nat + shows "[n] \ \ \ carrier Zp" + by simp + +lemma(in padic_integers) Zp_nat_mult_closed: + fixes n::nat + assumes "x \ carrier Zp" + shows "[n] \ x \ carrier Zp" + by (simp add: assms) + +lemma(in padic_integers) Zp_int_inc_closed: + fixes n::int + shows "[n] \ \ \ carrier Zp" + by simp + +lemma(in padic_integers) Zp_int_mult_closed: + fixes n::int + assumes "x \ carrier Zp" + shows "[n] \ x \ carrier Zp" + by (simp add: assms) + +text\The following lemmas give a concrete description of the inclusion of integers and natural numbers into $\mathbb{Z}_p$:\ + +lemma(in padic_integers) Zp_nat_inc_rep: + fixes n::nat + shows "[n] \ \ = (\ m. p_residue m n)" + apply(induction n) + apply (simp add: zero_rep) +proof- + case (Suc n) + fix n + assume A: "[n] \ \ = (\m. p_residue m (int n))" + then have 0: "[Suc n] \ \ = [n]\\ \ \" by auto + show "[Suc n] \ \ = (\m. p_residue m (Suc n))" + proof fix m + show "([Suc n] \ \) m = p_residue m (int (Suc n)) " + proof(cases "m=0") + case True + have 0: "([Suc n] \ \) m \ carrier (Zp_res_ring m)" + using Zp_nat_inc_closed padic_set_res_closed + by (simp add: residues_closed) + then have "([Suc n] \ \) m = 0" + using p_res_ring_0 True by blast + then show ?thesis + by (metis True p_res_ring_0' p_residue_range') + next + case False + then have R: "residues (p^m)" + by (simp add: prime residues_n) + have "([Suc n] \ \) m = ([n]\\ \ \) m" + by (simp add: "0") + then have P0: "([Suc n] \ \) m = p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> \\<^bsub>Zp_res_ring m\<^esub>" + using A False Zp_def padic_add_res padic_one_def Zp_defs(5) + padic_integers.residue_of_one(1) padic_integers_axioms by auto + then have P1:"([Suc n] \ \) m = p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m (1::int)" + by (metis R ext p_residue_alt_def residue_add_assoc residue_add_comm residue_plus_zero_r + residue_times_one_r residue_times_zero_l residues.res_one_eq) + have P2: "p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m 1 = ((int n) mod (p^m)) \\<^bsub>Zp_res_ring m\<^esub> 1" + using R P0 P1 residue_def residues.res_one_eq + by (simp add: residues.res_one_eq p_residue_alt_def) + have P3:"((int n) mod (p^m)) \\<^bsub>Zp_res_ring m\<^esub> 1 = ((int n) + 1) mod (p^m)" + using R residue_ring_def by (simp add: mod_add_left_eq) + have "p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m 1 = (int (Suc n)) mod (p^m)" + by (metis P2 P3 add.commute of_nat_Suc p_residue_alt_def residue_add) + then show ?thesis + using False R P1 p_residue_def p_residue_alt_def + by auto + qed + qed +qed + +lemma(in padic_integers) Zp_nat_inc_res: + fixes n::nat + shows "([n] \ \) k = n mod (p^k)" + using Zp_nat_inc_rep p_residue_def + by (simp add: p_residue_alt_def) + +lemma(in padic_integers) Zp_int_inc_rep: + fixes n::int + shows "[n] \ \ = (\ m. p_residue m n )" +proof(induction n) + case (nonneg n) + then show ?case using Zp_nat_inc_rep + by (simp add: add_pow_int_ge) +next + case (neg n) + show "\n. [(- int (Suc n))] \ \ = (\m. p_residue m (- int (Suc n)))" + proof + fix n + fix m + show "([(- int (Suc n))] \ \) m = p_residue m (- int (Suc n))" + proof- + have "[(- int (Suc n))] \ \ = \ ([(int (Suc n))] \ \)" + using a_inv_def abelian_group.a_group add_pow_def cring.axioms(1) domain_def + negative_zless_0 ring_def R.add.int_pow_neg R.one_closed by blast + then have "([(- int (Suc n))] \ \) m = (\ ([(int (Suc n))] \ \)) m" + by simp + have "\ \ carrier Zp" + using cring.cring_simprules(6) domain_def by blast + have "([(int (Suc n))] \ \) = ([(Suc n)] \ \)" + by (metis add_pow_def int_pow_int) + then have "([(int (Suc n))] \ \) \ carrier Zp" using Zp_nat_inc_closed + by simp + then have P0: "([(- int (Suc n))] \ \) m = \\<^bsub>Zp_res_ring m\<^esub> (([(int (Suc n))] \ \) m)" + using Zp_def prime + using \[(- int (Suc n))] \ \ = \ ([int (Suc n)] \ \)\ padic_integers.Zp_residue_a_inv(1) + padic_integers_axioms by auto + have "(([(int (Suc n))] \ \) m) = (p_residue m (Suc n))" + using Zp_nat_inc_rep by (simp add: add_pow_int_ge) + then have P1: "([(- int (Suc n))] \ \) m = \\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n))" + using P0 by auto + have "\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) = p_residue m (- int (Suc n))" + proof(cases "m=0") + case True + then have 0:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub>(p_residue 0 (Suc n))" + by blast + then have 1:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub> (p_residue 1 (Suc n))" + by (metis p_res_ring_0' residue_a_inv_closed) + then have 2:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub> 0" + by (metis p_res_ring_0' residue_a_inv_closed) + then have 3:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =0" + using residue_1_prop p_res_ring_0' residue_a_inv_closed by presburger + have 4: "p_residue m (- int (Suc n)) \ carrier (Zp_res_ring 0)" + using p_res_ring_0 True residue_1_zero p_residue_range' by blast + then show ?thesis + using "3" True residue_1_zero + by (simp add: p_res_ring_0') + next + case False + then have R: "residues (p^m)" + using padic_integers.p_residues padic_integers_axioms by blast + have "\\<^bsub>Zp_res_ring m\<^esub> p_residue m (int (Suc n)) = \\<^bsub>Zp_res_ring m\<^esub> (int (Suc n)) mod (p^m) " + using R residue_def residues.neg_cong residues.res_neg_eq p_residue_alt_def + by auto + then have "\\<^bsub>Zp_res_ring m\<^esub> p_residue m (int (Suc n)) = (-(int (Suc n))) mod (p^m)" + using R residues.res_neg_eq by auto + then show ?thesis + by (simp add: p_residue_alt_def) + qed + then show ?thesis + using P1 by linarith + qed + qed +qed + +lemma(in padic_integers) Zp_int_inc_res: + fixes n::int + shows "([n] \ \) k = n mod (p^k)" + using Zp_int_inc_rep p_residue_def + by (simp add: p_residue_alt_def) + +abbreviation(in padic_integers)(input) \

where +"\

\ [p] \ \" + +lemma(in padic_integers) p_natpow_prod: +"\

[^](n::nat) \ \

[^](m::nat) = \

[^](n + m)" + by (simp add: R.nat_pow_mult) + +lemma(in padic_integers) p_natintpow_prod: + assumes "(m::int) \ 0" + shows "\

[^](n::nat) \ \

[^]m = \

[^](n + m)" + using p_natpow_prod[of n "nat m"] assms int_pow_def[of Zp \

m] int_pow_def[of Zp \

"n + m"] + by (metis (no_types, lifting) int_nat_eq int_pow_int of_nat_add) + +lemma(in padic_integers) p_intnatpow_prod: + assumes "(n::int) \ 0" + shows "\

[^]n \ \

[^](m::nat) = \

[^](m + n)" + using p_natintpow_prod[of n m] assms mult_comm[of "\

[^]n" "\

[^]m"] + by simp + +lemma(in padic_integers) p_int_pow_prod: + assumes "(n::int) \ 0" + assumes "(m::int) \ 0" + shows "\

[^]n \ \

[^]m = \

[^](m + n)" +proof- + have "nat n + nat m = nat (n + m)" + using assms + by (simp add: nat_add_distrib) + then have "\

[^] (nat n + nat m) = \

[^](n + m)" + using assms + by (simp add: \nat n + nat m = nat (n + m)\) + then show ?thesis using assms p_natpow_prod[of "nat n" "nat m"] + by (smt pow_nat) +qed + +lemma(in padic_integers) p_natpow_prod_Suc: +"\

\ \

[^](m::nat) = \

[^](Suc m)" +"\

[^](m::nat) \ \

= \

[^](Suc m)" + using R.nat_pow_Suc2 R.nat_pow_Suc by auto + +lemma(in padic_integers) power_residue: + assumes "a \ carrier Zp" + assumes "k > 0" + shows "(a[^]\<^bsub>Zp\<^esub> (n::nat)) k = (a k)^n mod (p^k)" + apply(induction n) + using p_residues assms(2) residue_of_one(1) residues.one_cong apply auto[1] + by (simp add: assms(1) mod_mult_left_eq power_commutes residue_of_prod') + +(*************************************************************************************************) +(*************************************************************************************************) +(*****************************) section\The Valuation on $\mathbb{Z}_p$\ (**********************************) +(*************************************************************************************************) +(*************************************************************************************************) + + (**********************************************************************) + (**********************************************************************) + subsection\The Integer-Valued and Extended Integer-Valued Valuations\ + (**********************************************************************) + (**********************************************************************) +fun fromeint :: "eint \ int" where + "fromeint (eint x) = x" + +text\The extended-integer-valued $p$-adic valuation on $\mathbb{Z}_p$:\ + +definition(in padic_integers) val_Zp where +"val_Zp = (\x. (if (x = \) then (\::eint) else (eint (padic_val p x))))" + +text\We also define an integer-valued valuation on the nonzero elements of $\mathbb{Z}_p$, for simplified reasoning\ + +definition(in padic_integers) ord_Zp where +"ord_Zp = padic_val p" + +text\Ord of additive inverse\ + +lemma(in padic_integers) ord_Zp_of_a_inv: + assumes "a \ nonzero Zp" + shows "ord_Zp a = ord_Zp (\a)" + using ord_Zp_def Zp_def assms + padic_val_a_inv prime + by (simp add: domain.nonzero_memE(1) padic_int_is_domain) + +lemma(in padic_integers) val_Zp_of_a_inv: + assumes "a \ carrier Zp" + shows "val_Zp a = val_Zp (\a)" + using R.add.inv_eq_1_iff Zp_def assms padic_val_a_inv prime val_Zp_def by auto + +text\Ord-based criterion for being nonzero:\ + +lemma(in padic_integers) ord_of_nonzero: + assumes "x \carrier Zp" + assumes "ord_Zp x \0" + shows "x \ \" + "x \ nonzero Zp" +proof- + show "x \ \" + proof + assume "x = \" + then have "ord_Zp x = -1" + using ord_Zp_def padic_val_def Zp_def Zp_defs(2) by auto + then show False using assms(2) by auto + qed + then show "x \ nonzero Zp" + using nonzero_def assms(1) + by (simp add: nonzero_def) +qed + +lemma(in padic_integers) not_nonzero_Zp: + assumes "x \ carrier Zp" + assumes "x \ nonzero Zp" + shows "x = \" + using assms(1) assms(2) nonzero_def by fastforce + +lemma(in padic_integers) not_nonzero_Qp: + assumes "x \ carrier Q\<^sub>p" + assumes "x \ nonzero Q\<^sub>p" + shows "x = \\<^bsub>Q\<^sub>p\<^esub>" + using assms(1) assms(2) nonzero_def by force + +text\Relationship between val and ord\ + +lemma(in padic_integers) val_ord_Zp: + assumes "a \ \" + shows "val_Zp a = eint (ord_Zp a)" + by (simp add: assms ord_Zp_def val_Zp_def) + +lemma(in padic_integers) ord_pos: + assumes "x \ carrier Zp" + assumes "x \ \" + shows "ord_Zp x \ 0" +proof- + have "x \padic_zero p" + using Zp_def assms(2) Zp_defs(2) by auto + then have "ord_Zp x = int (LEAST k. x (Suc k) \ \\<^bsub>residue_ring (p^Suc k)\<^esub>)" + using ord_Zp_def padic_val_def by auto + then show ?thesis + by linarith +qed + +lemma(in padic_integers) val_pos: + assumes "x \ carrier Zp" + shows "val_Zp x \ 0" + unfolding val_Zp_def using assms + by (metis (full_types) eint_0 eint_ord_simps(1) eint_ord_simps(3) ord_Zp_def ord_pos) + +text\For passing between nat and int castings of ord\ + +lemma(in padic_integers) ord_nat: + assumes "x \ carrier Zp" + assumes "x \\" + shows "int (nat (ord_Zp x)) = ord_Zp x" + using ord_pos by (simp add: assms(1) assms(2)) + +lemma(in padic_integers) zero_below_ord: + assumes "x \ carrier Zp" + assumes "n \ ord_Zp x" + shows "x n = 0" +proof- + have "x n = \\<^bsub>residue_ring (p^n)\<^esub>" + using ord_Zp_def zero_below_val Zp_def assms(1) assms(2) prime padic_int_simps(5) + by auto + then show ?thesis using residue_ring_def + by simp +qed + +lemma(in padic_integers) zero_below_val_Zp: + assumes "x \ carrier Zp" + assumes "n \ val_Zp x" + shows "x n = 0" + by (metis assms(1) assms(2) eint_ord_simps(1) ord_Zp_def residue_of_zero(2) val_Zp_def zero_below_ord) + +lemma(in padic_integers) below_ord_zero: + assumes "x \ carrier Zp" + assumes "x (Suc n) \ 0" + shows "n \ ord_Zp x" +proof- + have 0: "x \ padic_set p" + using Zp_def assms(1) Zp_defs(3) + by auto + have 1: "x (Suc n) \ \\<^bsub>residue_ring (p^(Suc n))\<^esub>" + using residue_ring_def assms(2) by auto + have "of_nat n \ (padic_val p x )" + using 0 1 below_val_zero prime by auto + then show ?thesis using ord_Zp_def by auto +qed + +lemma(in padic_integers) below_val_Zp_zero: + assumes "x \ carrier Zp" + assumes "x (Suc n) \ 0" + shows "n \ val_Zp x" + by (metis Zp_def assms(1) assms(2) eint_ord_simps(1) padic_integers.below_ord_zero + padic_integers.residue_of_zero(2) padic_integers.val_ord_Zp padic_integers_axioms) + +lemma(in padic_integers) nonzero_imp_ex_nonzero_res: + assumes "x \ carrier Zp" + assumes "x \ \" + shows "\k. x (Suc k) \ 0" +proof- + have 0: "x 0 = 0" + using Zp_def assms(1) padic_int_simps(5) padic_set_zero_res prime by auto + have "\k. k > 0 \ x k \ 0" + apply(rule ccontr) using 0 Zp_defs unfolding padic_zero_def + by (metis assms(2) ext neq0_conv) + then show ?thesis + using not0_implies_Suc by blast +qed + +lemma(in padic_integers) ord_suc_nonzero: + assumes "x \ carrier Zp" + assumes "x \\" + assumes "ord_Zp x = n" + shows "x (Suc n) \ 0" +proof- + obtain k where k_def: "x (Suc k) \ 0" + using assms(1) nonzero_imp_ex_nonzero_res assms(2) by blast + then show ?thesis + using assms LeastI nonzero_imp_ex_nonzero_res unfolding ord_Zp_def padic_val_def + by (metis (mono_tags, lifting) Zp_defs(2) k_def of_nat_eq_iff padic_zero_def padic_zero_simp(1)) +qed + +lemma(in padic_integers) above_ord_nonzero: + assumes "x \ carrier Zp" + assumes "x \\" + assumes "n > ord_Zp x" + shows "x n \ 0" +proof- + have P0: "n \ (Suc (nat (ord_Zp x)))" + by (simp add: Suc_le_eq assms(1) assms(2) assms(3) nat_less_iff ord_pos) + then have P1: "p_residue (Suc (nat (ord_Zp x))) (x n) = x (Suc (nat (ord_Zp x)))" + using assms(1) p_residue_padic_int by blast + then have P2: "p_residue (Suc (nat (ord_Zp x))) (x n) \ 0" + using Zp_def assms(1) assms(2) ord_nat padic_integers.ord_suc_nonzero + padic_integers_axioms by auto + then show ?thesis + using P0 P1 assms(1) p_residue_padic_int[of x "(Suc (nat (ord_Zp x)))" n] p_residue_def + by (metis ord_Zp_def padic_int_simps(2) padic_integers.zero_rep padic_integers_axioms padic_zero_simp(2)) +qed + +lemma(in padic_integers) ord_Zp_geq: + assumes "x \ carrier Zp" + assumes "x n = 0" + assumes "x \\" + shows "ord_Zp x \ n" +proof(rule ccontr) + assume "\ int n \ ord_Zp x" + then show False using assms + using above_ord_nonzero by auto +qed + +lemma(in padic_integers) ord_equals: + assumes "x \ carrier Zp" + assumes "x (Suc n) \ 0" + assumes "x n = 0" + shows "ord_Zp x = n" + using assms(1) assms(2) assms(3) below_ord_zero ord_Zp_geq residue_of_zero(2) + by fastforce + +lemma(in padic_integers) ord_Zp_p: +"ord_Zp \

= (1::int)" +proof- + have "ord_Zp \

= int 1" + apply(rule ord_equals[of \

]) + using Zp_int_inc_res[of p] prime_gt_1_int prime by auto + then show ?thesis + by simp +qed + +lemma(in padic_integers) ord_Zp_one: +"ord_Zp \ = 0" +proof- + have "ord_Zp ([(1::int)]\\) = int 0" + apply(rule ord_equals) + using Zp_int_inc_res[of 1] prime_gt_1_int prime by auto + then show ?thesis + by simp +qed + +text\ord is multiplicative on nonzero elements of Zp\ + +lemma(in padic_integers) ord_Zp_mult: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + shows "(ord_Zp (x \\<^bsub>Zp\<^esub> y)) = (ord_Zp x) + (ord_Zp y)" + using val_prod[of p x y] prime assms Zp_defs Zp_def nonzero_memE(2) ord_Zp_def + nonzero_closed nonzero_memE(2) + by auto + +lemma(in padic_integers) ord_Zp_pow: + assumes "x \ nonzero Zp" + shows "ord_Zp (x[^](n::nat)) = n*(ord_Zp x)" +proof(induction n) + case 0 + have "x[^](0::nat) = \" + using assms(1) nonzero_def by simp + then show ?case + by (simp add: ord_Zp_one) +next + case (Suc n) + fix n + assume IH: "ord_Zp (x [^] n) = int n * ord_Zp x " + have N: "(x [^] n) \ nonzero Zp" + proof- + have "ord_Zp x \ 0" + using assms + by (simp add: nonzero_closed nonzero_memE(2) ord_pos) + then have "ord_Zp (x [^] n) \ 0" + using IH assms by simp + then have 0: "(x [^] n) \ \" + using ord_of_nonzero(1) by force + have 1: "(x [^] n) \ carrier Zp" + by (simp add: nonzero_closed assms) + then show ?thesis + using "0" not_nonzero_Zp by blast + qed + have "x[^](Suc n) = x \(x[^]n)" + using nonzero_closed assms R.nat_pow_Suc2 by blast + then have "ord_Zp (x[^](Suc n)) =(ord_Zp x) + ord_Zp (x[^]n)" + using N Zp_def assms padic_integers.ord_Zp_mult padic_integers_axioms by auto + then have "ord_Zp (x[^](Suc n)) =(ord_Zp x) +(int n * ord_Zp x)" + by (simp add: IH) + then have "ord_Zp (x[^](Suc n)) =(1*(ord_Zp x)) +(int n) * (ord_Zp x)" + by simp + then have "ord_Zp (x[^](Suc n)) =(1+ (int n)) * ord_Zp x" + by (simp add: comm_semiring_class.distrib) + then show "ord_Zp (x[^](Suc n)) = int (Suc n)*ord_Zp x" + by simp +qed + +lemma(in padic_integers) val_Zp_pow: + assumes "x \ nonzero Zp" + shows "val_Zp (x[^](n::nat)) = (n*(ord_Zp x))" + using Zp_def domain.nat_pow_nonzero[of Zp] domain_axioms nonzero_memE assms ord_Zp_def + padic_integers.ord_Zp_pow padic_integers_axioms val_Zp_def + nonzero_memE(2) + by fastforce + +lemma(in padic_integers) val_Zp_pow': + assumes "x \ nonzero Zp" + shows "val_Zp (x[^](n::nat)) = n*(val_Zp x)" + by (metis Zp_def assms not_nonzero_memI padic_integers.val_Zp_pow padic_integers.val_ord_Zp padic_integers_axioms times_eint_simps(1)) + +lemma(in padic_integers) ord_Zp_p_pow: +"ord_Zp (\

[^](n::nat)) = n" + using ord_Zp_pow ord_Zp_p Zp_def Zp_nat_inc_closed ord_of_nonzero(2) padic_integers_axioms int_inc_closed + Zp_int_inc_closed by auto + +lemma(in padic_integers) ord_Zp_p_int_pow: + assumes "n \0" + shows "ord_Zp (\

[^](n::int)) = n" + by (metis assms int_nat_eq int_pow_int ord_Zp_def ord_Zp_p_pow) + +lemma(in padic_integers) val_Zp_p: +"(val_Zp \

) = 1" + using Zp_def ord_Zp_def padic_val_def val_Zp_def ord_Zp_p Zp_defs(2) one_eint_def + by auto + +lemma(in padic_integers) val_Zp_p_pow: +"val_Zp (\

[^](n::nat)) = eint n" +proof- + have "(\

[^](n::nat)) \ \" + by (metis mult_zero_l n_not_Suc_n of_nat_eq_iff ord_Zp_def ord_Zp_p_pow p_natpow_prod_Suc(1)) + then show ?thesis + using ord_Zp_p_pow by (simp add: ord_Zp_def val_Zp_def) +qed + +lemma(in padic_integers) p_pow_res: + assumes "(n::nat) \ m" + shows "(\

[^]n) m = 0" + by (simp add: assms ord_Zp_p_pow zero_below_ord) + +lemma(in padic_integers) p_pow_factor: + assumes "(n::nat) \ m" + shows "(h \ (\

[^]n)) m = 0" "(h \ (\

[^]n)) m = \\<^bsub>Zp_res_ring n\<^esub>" + using assms p_pow_res p_res_ring_zero + by(auto simp: residue_of_zero Zp_residue_mult_zero(2)) + + (**********************************************************************) + (**********************************************************************) + subsection\The Ultrametric Inequality\ + (**********************************************************************) + (**********************************************************************) +text\Ultrametric inequality for ord\ + +lemma(in padic_integers) ord_Zp_ultrametric: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "x \ y \ nonzero Zp" + shows "ord_Zp (x \ y) \ min (ord_Zp x) (ord_Zp y)" + unfolding ord_Zp_def + using padic_val_ultrametric[of p x y] Zp_defs assms nonzero_memE Zp_def prime + nonzero_closed nonzero_memE(2) by auto + +text\Variants of the ultrametric inequality\ + +lemma (in padic_integers) ord_Zp_ultrametric_diff: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "x \ y " + shows "ord_Zp (x \ y) \ min (ord_Zp x) (ord_Zp y)" + using assms ord_Zp_ultrametric[of x "\ y"] + unfolding a_minus_def + by (metis (no_types, lifting) R.a_transpose_inv R.add.inv_closed R.add.m_closed R.l_neg nonzero_closed ord_Zp_of_a_inv ord_of_nonzero(2) ord_pos) + +lemma(in padic_integers) ord_Zp_not_equal_imp_notequal: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "ord_Zp x \ (ord_Zp y)" + shows "x \ y" "x \ y \\" "x \ y \\" + using assms + apply blast + using nonzero_closed assms(1) assms(2) assms(3) apply auto[1] + using nonzero_memE assms + using R.minus_equality nonzero_closed + Zp_def padic_integers.ord_Zp_of_a_inv + padic_integers_axioms by auto + +lemma(in padic_integers) ord_Zp_ultrametric_eq: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "ord_Zp x > (ord_Zp y)" + shows "ord_Zp (x \ y) = ord_Zp y" +proof- + have 0: "ord_Zp (x \ y) \ ord_Zp y" + using assms ord_Zp_not_equal_imp_notequal[of x y] + ord_Zp_ultrametric[of x y] nonzero_memE not_nonzero_Zp + nonzero_closed by force + have 1: "ord_Zp y \ min (ord_Zp(x \ y)) (ord_Zp x)" + proof- + have 0: "x \ y \ x" + using assms nonzero_memE + by (simp add: nonzero_closed nonzero_memE(2)) + have 1: "x \ y \ nonzero Zp" + using ord_Zp_not_equal_imp_notequal[of x y] + nonzero_closed assms(1) assms(2) assms(3) + not_nonzero_Zp by force + then show ?thesis + using 0 assms(1) assms(2) assms(3) ord_Zp_ultrametric_diff[of "x \ y" x] + by (simp add: R.minus_eq nonzero_closed R.r_neg1 add_comm) + qed + then show ?thesis + using 0 assms(3) + by linarith +qed + +lemma(in padic_integers) ord_Zp_ultrametric_eq': + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "ord_Zp x > (ord_Zp y)" + shows "ord_Zp (x \ y) = ord_Zp y" + using assms ord_Zp_ultrametric_eq[of x "\ y"] + unfolding a_minus_def + by (metis R.add.inv_closed R.add.inv_eq_1_iff nonzero_closed not_nonzero_Zp ord_Zp_of_a_inv) + +lemma(in padic_integers) ord_Zp_ultrametric_eq'': + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "ord_Zp x > (ord_Zp y)" + shows "ord_Zp (y \ x) = ord_Zp y" + by (metis R.add.inv_closed R.minus_eq + nonzero_closed Zp_def add_comm + assms(1) assms(2) assms(3) + ord_Zp_of_a_inv ord_of_nonzero(2) + ord_pos padic_integers.ord_Zp_ultrametric_eq padic_integers_axioms) + +lemma(in padic_integers) ord_Zp_not_equal_ord_plus_minus: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + assumes "ord_Zp x \ (ord_Zp y)" + shows "ord_Zp (x \ y) = ord_Zp (x \ y)" + apply(cases "ord_Zp x > ord_Zp y") + using assms + apply (simp add: ord_Zp_ultrametric_eq ord_Zp_ultrametric_eq') + using assms nonzero_memI + by (smt add_comm ord_Zp_ultrametric_eq ord_Zp_ultrametric_eq'') + +text\val is multiplicative on nonzero elements\ + +lemma(in padic_integers) val_Zp_mult0: + assumes "x \ carrier Zp" + assumes "x \\" + assumes "y \ carrier Zp" + assumes "y \\" + shows "(val_Zp (x \\<^bsub>Zp\<^esub> y)) = (val_Zp x) + (val_Zp y)" + apply(cases "x \\<^bsub>Zp\<^esub> y = \") + using assms(1) assms(2) assms(3) assms(4) integral_iff val_ord_Zp ord_Zp_mult nonzero_memI + apply (simp add: integral_iff) + using assms ord_Zp_mult[of x y] val_ord_Zp + by (simp add: nonzero_memI) + +text\val is multiplicative everywhere\ +lemma(in padic_integers) val_Zp_mult: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + shows "(val_Zp (x \\<^bsub>Zp\<^esub> y)) = (val_Zp x) + (val_Zp y)" + using assms(1) assms(2) integral_iff val_ord_Zp ord_Zp_mult nonzero_memI val_Zp_mult0 val_Zp_def + by simp + +lemma(in padic_integers) val_Zp_ultrametric0: + assumes "x \ carrier Zp" + assumes "x \\" + assumes "y \ carrier Zp" + assumes "y \\" + assumes "x \ y \ \" + shows "min (val_Zp x) (val_Zp y) \ val_Zp (x \ y) " + apply(cases "x \ y = \") + using assms apply blast + using assms ord_Zp_ultrametric[of x y] nonzero_memI val_ord_Zp[of x] val_ord_Zp[of y] val_ord_Zp[of "x \ y"] + by simp + +text\Unconstrained ultrametric inequality\ + +lemma(in padic_integers) val_Zp_ultrametric: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + shows " min (val_Zp x) (val_Zp y) \ val_Zp (x \ y)" + apply(cases "x = \") + apply (simp add: assms(2)) + apply(cases "y = \") + apply (simp add: assms(1)) + apply(cases "x \ y = \") + apply (simp add: val_Zp_def) + using assms val_Zp_ultrametric0[of x y] + by simp + +text\Variants of the ultrametric inequality\ + +lemma (in padic_integers) val_Zp_ultrametric_diff: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + shows "val_Zp (x \ y) \ min (val_Zp x) (val_Zp y)" + using assms val_Zp_ultrametric[of x "\y"] unfolding a_minus_def + by (metis R.add.inv_closed R.add.inv_eq_1_iff nonzero_memI ord_Zp_def ord_Zp_of_a_inv val_Zp_def) + +lemma(in padic_integers) val_Zp_not_equal_imp_notequal: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "val_Zp x \ val_Zp y" + shows "x \ y" "x \ y \\" "x \ y \\" + using assms(3) apply auto[1] + using assms(1) assms(2) assms(3) R.r_right_minus_eq apply blast + by (metis R.add.inv_eq_1_iff assms(1) assms(2) assms(3) R.minus_zero R.minus_equality + not_nonzero_Zp ord_Zp_def ord_Zp_of_a_inv val_ord_Zp) + +lemma(in padic_integers) val_Zp_ultrametric_eq: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "val_Zp x > val_Zp y" + shows "val_Zp (x \ y) = val_Zp y" + apply(cases "x \ \ \ y \ \ \ x \ y") + using assms ord_Zp_ultrametric_eq[of x y] val_ord_Zp nonzero_memE + using not_nonzero_memE val_Zp_not_equal_imp_notequal(3) apply force + unfolding val_Zp_def + using assms(2) assms(3) val_Zp_def by force + +lemma(in padic_integers) val_Zp_ultrametric_eq': + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "val_Zp x > (val_Zp y)" + shows "val_Zp (x \ y) = val_Zp y" + using assms val_Zp_ultrametric_eq[of x "\ y"] + unfolding a_minus_def + by (metis R.add.inv_closed R.r_neg val_Zp_not_equal_imp_notequal(3)) + +lemma(in padic_integers) val_Zp_ultrametric_eq'': + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "val_Zp x > (val_Zp y)" + shows "val_Zp (y \ x) = val_Zp y" +proof- + have 0: "y \ x = \ (x \ y)" + using assms(1,2) unfolding a_minus_def + by (simp add: R.add.m_comm R.minus_add) + have 1: "val_Zp (x \ y) = val_Zp y" + using assms val_Zp_ultrametric_eq' by blast + have 2: "val_Zp (x \ y) = val_Zp (y \ x)" + unfolding 0 unfolding a_minus_def + by(rule val_Zp_of_a_inv, rule R.ring_simprules, rule assms, rule R.ring_simprules, rule assms) + show ?thesis using 1 unfolding 2 by blast +qed + +lemma(in padic_integers) val_Zp_not_equal_ord_plus_minus: + assumes "x \ carrier Zp" + assumes "y \ carrier Zp" + assumes "val_Zp x \ (val_Zp y)" + shows "val_Zp (x \ y) = val_Zp (x \ y)" + by (metis R.add.inv_closed R.minus_eq R.r_neg R.r_zero add_comm assms(1) assms(2) assms(3) not_nonzero_Zp ord_Zp_def ord_Zp_not_equal_ord_plus_minus val_Zp_def val_Zp_not_equal_imp_notequal(3)) + + (**********************************************************************) + (**********************************************************************) + subsection\Units of $\mathbb{Z}_p$\ + (**********************************************************************) + (**********************************************************************) +text\Elements with valuation 0 in Zp are the units\ + +lemma(in padic_integers) val_Zp_0_criterion: + assumes "x \ carrier Zp" + assumes "x 1 \ 0" + shows "val_Zp x = 0" + unfolding val_Zp_def + using Zp_def assms(1) assms(2) ord_equals padic_set_zero_res prime + by (metis One_nat_def Zp_defs(3) of_nat_0 ord_Zp_def residue_of_zero(2) zero_eint_def) + +text\Units in Zp have val 0\ + +lemma(in padic_integers) unit_imp_val_Zp0: + assumes "x \ Units Zp" + shows "val_Zp x = 0" + apply(rule val_Zp_0_criterion) + apply (simp add: R.Units_closed assms) + using assms residue_of_prod[of x "inv x" 1] residue_of_one(2)[of 1] R.Units_r_inv[of x] + comm_monoid.UnitsI[of "R 1"] p_res_ring_1_field + by (metis le_eq_less_or_eq residue_of_prod residue_times_zero_r zero_le_one zero_neq_one) + +text\Elements in Zp with ord 0 are units\ + +lemma(in padic_integers) val_Zp0_imp_unit0: + assumes "val_Zp x = 0" + assumes "x \ carrier Zp" + fixes n::nat + shows "(x (Suc n)) \ Units (Zp_res_ring (Suc n))" + unfolding val_Zp_def +proof- + have p_res_ring: "residues (p^(Suc n))" + using p_residues by blast + have "\ n. coprime (x (Suc n)) p" + proof- + fix n + show "coprime (x (Suc n)) p" + proof- + have "\ \ coprime (x (Suc n)) p" + proof + assume "\ coprime (x (Suc n)) p" + then have "p dvd (x (Suc n))" using prime + by (meson coprime_commute prime_imp_coprime prime_nat_int_transfer) + then obtain k where "(x (Suc n)) = k*p" + by fastforce + then have S:"x (Suc n) mod p = 0" + by simp + have "x 1 = 0" + proof- + have "Suc n \ 1" + by simp + then have "x 1 = p_residue 1 (x (Suc n))" + using p_residue_padic_int assms(2) by presburger + then show ?thesis using S + by (simp add: p_residue_alt_def) + qed + have "x \\" + proof- + have "ord_Zp x \ ord_Zp \" + using Zp_def ord_Zp_def padic_val_def assms(1) ord_of_nonzero(1) R.zero_closed + Zp_defs(2) val_Zp_def + by auto + then show ?thesis + by blast + qed + then have "x 1 \ 0" + using assms(1) assms(2) ord_suc_nonzero + unfolding val_Zp_def + by (simp add: ord_Zp_def zero_eint_def) + then show False + using \x 1 = 0\ by blast + qed + then show ?thesis + by auto + qed + qed + then have "\ n. coprime (x (Suc n)) (p^(Suc n))" + by simp + then have "coprime (x (Suc n)) (p^(Suc n))" + by blast + then show ?thesis using assms residues.res_units_eq p_res_ring + by (metis (no_types, lifting) mod_pos_pos_trivial p_residue_ring_car_memE(1) + p_residue_ring_car_memE(2) residues.m_gt_one residues.mod_in_res_units residues_closed) +qed + +lemma(in padic_integers) val_Zp0_imp_unit0': + assumes "val_Zp x = 0" + assumes "x \ carrier Zp" + assumes "(n::nat) > 0" + shows "(x n) \ Units (Zp_res_ring n)" + using assms val_Zp0_imp_unit0 gr0_implies_Suc by blast + +lemma(in cring) ring_hom_Units_inv: + assumes "a \ Units R" + assumes "cring S" + assumes "h \ ring_hom R S" + shows "h (inv a) = inv\<^bsub>S\<^esub> h a" "h a \ Units S" +proof- + have 0:"h (inv a) \\<^bsub>S\<^esub> h a = \\<^bsub>S\<^esub>" + using assms Units_closed Units_inv_closed + by (metis (no_types, lifting) Units_l_inv ring_hom_mult ring_hom_one) + then show 1: "h (inv a) = inv\<^bsub>S\<^esub> h a" + by (metis Units_closed Units_inv_closed assms(1) assms(2) assms(3) comm_monoid.is_invI(1) cring_def ring_hom_closed) + show "h a \ Units S" + apply(rule comm_monoid.UnitsI[of S "inv\<^bsub>S\<^esub> h a"]) using 0 1 assms + using cring.axioms(2) apply blast + apply (metis "1" Units_inv_closed assms(1) assms(3) ring_hom_closed) + apply (meson Units_closed assms(1) assms(3) ring_hom_closed) + using "0" "1" by auto +qed + +lemma(in padic_integers) val_Zp_0_imp_unit: + assumes "val_Zp x = 0" + assumes "x \ carrier Zp" + shows "x \ Units Zp" +proof- + obtain y where y_def: " y= (\n. (if n=0 then 0 else (m_inv (Zp_res_ring n) (x n))))" + by blast + have 0: "\m. m > 0 \ y m = inv \<^bsub>Zp_res_ring m\<^esub> (x m)" + using y_def by auto + have 1: "\m. m > 0 \ inv\<^bsub>Zp_res_ring m\<^esub> (x m) \ carrier (Zp_res_ring m)" + proof- fix m::nat assume A: "m > 0" then show "inv\<^bsub>Zp_res_ring m\<^esub> (x m) \ carrier (Zp_res_ring m)" + using assms val_Zp0_imp_unit0' monoid.Units_inv_closed[of "Zp_res_ring m" "x m"] + by (smt One_nat_def Zp_def Zp_defs(2) cring.axioms(1) of_nat_0 ord_Zp_def + padic_integers.R_cring padic_integers.ord_suc_nonzero padic_integers.val_Zp_0_criterion padic_integers_axioms padic_val_def ring_def) + qed + have 2: "y \ padic_set p" + proof(rule padic_set_memI) + show 20: "\m. y m \ carrier (residue_ring (p ^ m))" + proof- fix m show "y m \ carrier (residue_ring (p ^ m))" + apply(cases "m = 0") + using y_def 0[of m] 1[of m] + by(auto simp: residue_ring_def y_def) + qed + show "\m n. m < n \ residue (p ^ m) (y n) = y m" + proof- fix m n::nat assume A: "m < n" + show "residue (p ^ m) (y n) = y m" + proof(cases "m = 0") + case True + then show ?thesis + by (simp add: residue_1_zero y_def) + next + case False + have hom: "residue (p ^ m) \ ring_hom (Zp_res_ring n) (Zp_res_ring m)" + using A False prime residue_hom_p by auto + have inv: "y n = inv\<^bsub>Zp_res_ring n\<^esub> x n" using A + by (simp add: False y_def) + have unit: "x n \ Units (Zp_res_ring n)" + using A False Zp_def assms(1) assms(2) val_Zp0_imp_unit0' prime + by (metis gr0I gr_implies_not0) + have F0: "residue (p ^ m) (x n) = x m" + using A Zp_defs(3) assms(2) padic_set_res_coherent prime by auto + have F1: "residue (p ^ m) (y n) = inv\<^bsub>Zp_res_ring m\<^esub> x m" + using F0 R_cring A hom inv unit cring.ring_hom_Units_inv[of "Zp_res_ring n" "x n" "Zp_res_ring m" "residue (p ^ m)"] + False + by auto + then show ?thesis + by (simp add: False y_def) + qed + qed + qed + have 3: "y \ x = \" + proof + fix m + show "(y \ x) m = \ m" + proof(cases "m=0") + case True + then have L: "(y \ x) m = 0" + using Zp_def "1" assms(2) Zp_residue_mult_zero(1) y_def + by auto + have R: "\ m = 0" + by (simp add: True cring.cring_simprules(6) domain.axioms(1) ord_Zp_one zero_below_ord) + then show ?thesis using L R by auto + next + case False + have P: "(y \ x) m = (y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m)" + using Zp_def residue_of_prod by auto + have "(y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m) = 1" + proof- + have "p^m > 1" + using False prime prime_gt_1_int by auto + then have "residues (p^m)" + using less_imp_of_nat_less residues.intro by fastforce + have "cring (residue_ring (p^m))" + using residues.cring \residues (p ^ m)\ + by blast + then have M: "monoid (residue_ring (p^m))" + using cring_def ring_def by blast + have U: "(x m) \ Units (residue_ring (p^m))" + using False Zp_def assms(1) assms(2) padic_integers.val_Zp0_imp_unit0' padic_integers_axioms by auto + have I: "y m = m_inv (residue_ring (p^m)) (x m)" + by (simp add: False y_def) + have "(y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m) = \\<^bsub>residue_ring (p^m)\<^esub>" + using M U I by (simp add: monoid.Units_l_inv) + then show ?thesis + using residue_ring_def by simp + qed + then show ?thesis + using P Zp_def False residue_of_one(2) by auto + qed + qed + have 4: "y \ carrier Zp" + using 2 Zp_defs by auto + show ?thesis + apply(rule R.UnitsI[of y]) + using assms 4 3 by auto +qed + +text\Definition of ord on a fraction is independent of the choice of representatives\ + +lemma(in padic_integers) ord_Zp_eq_frac: + assumes "a \ nonzero Zp" + assumes "b \ nonzero Zp" + assumes "c \ nonzero Zp" + assumes "d \ nonzero Zp" + assumes "a \ d = b \ c" + shows "(ord_Zp a) - (ord_Zp b) = (ord_Zp c) - (ord_Zp d)" +proof- + have "ord_Zp (a \ d) = ord_Zp (b \ c)" + using assms + by presburger + then have "(ord_Zp a) + (ord_Zp d) = (ord_Zp b) + (ord_Zp c)" + using assms(1) assms(2) assms(3) assms(4) ord_Zp_mult by metis + then show ?thesis + by linarith +qed + +lemma(in padic_integers) val_Zp_eq_frac_0: + assumes "a \ nonzero Zp" + assumes "b \ nonzero Zp" + assumes "c \ nonzero Zp" + assumes "d \ nonzero Zp" + assumes "a \ d = b \ c" + shows "(val_Zp a) - (val_Zp b) = (val_Zp c) - (val_Zp d)" +proof- + have 0:"(val_Zp a) - (val_Zp b) = (ord_Zp a) - (ord_Zp b)" + using assms nonzero_memE Zp_defs(2) ord_Zp_def val_Zp_def by auto + have 1: "(val_Zp c) - (val_Zp d) = (ord_Zp c) - (ord_Zp d)" + using assms nonzero_memE val_ord_Zp[of c] val_ord_Zp[of d] + by (simp add: nonzero_memE(2)) + then show ?thesis + using "0" assms(1) assms(2) assms(3) assms(4) assms(5) ord_Zp_eq_frac + by presburger +qed + +(*************************************************************************************************) +(*************************************************************************************************) +(*****************) section\Angular Component Maps on $\mathbb{Z}_p$\ (*********************) +(*************************************************************************************************) +(*************************************************************************************************) +text\The angular component map on $\mathbb{Z}_p$ is just the map which normalizes a point $x \in \mathbb{Z}_p$ by mapping it to a point with valuation $0$. It is explicitly defined as the mapping $x \mapsto p^{-ord (p)}*x$ for nonzero $x$, and $0 \mapsto 0$. By composing these maps with reductions mod $p^n$ we get maps which are equal to the standard residue maps on units of $\mathbb{Z}_p$, but in general unequal elsewhere. Both the angular component map and the angular component map mod $p^n$ are homomorpshims from the multiplicative group of units of $\mathbb{Z}_p$ to the multiplicative group of units of the residue rings, and play a key role in first-order model-theoretic formalizations of the $p$-adics (see, for example \cite{10.2307/2274477}, or \cite{Denef1986}). \ + + +lemma(in cring) int_nat_pow_rep: +"[(k::int)]\\ [^] (n::nat) = [(k^n)]\\" + apply(induction n) + by (auto simp: add.int_pow_pow add_pow_rdistr_int mult.commute) + +lemma(in padic_integers) p_pow_rep0: + fixes n::nat + shows "\

[^]n = [(p^n)]\\" + using R.int_nat_pow_rep by auto + +lemma(in padic_integers) p_pow_nonzero: + shows "(\

[^](n::nat)) \ carrier Zp" + "(\

[^](n::nat)) \ \" + apply simp + using Zp_def nat_pow_nonzero domain_axioms nonzero_memE int_inc_closed ord_Zp_p + padic_integers.ord_of_nonzero(2) padic_integers_axioms Zp_int_inc_closed + nonzero_memE(2) + by (metis ord_of_nonzero(2) zero_le_one) + +lemma(in padic_integers) p_pow_nonzero': + shows "(\

[^](n::nat)) \ nonzero Zp" + using nonzero_def p_pow_nonzero + by (simp add: nonzero_def) + +lemma(in padic_integers) p_pow_rep: + fixes n::nat + shows "(\

[^]n) k = (p^n) mod (p^k)" + by (simp add: R.int_nat_pow_rep Zp_int_inc_res) + +lemma(in padic_integers) p_pow_car: + assumes " (n::int)\ 0" + shows "(\

[^]n) \ carrier Zp" +proof- + have "(\

[^]n) = (\

[^](nat n))" + by (metis assms int_nat_eq int_pow_int) + then show ?thesis + by simp +qed + +lemma(in padic_integers) p_int_pow_nonzero: + assumes "(n::int) \0" + shows "(\

[^]n) \ nonzero Zp" + by (metis assms not_nonzero_Zp ord_Zp_p_int_pow ord_of_nonzero(1) p_pow_car) + +lemma(in padic_integers) p_nonzero: + shows "\

\ nonzero Zp" + using p_int_pow_nonzero[of 1] + by (simp add: ord_Zp_p ord_of_nonzero(2)) + +text\Every element of Zp is a unit times a power of p.\ + +lemma(in padic_integers) residue_factor_unique: + assumes "k>0" + assumes "x \ carrier Zp" + assumes "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" + shows "u = (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k)))" +proof- + obtain P where + P_def: "P = (\ u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k)))" + by simp + have 0: "P u" + using P_def assms(3) by blast + have 1: "\ v. P v \ v = u" + by (metis P_def assms(3) mult_cancel_right + not_prime_0 power_not_zero prime) + have "u = (THE u. P u)" + by (metis 0 1 the_equality) + then show ?thesis using P_def + by blast +qed + +lemma(in padic_integers) residue_factor_exists: + assumes "m = nat (ord_Zp x)" + assumes "k > 0" + assumes "x \ carrier Zp" + assumes "x \\" + obtains u where "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" +proof- + have X0: "(x (m+k)) \ carrier (Zp_res_ring (m+k)) " + using Zp_def assms(3) padic_set_res_closed residues_closed + by blast + then have X1: "(x (m+k)) \ 0" + using p_residues assms(2) residues.res_carrier_eq by simp + then have X2: "(x (m+k)) > 0" + using assms(1) assms(2) assms(3) assms(4) above_ord_nonzero + by (metis add.right_neutral add_cancel_right_right + add_gr_0 int_nat_eq less_add_same_cancel1 + less_imp_of_nat_less not_gr_zero of_nat_0_less_iff of_nat_add ord_pos) + have 0: "x m = 0" + using Zp_def assms(1) assms(3) zero_below_val ord_nat zero_below_ord[of x m] + assms(4) ord_Zp_def by auto + then have 1: "x (m +k) mod p^m = 0" + using assms(2) assms(3) p_residue_padic_int residue_def + by (simp add: p_residue_alt_def) + then have "\ u. u*(p^m) = (x (m+k))" + by auto + then obtain u where U0: " u*(p^m) = (x (m+k))" + by blast + have I: "(p^m) > 0 " + using prime + by (simp add: prime_gt_0_int) + then have U1: "(u* p^m) = (x (m+k))" + by (simp add: U0) + have U2: "u \ 0" + using I U1 X1 + by (metis U0 less_imp_triv mult.right_neutral mult_less_cancel_left + of_nat_zero_less_power_iff power.simps(1) times_int_code(1)) + have X3: "(x (m+k)) < p^(m+k)" + using assms(3) X0 p_residues assms(2) residues.res_carrier_eq by auto + have U3: "u < p^k" + proof(rule ccontr) + assume "\ u < (p ^ k)" + then have "(p^k) \ u" + by simp + then have " (p^k * p^m) \ u*(p^m)" + using I by simp + then have "p^(m + k) \ (x (m+k))" + by (simp add: U0 add.commute semiring_normalization_rules(26)) + then show False + using X3 by linarith + qed + then have "u \ carrier (Zp_res_ring k)" + using assms(2) p_residues residues.res_carrier_eq U3 U2 by auto + then show ?thesis using U1 that by blast +qed + +definition(in padic_integers) normalizer where +"normalizer m x = (\k. if (k=0) then 0 else (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m + k)) ) )" + +definition(in padic_integers) ac_Zp where +"ac_Zp x = normalizer (nat (ord_Zp x)) x" + +lemma(in padic_integers) ac_Zp_equation: + assumes "x \ carrier Zp" + assumes "x \\" + assumes "k > 0" + assumes "m = nat (ord_Zp x)" + shows "(ac_Zp x k) \ carrier (Zp_res_ring k) \ (ac_Zp x k)*(p^m) = (x (m+k))" +proof- + have K0: "k >0" + using assms nat_neq_iff by blast + have KM: "m+ k > m" + using assms(3) assms(4) by linarith + obtain u where U0: "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" + using assms(1) assms(2) assms(3) assms(4) residue_factor_exists by blast + have RHS: "ac_Zp x k = (THE u. u \ carrier (Zp_res_ring k) \ u*(p^m) = (x (m+k)))" + proof- + have K: "k \0" + by (simp add: K0) + have "ac_Zp x k = normalizer (nat (ord_Zp x)) x k" + using ac_Zp_def by presburger + then have "ac_Zp x k = normalizer m x k" + using assms by blast + then show "ac_Zp x k = (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m + k)) ) " + using K unfolding normalizer_def p_residue_def + by simp + qed + have LHS:"u = (THE u. u \ carrier (Zp_res_ring k) \ u*(p^m) = (x (m+k)))" + using assms U0 K0 assms(1) residue_factor_unique[of k x u m] by metis + then have "u = ac_Zp x k" + by (simp add: RHS) + then show ?thesis using U0 by auto +qed + +lemma(in padic_integers) ac_Zp_res: + assumes "m >k" + assumes "x \ carrier Zp" + assumes "x \\" + shows "p_residue k (ac_Zp x m) = (ac_Zp x k)" +proof(cases "k =0") + case True + then show ?thesis + unfolding ac_Zp_def normalizer_def + by (meson p_res_ring_0' p_residue_range') +next + case False + obtain n where n_def: "n = nat (ord_Zp x)" + by blast + have K0: "k >0" using False by simp + obtain uk where Uk0: "uk = (ac_Zp x k)" + by simp + obtain um where Um0: "um = (ac_Zp x m)" + by simp + have Uk1: "uk \ carrier (Zp_res_ring k) \ uk*(p^n) = (x (n+k))" + using K0 Uk0 ac_Zp_equation assms(2) assms(3) n_def by metis + have Um1: "um \ carrier (Zp_res_ring m) \ um*(p^n) = (x (n+m))" + using Uk1 Um0 ac_Zp_equation assms(1) assms(3) n_def assms(2) + by (metis neq0_conv not_less0) + have "um mod (p^k) = uk" + proof- + have "(x (n+m)) mod (p^(n + k)) = (x (n+k))" + using assms(1) assms(3) p_residue_padic_int p_residue_def n_def + by (simp add: assms(2) p_residue_alt_def) + then have "(p^(n + k)) dvd (x (n+m)) - (x (n+k))" + by (metis dvd_minus_mod) + then obtain d where "(x (n+m)) - (x (n+k)) = (p^(n+k))*d" + using dvd_def by blast + then have "((um*(p^n)) - (uk*(p^n))) = p^(n+k)*d" + using Uk1 Um1 by auto + then have "((um - uk)*(p^n)) = p^(n+k)*d" + by (simp add: left_diff_distrib) + then have "((um - uk)*(p^n)) = ((p^k)*d)*(p^n)" + by (simp add: power_add) + then have "(um - uk) = ((p^k)*d)" + using prime by auto + then have "um mod p^k = uk mod p^k" + by (simp add: mod_eq_dvd_iff) + then show ?thesis using Uk1 + by (metis mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) + qed + then show ?thesis + by (simp add: Uk0 Um0 p_residue_alt_def) +qed + +lemma(in padic_integers) ac_Zp_in_Zp: + assumes "x \ carrier Zp" + assumes "x \\" + shows "ac_Zp x \ carrier Zp" +proof- + have "ac_Zp x \ padic_set p" + proof(rule padic_set_memI) + show "\m. ac_Zp x m \ carrier (residue_ring (p ^ m))" + proof- + fix m + show "ac_Zp x m \ carrier (residue_ring (p ^ m))" + proof(cases "m = 0") + case True + then have "ac_Zp x m = 0" + unfolding ac_Zp_def normalizer_def by auto + then show ?thesis + by (simp add: True residue_ring_def) + next + case False + then have "m>0" + by blast + then show ?thesis + using ac_Zp_equation + by (metis assms(1) assms(2)) + qed + qed + show "\m n. m < n \ residue (p ^ m) (ac_Zp x n) = ac_Zp x m" + using ac_Zp_res + by (simp add: assms(1) assms(2) p_residue_def) + qed + then show ?thesis + by (simp add: Zp_defs(3)) +qed + +lemma(in padic_integers) ac_Zp_is_Unit: + assumes "x \ carrier Zp" + assumes "x \\" + shows "ac_Zp x \ Units Zp" +proof(rule val_Zp_0_imp_unit) + show "ac_Zp x \ carrier Zp" + by (simp add: ac_Zp_in_Zp assms(1) assms(2)) + obtain m where M: "m = nat (ord_Zp x)" + by blast + have AC1: "(ac_Zp x 1)*(p^m) = (x (m+1))" + using M ac_Zp_equation assms(1) assms(2) + by (metis One_nat_def lessI) + have "(x (m+1)) \0" + using M assms + by (metis Suc_eq_plus1 Suc_le_eq nat_int nat_mono nat_neq_iff ord_Zp_geq) + then have "(ac_Zp x 1) \ 0" + using AC1 by auto + then show "val_Zp (ac_Zp x) = 0" + using \ac_Zp x \ carrier Zp\ val_Zp_0_criterion + by blast +qed + +text\The typical defining equation for the angular component map.\ + +lemma(in padic_integers) ac_Zp_factors_x: + assumes "x \ carrier Zp" + assumes "x \\" + shows "x = (\

[^](nat (ord_Zp x))) \ (ac_Zp x)" "x = (\

[^](ord_Zp x)) \ (ac_Zp x)" +proof- + show "x = (\

[^](nat (ord_Zp x)))\ (ac_Zp x)" + proof + fix k + show "x k = ((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k" + proof(cases "k=0") + case True + then show ?thesis + using Zp_def Zp_defs(3) Zp_residue_mult_zero(1) ac_Zp_in_Zp + assms(1) assms(2) mult_comm padic_set_zero_res prime by auto + next + case False + show ?thesis + proof(cases "k \ ord_Zp x") + case True + have 0: "x k = 0" + using True assms(1) zero_below_ord by blast + have 1: "(\

[^](nat (ord_Zp x))) k = 0" + using True assms(1) assms(2) ord_Zp_p_pow ord_nat p_pow_nonzero(1) zero_below_ord + by presburger + have "((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k = (\

[^](nat (ord_Zp x))) k * (ac_Zp x) k mod p^k" + using Zp_def padic_mult_res residue_ring_def + using residue_of_prod' by blast + then have "((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k = 0" + by (simp add: "1") + then show ?thesis using 0 + by metis + next + case False + obtain n where N: "n = nat (ord_Zp x)" + by metis + obtain m where M0: "k = n + m" + using False N le_Suc_ex ord_Zp_def by fastforce + have M1: "m >0" + using M0 False N assms(1) assms(2) ord_nat + by (metis Nat.add_0_right gr0I le_refl less_eq_int_code(1) + nat_eq_iff2 neq0_conv of_nat_eq_0_iff of_nat_mono) + have E1: "(ac_Zp x m)*(p^n) = (x k)" + using M0 M1 N ac_Zp_equation assms(1) assms(2) by blast + have E2: "(ac_Zp x k)*(p^n) = (x (n + k))" + using M0 M1 N ac_Zp_equation assms(1) assms(2) add_gr_0 + by presburger + have E3: "((ac_Zp x k) mod (p^k))*((p^n) mod p^k) mod (p^k) = (x (n + k)) mod (p^k)" + by (metis E2 mod_mult_left_eq mod_mult_right_eq) + have E4: "((ac_Zp x k) mod (p^k))*(p^n) mod (p^k) = (x k)" + using E2 assms(1) le_add2 mod_mult_left_eq p_residue_padic_int p_residue_def + by (metis Zp_int_inc_rep Zp_int_inc_res) + + have E5: "(ac_Zp x k)*((p^n) mod p^k) mod (p^k) = (x k)" + using E2 assms(1) p_residue_padic_int p_residue_def by (metis E3 E4 mod_mult_left_eq) + have E6: "(ac_Zp x k) \\<^bsub>(Zp_res_ring k)\<^esub> ((p^n) mod p^k) = (x k)" + using E5 M0 M1 p_residues residues.res_mult_eq by auto + have E7: " ((p^n) mod p^k) \\<^bsub>(Zp_res_ring k)\<^esub>(ac_Zp x k) = (x k)" + by (simp add: E6 residue_mult_comm) + have E8: "((\

[^](nat (ord_Zp x))) k) \\<^bsub>(Zp_res_ring k)\<^esub> (ac_Zp x k) = (x k)" + using E7 N p_pow_rep + by metis + then show ?thesis + by (simp add: residue_of_prod) + qed + qed + qed + then show "x = (\

[^](ord_Zp x)) \ (ac_Zp x)" + by (metis assms(1) assms(2) int_pow_int ord_nat) +qed + +lemma(in padic_integers) ac_Zp_factors': + assumes "x \ nonzero Zp" + shows "x = [p] \ \ [^] ord_Zp x \ ac_Zp x" + using assms nonzero_memE + by (simp add: nonzero_closed nonzero_memE(2) ac_Zp_factors_x(2)) + +lemma(in padic_integers) ac_Zp_mult: + assumes "x \ nonzero Zp" + assumes "y \ nonzero Zp" + shows "ac_Zp (x \ y) = (ac_Zp x) \ (ac_Zp y)" +proof- + have P0: "x = (\

[^](nat (ord_Zp x))) \ (ac_Zp x)" + using nonzero_memE ac_Zp_factors_x assms(1) + by (simp add: nonzero_closed nonzero_memE(2)) + have P1: "y = (\

[^](nat (ord_Zp y))) \ (ac_Zp y)" + using nonzero_memE ac_Zp_factors_x assms(2) + by (simp add: nonzero_closed nonzero_memE(2)) + + have "x \ y = (\

[^](nat (ord_Zp (x \ y)))) \ (ac_Zp (x \ y))" + proof- + have "x \ y \ nonzero Zp" + by (simp add: assms(1) assms(2) nonzero_mult_closed) + then show ?thesis + using nonzero_closed nonzero_memE(2) Zp_def + padic_integers.ac_Zp_factors_x(1) padic_integers_axioms + by blast + qed + then have "x \ y = (\

[^](nat ((ord_Zp x) + (ord_Zp y)))) \ (ac_Zp (x \ y))" + using assms ord_Zp_mult[of x y] + by (simp add: Zp_def) + then have "x \ y = (\

[^]((nat (ord_Zp x)) + nat (ord_Zp y))) \ (ac_Zp (x \ y))" + using nonzero_closed nonzero_memE(2) assms(1) assms(2) + nat_add_distrib ord_pos by auto + then have "x \ y = (\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y))" + using p_natpow_prod + by metis + then have P2: "(\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y)) + = ((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) \ ((\

[^](nat (ord_Zp y))) \ (ac_Zp y))" + using P0 P1 + by metis + have "(\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y)) + = ((\

[^](nat (ord_Zp x))) \ ((\

[^](nat (ord_Zp y))) \ (ac_Zp x)) \ (ac_Zp y))" + by (metis P0 P1 Zp_def \x \ y = [p] \ \ [^] nat (ord_Zp x) \ [p] \ \ [^] nat (ord_Zp y) \ ac_Zp (x \ y)\ + mult_comm padic_integers.mult_assoc padic_integers_axioms) + then have "((\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y)))) \ (ac_Zp (x \ y)) + =((\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y)))) \ ((ac_Zp x) \ (ac_Zp y))" + using Zp_def mult_assoc by auto + then show ?thesis + by (metis (no_types, lifting) R.m_closed + \x \ y = [p] \ \ [^] nat (ord_Zp x) \ [p] \ \ [^] nat (ord_Zp y) \ ac_Zp (x \ y)\ + ac_Zp_in_Zp assms(1) assms(2) integral_iff m_lcancel + nonzero_closed nonzero_memE(2) p_pow_nonzero(1)) +qed + +lemma(in padic_integers) ac_Zp_one: +"ac_Zp \ = \" + by (metis R.one_closed Zp_def ac_Zp_factors_x(2) int_pow_0 ord_Zp_one padic_integers.ac_Zp_in_Zp padic_integers_axioms padic_one_id prime zero_not_one) + +lemma(in padic_integers) ac_Zp_inv: + assumes "x \ Units Zp" + shows "ac_Zp (inv\<^bsub>Zp\<^esub> x) = inv\<^bsub>Zp\<^esub> (ac_Zp x)" +proof- + have "x \ (inv\<^bsub>Zp\<^esub> x) = \" + using assms by simp + then have "(ac_Zp x) \ (ac_Zp (inv\<^bsub>Zp\<^esub> x)) = ac_Zp \" + using ac_Zp_mult[of x "(inv x)"] R.Units_nonzero + assms zero_not_one by auto + then show ?thesis + using R.invI(2)[of "(ac_Zp x)" "(ac_Zp (inv\<^bsub>Zp\<^esub> x))"] assms ac_Zp_in_Zp ac_Zp_one + by (metis (no_types, lifting) R.Units_closed R.Units_inv_closed + R.Units_r_inv integral_iff R.inv_unique ac_Zp_is_Unit) +qed + +lemma(in padic_integers) ac_Zp_of_Unit: + assumes "val_Zp x = 0" + assumes "x \ carrier Zp" + shows "ac_Zp x = x" + using assms unfolding val_Zp_def + by (metis R.one_closed Zp_def ac_Zp_factors_x(2) ac_Zp_one eint.inject infinity_ne_i0 mult_assoc + ord_Zp_def ord_Zp_one padic_integers.ac_Zp_in_Zp padic_integers_axioms padic_one_id prime zero_eint_def zero_not_one) + +lemma(in padic_integers) ac_Zp_p: +"(ac_Zp \

) = \" +proof- + have 0: "\

= \

[^] nat (ord_Zp \

) \ ac_Zp \

" + using ac_Zp_factors_x[of \

] ord_Zp_p ord_of_nonzero(1) + by auto + have 1: "\

[^] nat (ord_Zp \

) = \

" + by (metis One_nat_def nat_1 ord_Zp_p p_pow_rep0 power_one_right) + then have 2: "\

= \

\ ac_Zp \

" + using "0" by presburger + have "ac_Zp \

\ carrier Zp" + using ac_Zp_in_Zp[of \

] + by (simp add: ord_Zp_p ord_of_nonzero(1)) + then show ?thesis + by (metis "1" "2" m_lcancel R.one_closed R.r_one + Zp_int_inc_closed p_pow_nonzero(2)) +qed + +lemma(in padic_integers) ac_Zp_p_nat_pow: +"(ac_Zp (\

[^] (n::nat))) = \" + apply(induction n) + apply (simp add: ac_Zp_one) + using ac_Zp_mult ac_Zp_p int_nat_pow_rep nat_pow_Suc2 R.nat_pow_one + R.one_closed p_natpow_prod_Suc(1) p_nonzero p_pow_nonzero' p_pow_rep0 + by auto + +text\Facts for reasoning about integer powers in an arbitrary commutative monoid:\ + +lemma(in monoid) int_pow_add: + fixes n::int + fixes m::int + assumes "a \ Units G" + shows "a [^] (n + m) = (a [^] n) \ (a [^] m)" +proof- + have 0: "group (units_of G)" + by (simp add: units_group) + have 1: "a \ carrier (units_of G)" + by (simp add: assms units_of_carrier) + have "\n::int. a [^] n = a [^]\<^bsub>units_of G\<^esub> n" + proof- fix k::int show "a [^] k = a [^]\<^bsub>units_of G\<^esub> k" using 1 assms units_of_pow + by (metis Units_pow_closed int_pow_def nat_pow_def units_of_inv units_of_pow) + qed + have 2: "a [^]\<^bsub>units_of G\<^esub> (n + m) = (a [^]\<^bsub>units_of G\<^esub> n) \\<^bsub>units_of G\<^esub> (a [^]\<^bsub>units_of G\<^esub> m)" + by (simp add: "1" group.int_pow_mult units_group) + show ?thesis using 0 1 2 + by (simp add: \\n. a [^] n = a [^]\<^bsub>units_of G\<^esub> n\ units_of_mult) +qed + +lemma(in monoid) int_pow_unit_closed: + fixes n::int + assumes "a \ Units G" + shows "a[^] n \ Units G" + apply(cases "n \ 0") + using units_of_def[of G] units_group Units_inv_Units[of a] + Units_pow_closed[of "inv a"] Units_pow_closed[of a] + apply (metis assms pow_nat) + using units_of_def[of G] units_group Units_inv_Units[of a] + Units_pow_closed[of "inv a"] Units_pow_closed[of a] + by (simp add: assms int_pow_def nat_pow_def) + +lemma(in monoid) nat_pow_of_inv: + fixes n::nat + assumes "a \ Units G" + shows "inv a[^] n = inv (a[^] n)" + by (metis (no_types, hide_lams) Units_inv_Units Units_inv_closed Units_inv_inv Units_pow_closed + Units_r_inv assms inv_unique' nat_pow_closed nat_pow_one pow_mult_distrib) + +lemma(in monoid) int_pow_of_inv: + fixes n::int + assumes "a \ Units G" + shows "inv a[^] n = inv (a[^] n)" + apply(cases "n \0") + apply (metis assms nat_pow_of_inv pow_nat) + by (metis assms int_pow_def2 nat_pow_of_inv) + +lemma(in monoid) int_pow_inv: + fixes n::int + assumes "a \ Units G" + shows "a[^] -n = inv a[^] n" + apply(cases "n =0") + apply simp + apply(cases "n > 0") + using int_pow_def2[of G a "-n"] int_pow_of_inv + apply (simp add: assms) + using assms int_pow_def2[of G a "-n"] int_pow_def2[of G a "n"] int_pow_def2[of G "inv a"] + int_pow_of_inv[of a n] Units_inv_Units[of a] Units_inv_inv Units_pow_closed[of a] + by (metis linorder_not_less nat_0_iff nat_eq_iff2 nat_zero_as_int neg_0_less_iff_less) + +lemma(in monoid) int_pow_inv': + fixes n::int + assumes "a \ Units G" + shows "a[^] -n = inv (a[^] n)" + by (simp add: assms int_pow_inv int_pow_of_inv) + +lemma(in comm_monoid) inv_of_prod: + assumes "a \ Units G" + assumes "b \ Units G" + shows "inv (a \ b) = (inv a) \ (inv b)" + by (metis Units_m_closed assms(1) assms(2) comm_monoid.m_comm comm_monoid_axioms + group.inv_mult_group monoid.Units_inv_closed monoid_axioms units_group + units_of_carrier units_of_inv units_of_mult) + + +(*************************************************************************************************) +(*************************************************************************************************) +(************)section\Behaviour of $val\_Zp$ and $ord\_Zp$ on Natural Numbers and Integers\ (***********) +(*************************************************************************************************) +(*************************************************************************************************) + +text\If f and g have an equal residue at k, then they differ by a multiple of $p^k$.\ +lemma(in padic_integers) eq_residue_mod: + assumes "f \ carrier Zp" + assumes "g \ carrier Zp" + assumes "f k = g k" + shows "\h. h \ carrier Zp \ f = g \ (\

[^]k)\h" +proof(cases "f = g") + case True + then show ?thesis + using Zp_int_inc_zero' assms(1) by auto +next + case False + have "(f \ g) k = 0" + using assms + by (metis R.r_right_minus_eq residue_of_diff residue_of_zero(2)) + then have "ord_Zp (f \ g) \ k" + using False assms + by (simp add: ord_Zp_geq) + then obtain m::int where m_def: "m \ 0 \ ord_Zp (f \ g) = k + m" + using zle_iff_zadd by auto + have "f \ g = \

[^](k + m) \ ac_Zp (f \ g)" + using ac_Zp_factors_x(2)[of "f \ g"] False m_def assms(1) assms(2) by auto + then have 0: "f \ g = \

[^]k \ \

[^] m \ ac_Zp (f \ g)" + by (simp add: Zp_def m_def padic_integers.p_natintpow_prod padic_integers_axioms) + have "\

[^]k \ \

[^] m \ ac_Zp (f \ g) \ carrier Zp" + using assms "0" by auto + then have "f = g \ \

[^]k \ \

[^] m \ ac_Zp (f \ g)" + using 0 assms R.ring_simprules + by simp + then show ?thesis using mult_assoc + by (metis "0" False R.m_closed R.r_right_minus_eq \[p] \ \ [^] k \ [p] \ \ [^] m \ ac_Zp (f \ g) \ carrier Zp\ ac_Zp_in_Zp assms(1) assms(2) m_def p_pow_car) +qed + +lemma(in padic_integers) eq_residue_mod': + assumes "f \ carrier Zp" + assumes "g \ carrier Zp" + assumes "f k = g k" + obtains h where "h \ carrier Zp \ f = g \ (\

[^]k)\h" + using assms eq_residue_mod by meson + +text\Valuations of integers which do not divide $p$:\ + +lemma(in padic_integers) ord_Zp_p_nat_unit: + assumes "(n::nat) mod p \ 0" + shows "ord_Zp ([n]\\) = 0" + using ord_equals[of "[n]\\" "0::nat"] + by (simp add: Zp_nat_inc_res assms) + +lemma(in padic_integers) val_Zp_p_nat_unit: + assumes "(n::nat) mod p \ 0" + shows "val_Zp ([n]\\) = 0" + unfolding val_Zp_def + using assms ord_Zp_def ord_Zp_p_nat_unit ord_of_nonzero(1) zero_eint_def by auto + +lemma(in padic_integers) nat_unit: + assumes "(n::nat) mod p \ 0" + shows "([n]\\) \ Units Zp " + using Zp_nat_mult_closed val_Zp_p_nat_unit + by (simp add: assms val_Zp_0_imp_unit ord_Zp_p_nat_unit) + +lemma(in padic_integers) ord_Zp_p_int_unit: + assumes "(n::int) mod p \ 0" + shows "ord_Zp ([n]\\) = 0" + by (metis One_nat_def Zp_int_inc_closed Zp_int_inc_res assms mod_by_1 of_nat_0 ord_equals power_0 power_one_right) + +lemma(in padic_integers) val_Zp_p_int_unit: + assumes "(n::int) mod p \ 0" + shows "val_Zp ([n]\\) = 0" + unfolding val_Zp_def + using assms ord_Zp_def ord_Zp_p_int_unit ord_of_nonzero(1) zero_eint_def by auto + +lemma(in padic_integers) int_unit: + assumes "(n::int) mod p \ 0" + shows "([n]\\) \ Units Zp " + by (simp add: assms val_Zp_0_imp_unit val_Zp_p_int_unit) + +lemma(in padic_integers) int_decomp_ord: + assumes "n = l*(p^k)" + assumes "l mod p \ 0" + shows "ord_Zp ([n]\\) = k" +proof- + have 0: "n = l * (p^k)" + using assms(1) + by simp + then have "(l * (p ^ k) mod (p ^ (Suc k))) \ 0" + using Zp_def Zp_nat_inc_zero assms(2) p_nonzero nonzero_memE + padic_integers_axioms R.int_inc_zero nonzero_memE(2) by auto + then have 3: "(l * p ^ k) mod (p ^ (Suc k)) \ 0" + by presburger + show ?thesis + using "0" "3" Zp_int_inc_res ord_equals by auto +qed + +lemma(in padic_integers) int_decomp_val: + assumes "n = l*(p^k)" + assumes "l mod p \ 0" + shows "val_Zp ([n]\\) = k" + using Zp_def assms(1) assms(2) R.int_inc_closed ord_of_nonzero(1) int_decomp_ord padic_integers_axioms val_ord_Zp + by auto + +text\$\mathbb{Z}_p$ has characteristic zero:\ + +lemma(in padic_integers) Zp_char_0: + assumes "(n::int) > 0" + shows "[n]\\ \ \" +proof- + have "prime (nat p)" + using prime prime_nat_iff_prime + by blast + then obtain l0 k where 0: "nat n = l0*((nat p)^k) \ \ (nat p) dvd l0 " + using prime assms prime_power_canonical[of "nat p" "nat n"] + by auto + obtain l where l_def: "l = int l0" + by blast + have 1: "n = l*(p^k) \ \ p dvd l " + using 0 l_def + by (smt assms int_dvd_int_iff int_nat_eq of_nat_mult of_nat_power prime prime_gt_0_int) + show ?thesis + apply(cases "l = 1") + using 1 p_pow_nonzero(2) p_pow_rep0 apply auto[1] + using 1 by (simp add: dvd_eq_mod_eq_0 int_decomp_ord ord_of_nonzero(1)) +qed + +lemma(in padic_integers) Zp_char_0': + assumes "(n::nat) > 0" + shows "[n]\\ \ \" +proof- + have "[n]\\ = [(int n)]\\" + using assms + by (simp add: add_pow_def int_pow_int) + then show ?thesis using assms Zp_char_0[of "int n"] + by simp +qed + +lemma (in domain) not_eq_diff_nonzero: + assumes "a \ b" + assumes "a \ carrier R" + assumes "b \carrier R" + shows "a \ b \ nonzero R" + by (simp add: nonzero_def assms(1) assms(2) assms(3)) + +lemma (in domain) minus_a_inv: + assumes "a \ carrier R" + assumes "b \ carrier R" + shows "a \ b = \ (b \ a)" + by (simp add: add.inv_mult_group assms(1) assms(2) minus_eq) + +lemma(in ring) plus_diff_simp: + assumes "a \ carrier R" + assumes "b \ carrier R" + assumes "c \ carrier R" + assumes "X = a \ b" + assumes "Y = c \ a" + shows "X \ Y = c \ b" + using assms + unfolding a_minus_def + using ring_simprules + by (simp add: r_neg r_neg2) + +lemma (in padic_integers) Zp_residue_eq: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "val_Zp (a \ b) > k" + shows "(a k) = (b k)" +proof- + have 0: "(a \ b) k = a k \\<^bsub>Zp_res_ring k\<^esub> b k" + using assms + by (simp add: residue_of_diff) + have 1: "(a \ b) k = 0" + using assms zero_below_val + by (smt R.minus_closed Zp_def eint_ord_simps(2) padic_integers.p_res_ring_zero + padic_integers.residue_of_zero(1) padic_integers.val_ord_Zp padic_integers.zero_below_ord padic_integers_axioms) + show ?thesis + apply(cases "k = 0") + apply (metis assms(1) assms(2) p_res_ring_0' residues_closed) + using 0 1 assms p_residues R_cring Zp_def assms(1) assms(2) cring_def padic_set_res_closed + residues.res_zero_eq ring.r_right_minus_eq + by (metis Zp_defs(3) linorder_neqE_nat not_less0 p_res_ring_zero) +qed + +lemma (in padic_integers) Zp_residue_eq2: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "(a k) = (b k)" + assumes "a \ b" + shows "val_Zp (a \ b) \ k" +proof- + have "(a \ b) k = 0" + using assms residue_of_diff + by (simp add: Zp_def padic_integers.residue_of_diff' padic_integers_axioms) + then show ?thesis + using assms(1) assms(2) ord_Zp_def ord_Zp_geq val_Zp_def by auto +qed + +lemma (in padic_integers) equal_val_Zp: + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "c \ carrier Zp" + assumes "val_Zp a = val_Zp b" + assumes "val_Zp (c \ a) > val_Zp b" + shows "val_Zp c = val_Zp b" +proof- + have 0: "val_Zp c = val_Zp (c \ a \ a)" + using assms + by (simp add: R.l_neg R.minus_eq add_assoc) + have "val_Zp c \ min (val_Zp (c \ a)) (val_Zp a)" + using val_Zp_ultrametric[of "(c \ a)" a] assms(1) + assms(3) ord_Zp_ultrametric_eq'' + by (simp add: "0") + then have 1: "val_Zp c \ (val_Zp a)" + by (metis assms(4) assms(5) dual_order.order_iff_strict less_le_trans min_le_iff_disj) + have "val_Zp c = (val_Zp a)" + proof(rule ccontr) + assume A: "val_Zp c \ val_Zp a" + then have 0: "val_Zp c > val_Zp a" + using 1 A by auto + then have "val_Zp (c \ (a \ c)) \ min (val_Zp c) (val_Zp (a \ c))" + by (simp add: assms(1) assms(3) val_Zp_ultrametric) + then have 1: "val_Zp a \ min (val_Zp c) (val_Zp (a \ c))" + using assms(1) assms(3) assms(4) assms(5) val_Zp_ultrametric_eq' 0 by auto + have 2: "val_Zp (a \ c) > val_Zp a" + using "0" assms(1) assms(3) assms(4) assms(5) + val_Zp_ultrametric_eq' by auto + then have "val_Zp a > val_Zp a" + using 0 1 2 val_Zp_of_a_inv + by (metis assms(1) assms(3) assms(4) assms(5) val_Zp_ultrametric_eq') + then show False + by blast + qed + then show ?thesis + using assms(4) + by simp +qed + +lemma (in padic_integers) equal_val_Zp': + assumes "a \ carrier Zp" + assumes "b \ carrier Zp" + assumes "c \ carrier Zp" + assumes "val_Zp a = val_Zp b" + assumes "val_Zp c > val_Zp b" + shows "val_Zp (a \ c) = val_Zp b" +proof- + have 0: "val_Zp b < val_Zp (a \ c \ a)" + by (simp add: R.minus_eq nonzero_closed R.r_neg1 add_comm assms(1) assms(3) assms(5)) + have 1: "val_Zp a \ val_Zp (\ c)" + using assms(3) assms(4) assms(5) + by (metis eq_iff not_less val_Zp_of_a_inv) + then show ?thesis + by (meson "0" R.semiring_axioms assms(1) assms(2) assms(3) assms(4) equal_val_Zp semiring.semiring_simprules(1)) +qed + +lemma (in padic_integers) val_Zp_of_minus: + assumes "a \ carrier Zp" + shows "val_Zp a = val_Zp (\ a)" + using assms not_nonzero_Zp ord_Zp_def ord_Zp_of_a_inv val_Zp_def + by auto + +end diff --git a/thys/Padic_Ints/ROOT b/thys/Padic_Ints/ROOT new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/ROOT @@ -0,0 +1,21 @@ +chapter AFP + +session Padic_Ints (AFP) = "HOL-Algebra" + + options [timeout = 1200] + sessions + "HOL-Number_Theory" + theories + Function_Ring + Cring_Poly + Supplementary_Ring_Facts + Extended_Int + Padic_Construction + Padic_Integers + Padic_Int_Topology + Padic_Int_Polynomials + Hensels_Lemma + Zp_Compact + document_files + "root.tex" + "root.bib" + diff --git a/thys/Padic_Ints/Supplementary_Ring_Facts.thy b/thys/Padic_Ints/Supplementary_Ring_Facts.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Supplementary_Ring_Facts.thy @@ -0,0 +1,160 @@ +theory Supplementary_Ring_Facts +imports "HOL-Algebra.Ring" + "HOL-Algebra.UnivPoly" + "HOL-Algebra.Subrings" + +begin + +section\Supplementary Ring Facts\ + +text\The nonzero elements of a ring.\ + +definition nonzero :: "('a, 'b) ring_scheme \ 'a set" where +"nonzero R = {a \ carrier R. a \ \\<^bsub>R\<^esub>}" + + +lemma zero_not_in_nonzero: +"\\<^bsub>R\<^esub> \ nonzero R" + unfolding nonzero_def by blast + +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) nonzero_closed: + assumes "a \ nonzero R" + shows "a \ carrier R" + using assms + by (simp add: nonzero_def) + +lemma(in domain) nonzero_mult_in_car: + assumes "a \ nonzero R" + assumes "b \ nonzero R" + shows "a \ b \ 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 \ nonzero R" + apply(rule nonzero_memI) + using assms nonzero_memE apply blast + using assms nonzero_memE + by (simp add: integral_iff) + +lemma(in domain) nonzero_one_closed: +"\ \ nonzero R" + by (simp add: nonzero_def) + +lemma(in domain) one_nonzero: +"\ \ nonzero R" + by (simp add: nonzero_one_closed) + +lemma(in domain) 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) + +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 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) + + +end diff --git a/thys/Padic_Ints/Zp_Compact.thy b/thys/Padic_Ints/Zp_Compact.thy new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/Zp_Compact.thy @@ -0,0 +1,767 @@ +theory Zp_Compact +imports Padic_Int_Topology +begin + +context padic_integers +begin + +lemma res_ring_car: +"carrier (Zp_res_ring k) = {0..p ^ k - 1}" + unfolding residue_ring_def by simp + +text\The refinement of a sequence by a function $nat \Rightarrow nat$\ +definition take_subseq :: "(nat \ 'a) \ (nat \ nat) \ (nat \ 'a)" where +"take_subseq s f = (\k. s (f k))" + +text\Predicate for increasing function on the natural numbers\ +definition is_increasing :: "(nat \ nat) \ bool" where +"is_increasing f = (\ n m::nat. n>m \ (f n) > (f m))" + +text\Elimination and introduction lemma for increasing functions\ +lemma is_increasingI: + assumes "\ n m::nat. n>m \ (f n) > (f m)" + shows "is_increasing f" + unfolding is_increasing_def + using assms + by blast + +lemma is_increasingE: + assumes "is_increasing f" + assumes " n> m" + shows "f n > f m" + using assms + unfolding is_increasing_def + by blast + +text\The subsequence predicate\ +definition is_subseq_of :: "(nat \ 'a) \ (nat \ 'a) \ bool" where +"is_subseq_of s s' = (\(f::nat \ nat). is_increasing f \ s' = take_subseq s f)" + +text\Subsequence introduction lemma\ +lemma is_subseqI: + assumes "is_increasing f" + assumes "s' = take_subseq s f" + shows "is_subseq_of s s'" + using assms + unfolding is_subseq_of_def + by auto + +lemma is_subseq_ind: + assumes "is_subseq_of s s'" + shows "\ l. s' k = s l" + using assms + unfolding is_subseq_of_def take_subseq_def by blast + +lemma is_subseq_closed: + assumes "s \ closed_seqs Zp" + assumes "is_subseq_of s s'" + shows "s' \ closed_seqs Zp" + apply(rule closed_seqs_memI) + using is_subseq_ind assms closed_seqs_memE + by metis + +text\Given a sequence and a predicate, returns the function from nat to nat which represents +the increasing sequences of indices n on which P (s n) holds.\ + +primrec seq_filter :: "(nat \'a) \ ('a \ bool) \ nat \ nat" where +"seq_filter s P (0::nat) = (LEAST k::nat. P (s k))"| +"seq_filter s P (Suc n) = (LEAST k:: nat. (P (s k)) \ k > (seq_filter s P n))" + +lemma seq_filter_pre_increasing: + assumes "\n::nat. \m. m > n \ P (s m)" + shows "seq_filter s P n < seq_filter s P (Suc n)" + apply(auto) +proof(induction n) + case 0 + have "\k. P (s k)" using assms(1) by blast + then have "\k::nat. (LEAST k::nat. (P (s k))) \ 0" + by blast + obtain k where "(LEAST k::nat. (P (s k))) = k" by simp + have "\l. l = (LEAST l::nat. (P (s l) \ l > k))" + by simp + thus ?case + by (metis (no_types, lifting) LeastI assms) +next + case (Suc n) + then show ?case + by (metis (no_types, lifting) LeastI assms) +qed + +lemma seq_filter_increasing: + assumes "\n::nat. \m. m > n \ P (s m)" + shows "is_increasing (seq_filter s P)" + by (metis assms seq_filter_pre_increasing is_increasingI lift_Suc_mono_less) + +definition filtered_seq :: "(nat \ 'a) \ ('a \ bool) \ (nat \ 'a)" where +"filtered_seq s P = take_subseq s (seq_filter s P)" + +lemma filter_exist: + assumes "s \ closed_seqs Zp" + assumes "\n::nat. \m. m > n \ P (s m)" + shows "\m. n\m \ P (s (seq_filter s P n))" +proof(induct n) + case 0 + then show ?case + using LeastI assms(2) by force +next + case (Suc n) + then show ?case + by (smt LeastI assms(2) seq_filter.simps(2)) +qed + +text\In a filtered sequence, every element satisfies the filtering predicate \ + +lemma fil_seq_pred: + assumes "s \ closed_seqs Zp" + assumes "s' = filtered_seq s P" + assumes "\n::nat. \m. m > n \ P (s m)" + shows "\m::nat. P (s' m)" +proof- + have "\k. P (s k)" using assms(3) + by blast + fix m + obtain k where kdef: "k = seq_filter s P m" by auto + have "\k. P (s k)" + using assms(3) by auto + then have "P (s k)" + by (metis (full_types) assms(1) assms(3) kdef le_refl less_imp_triv not_less_eq filter_exist ) + then have "s' m = s k" + by (simp add: assms(2) filtered_seq_def kdef take_subseq_def) + hence "P (s' m)" + by (simp add: \P (s k)\) + thus "\m. P (s' m)" using assms(2) assms(3) dual_order.strict_trans filter_exist filtered_seq_def + lessI less_Suc_eq_le take_subseq_def + by (metis (mono_tags, hide_lams) assms(1)) +qed + +definition kth_res_equals :: "nat \ int \ (padic_int \ bool)" where +"kth_res_equals k n a = (a k = n)" + +(*The characteristic function of the underlying set of a sequence*) +definition indicator:: "(nat \ 'a) \ ('a \ bool)" where +"indicator s a = (\n::nat. s n = a)" + + +text\Choice function for a subsequence with constant kth residue. Could be made constructive by +choosing the LEAST n if we wanted.\ + +definition const_res_subseq :: "nat \ padic_int_seq \ padic_int_seq" where +"const_res_subseq k s = (SOME s'::(padic_int_seq). (\ n. is_subseq_of s s' \ s' + = (filtered_seq s (kth_res_equals k n)) \ (\m. s' m k = n)))" + +text\The constant kth residue value for the sequence obtained by the previous function\ + +definition const_res :: "nat \ padic_int_seq \ int" where +"const_res k s = (THE n. (\ m. (const_res_subseq k s) m k = n))" + +definition maps_to_n:: "int \ (nat \ int) \ bool" where +"maps_to_n n f = (\(k::nat). f k \ {0..n})" + +definition drop_res :: "int \ (nat \ int) \ (nat \ int)" where +"drop_res k f n = (if (f n) = k then 0 else f n)" + +lemma maps_to_nE: + assumes "maps_to_n n f" + shows "(f k) \ {0..n}" + using assms + unfolding maps_to_n_def + by blast + +lemma maps_to_nI: + assumes "\n. f n \{0 .. k}" + shows "maps_to_n k f" + using assms maps_to_n_def by auto + + +lemma maps_to_n_drop_res: + assumes "maps_to_n (Suc n) f" + shows "maps_to_n n (drop_res (Suc n) f)" +proof- + fix k + have "drop_res (Suc n) f k \ {0..n}" + proof(cases "f k = Suc n") + case True + then have "drop_res (Suc n) f k = 0" + unfolding drop_res_def by auto + then show ?thesis + using assms local.drop_res_def maps_to_n_def by auto + next + case False + then show ?thesis + using assms atLeast0_atMost_Suc maps_to_n_def drop_res_def + by auto + qed + then have "\k. drop_res (Suc n) f k \ {0..n}" + using assms local.drop_res_def maps_to_n_def by auto + then show "maps_to_n n (drop_res (Suc n) f)" using maps_to_nI + using maps_to_n_def by blast +qed + +lemma drop_res_eq_f: + assumes "maps_to_n (Suc n) f" + assumes "\ (\m. \n. n>m \ (f n = (Suc k)))" + shows "\N. \n. n>N \ f n = drop_res (Suc k) f n" +proof- + have "\m. \n. n \ m \ (f n) \ (Suc k)" + using assms + by (meson Suc_le_eq nat_le_linear) + then have "\m. \n. n \ m \ (f n) = drop_res (Suc k) f n" + using drop_res_def by auto + then show ?thesis + by (meson less_Suc_eq_le order.asym) +qed + +lemma maps_to_n_infinite_seq: + shows "\f. maps_to_n (k::nat) f \ \l::int. \m. \n. n>m \ (f n = l)" +proof(induction k) + case 0 + then have "\n. f n \ {0}" + using maps_to_nE[of 0 f] by auto + then show " \l. \m. \n. m < n \ f n = l" + by blast +next + case (Suc k) + assume IH: "\f. maps_to_n k f \ \l. \m. \n. m < n \ f n = l" + fix f + assume A: "maps_to_n (Suc k) f" + show "\l. \m. \n. n>m \ (f n = l)" + proof(cases " \m. \n. n>m \ (f n = (Suc k))") + case True + then show ?thesis by blast + next + case False + then obtain N where N_def: "\n. n>N \ f n = drop_res (Suc k) f n" + using drop_res_eq_f drop_res_def + by fastforce + have " maps_to_n k (drop_res (Suc k) f) " + using A maps_to_n_drop_res by blast + then have " \l. \m. \n. m < n \ (drop_res (Suc k) f) n = l" + using IH by blast + then obtain l where l_def: "\m. \n. m < n \ (drop_res (Suc k) f) n = l" + by blast + have "\m. \n. n>m \ (f n = l)" + apply auto + proof- + fix m + show "\n>m. f n = l" + proof- + obtain n where N'_def: "(max m N) < n \ (drop_res (Suc k) f) n = l" + using l_def by blast + have "f n = (drop_res (Suc k) f) n" + using N'_def N_def + by simp + then show ?thesis + using N'_def by auto + qed + qed + then show ?thesis + by blast + qed +qed + +lemma int_nat_p_pow_minus: +"int (nat (p ^ k - 1)) = p ^ k - 1" + by (simp add: prime prime_gt_0_int) + +lemma maps_to_n_infinite_seq_res_ring: +"\f. f \ (UNIV::nat set) \ carrier (Zp_res_ring k) \ \l. \m. \n. n>m \ (f n = l)" +apply(rule maps_to_n_infinite_seq[of "nat (p^k - 1)"]) + unfolding maps_to_n_def res_ring_car int_nat_p_pow_minus by blast + +definition index_to_residue :: "padic_int_seq \ nat \ nat \ int" where +"index_to_residue s k m = ((s m) k)" + +lemma seq_maps_to_n: + assumes "s \ closed_seqs Zp" + shows "(index_to_residue s k) \ UNIV \ carrier (Zp_res_ring k)" +proof- + have A1: "\m. (s m) \ carrier Zp" + using assms closed_seqs_memE by auto + have A2: "\m. (s m k) \ carrier (Zp_res_ring k)" + using assms by (simp add: A1) + have "\m. index_to_residue s k m = s m k" + using index_to_residue_def + by auto + thus "index_to_residue s k \ UNIV \ carrier (residue_ring (p ^ k))" + using A2 by simp +qed + +lemma seq_pr_inc: + assumes "s \ closed_seqs Zp" + shows "\l. \m. \n > m. (kth_res_equals k l) (s n)" +proof- + fix k l m + have 0: "(kth_res_equals k l) (s m) \ (s m) k = l" + by (simp add: kth_res_equals_def) + have 1: "\k m. s m k = index_to_residue s k m" + by (simp add: index_to_residue_def) + have 2: "(index_to_residue s k) \ UNIV \ carrier (Zp_res_ring k)" + using seq_maps_to_n assms by blast + have 3: "\m. s m k \ carrier (Zp_res_ring k)" + proof- + fix m have 30: "s m k = index_to_residue s k m" + using 1 by blast + show " s m k \ carrier (Zp_res_ring k)" + unfolding 30 using 2 by blast + qed + obtain j where j_def: "j = nat (p^k - 1)" + by blast + have j_to_int: "int j = p^k - 1" + using j_def + by (simp add: prime prime_gt_0_int) + have "\l. \m. \n. n > m \ (index_to_residue s k n = l)" + by(rule maps_to_n_infinite_seq_res_ring[of _ k], rule seq_maps_to_n, rule assms) + hence "\l. \m. \n. n > m \ (s n k = l)" + by (simp add: index_to_residue_def) + thus "\l. \m. \n > m. (kth_res_equals k l) (s n)" + using kth_res_equals_def by auto +qed + +lemma kth_res_equals_subseq: + assumes "s \ closed_seqs Zp" + shows "\n. is_subseq_of s (filtered_seq s (kth_res_equals k n)) \ (\m. (filtered_seq s (kth_res_equals k n)) m k = n)" +proof- + obtain l where l_def: " \ m. \n > m. (kth_res_equals k l) (s n)" + using assms seq_pr_inc by blast + have 0: "is_subseq_of s (filtered_seq s (kth_res_equals k l))" + unfolding filtered_seq_def + apply(rule is_subseqI[of "seq_filter s (kth_res_equals k l)"]) + apply(rule seq_filter_increasing, rule l_def) + by blast + have 1: " (\m. (filtered_seq s (kth_res_equals k l)) m k = l)" + using l_def + by (meson assms kth_res_equals_def fil_seq_pred padic_integers_axioms) + show ?thesis using 0 1 by blast +qed + +lemma const_res_subseq_prop_0: + assumes "s \ closed_seqs Zp" + shows "\l. (((const_res_subseq k s) = filtered_seq s (kth_res_equals k l)) \ (is_subseq_of s (const_res_subseq k s)) \ (\m.(const_res_subseq k s) m k = l))" +proof- + have " \n. (is_subseq_of s (filtered_seq s (kth_res_equals k n)) \ (\m. (filtered_seq s (kth_res_equals k n)) m k = n))" + by (simp add: kth_res_equals_subseq assms) + then have "\s'. (\n. (is_subseq_of s s') \ (s' = filtered_seq s (kth_res_equals k n)) \ (\m. s' m k = n))" + by blast + then show ?thesis + using const_res_subseq_def[of k s] const_res_subseq_def someI_ex + by (smt const_res_subseq_def someI_ex) +qed + +lemma const_res_subseq_prop_1: + assumes "s \ closed_seqs Zp" + shows "(\m.(const_res_subseq k s) m k = (const_res k s) )" + using const_res_subseq_prop_0[of s] const_res_def[of k s] + by (smt assms const_res_subseq_def const_res_def the_equality) + +lemma const_res_subseq: + assumes "s \ closed_seqs Zp" + shows "is_subseq_of s (const_res_subseq k s)" + using assms const_res_subseq_prop_0[of s k] by blast + +lemma const_res_range: + assumes "s \ closed_seqs Zp" + assumes "k > 0" + shows "const_res k s \ carrier (Zp_res_ring k)" +proof- + have 0: "(const_res_subseq k s) 0 \ carrier Zp" + using const_res_subseq[of s k] is_subseq_closed[of s "const_res_subseq k s"] + assms(1) closed_seqs_memE by blast + have 1: "(const_res_subseq k s) 0 k \ carrier (Zp_res_ring k)" + using 0 by simp + then show ?thesis + using assms const_res_subseq_prop_1[of s k] + by (simp add: \s \ closed_seqs Zp\) +qed + +fun res_seq ::"padic_int_seq \ nat \ padic_int_seq" where +"res_seq s 0 = s"| +"res_seq s (Suc k) = const_res_subseq (Suc k) (res_seq s k)" + +lemma res_seq_res: + assumes "s \ closed_seqs Zp" + shows "(res_seq s k) \ closed_seqs Zp" + apply(induction k) + apply (simp add: assms) + by (simp add: const_res_subseq is_subseq_closed) + +lemma res_seq_res': + assumes "s \ closed_seqs Zp" + shows "\n. res_seq s (Suc k) n (Suc k) = const_res (Suc k) (res_seq s k)" + using assms res_seq_res[of s k] const_res_subseq_prop_1[of "(res_seq s k)" "Suc k" ] + by simp + +lemma res_seq_subseq: + assumes "s \ closed_seqs Zp" + shows "is_subseq_of (res_seq s k) (res_seq s (Suc k))" + by (metis assms const_res_subseq_prop_0 res_seq_res + res_seq.simps(2)) + +lemma is_increasing_id: +"is_increasing (\ n. n)" + by (simp add: is_increasingI) + +lemma is_increasing_comp: + assumes "is_increasing f" + assumes "is_increasing g" + shows "is_increasing (f \ g)" + using assms(1) assms(2) is_increasing_def + by auto + +lemma is_increasing_imp_geq_id[simp]: + assumes "is_increasing f" + shows "f n \n" + apply(induction n) + apply simp + by (metis (mono_tags, lifting) assms is_increasing_def + leD lessI not_less_eq_eq order_less_le_subst2) + +lemma is_subseq_ofE: + assumes "s \ closed_seqs Zp" + assumes "is_subseq_of s s'" + shows "\k. k \ n \ s' n = s k" +proof- + obtain f where "is_increasing f \ s' = take_subseq s f" + using assms(2) is_subseq_of_def by blast + then have " f n \ n \ s' n = s (f n)" + unfolding take_subseq_def + by simp + then show ?thesis by blast +qed + + +lemma is_subseq_of_id: + assumes "s \ closed_seqs Zp" + shows "is_subseq_of s s" +proof- + have "s = take_subseq s (\n. n)" + unfolding take_subseq_def + by auto + then show ?thesis using is_increasing_id + using is_subseqI + by blast +qed + +lemma is_subseq_of_trans: + assumes "s \ closed_seqs Zp" + assumes "is_subseq_of s s'" + assumes "is_subseq_of s' s''" + shows "is_subseq_of s s''" +proof- + obtain f where f_def: "is_increasing f \ s' = take_subseq s f" + using assms(2) is_subseq_of_def + by blast + obtain g where g_def: "is_increasing g \ s'' = take_subseq s' g" + using assms(3) is_subseq_of_def + by blast + have "s'' = take_subseq s (f \ g)" + proof + fix x + show "s'' x = take_subseq s (f \ g) x" + using f_def g_def unfolding take_subseq_def + by auto + qed + then show ?thesis + using f_def g_def is_increasing_comp is_subseq_of_def + by blast +qed + +lemma res_seq_subseq': + assumes "s \ closed_seqs Zp" + shows "is_subseq_of s (res_seq s k)" +proof(induction k) + case 0 + then show ?case using is_subseq_of_id + by (simp add: assms) +next + case (Suc k) + fix k + assume "is_subseq_of s (res_seq s k)" + then show "is_subseq_of s (res_seq s (Suc k)) " + using assms is_subseq_of_trans res_seq_subseq + by blast +qed + +lemma res_seq_subseq'': + assumes "s \ closed_seqs Zp" + shows "is_subseq_of (res_seq s n) (res_seq s (n + k))" + apply(induction k) + apply (simp add: assms is_subseq_of_id res_seq_res) + using add_Suc_right assms is_subseq_of_trans res_seq_res res_seq_subseq by presburger +(**) + +definition acc_point :: "padic_int_seq \ padic_int" where +"acc_point s k = (if (k = 0) then (0::int) else ((res_seq s k) 0 k))" + +lemma res_seq_res_1: + assumes "s \ closed_seqs Zp" + shows "res_seq s (Suc k) 0 k = res_seq s k 0 k" +proof- + obtain n where n_def: "res_seq s (Suc k) 0 = res_seq s k n" + by (metis assms is_subseq_of_def res_seq_subseq take_subseq_def) + have "res_seq s (Suc k) 0 k = res_seq s k n k" + using n_def by auto + thus ?thesis + using assms padic_integers.p_res_ring_0' + padic_integers_axioms res_seq.elims residues_closed + proof - + have "\n. s n \ carrier Zp" + by (simp add: assms closed_seqs_memE) + then show ?thesis + by (metis \res_seq s (Suc k) 0 k = res_seq s k n k\ assms padic_integers.p_res_ring_0' padic_integers_axioms res_seq.elims res_seq_res' residues_closed) + qed +qed + +lemma acc_point_cres: + assumes "s \ closed_seqs Zp" + shows "(acc_point s (Suc k)) = (const_res (Suc k) (res_seq s k))" +proof- + have "Suc k > 0" by simp + have "(res_seq s (Suc k)) = const_res_subseq (Suc k) (res_seq s k)" + by simp + then have "(const_res_subseq (Suc k) (res_seq s k)) 0 (Suc k) = const_res (Suc k) (res_seq s k)" + using assms res_seq_res' padic_integers_axioms by auto + have "acc_point s (Suc k) = res_seq s (Suc k) 0 (Suc k)" using acc_point_def by simp + then have "acc_point s (Suc k) = (const_res_subseq (Suc k) (res_seq s k)) 0 (Suc k)" + by simp + thus ?thesis + by (simp add: \(const_res_subseq (Suc k) (res_seq s k)) 0 (Suc k) = const_res (Suc k) (res_seq s k)\) +qed + +lemma acc_point_res: + assumes "s \ closed_seqs Zp" + shows "residue (p ^ k) (acc_point s (Suc k)) = acc_point s k" +proof(cases "k = 0") + case True + then show ?thesis + by (simp add: acc_point_def residue_1_zero) +next + case False + assume "k \ 0" show "residue (p ^ k) (acc_point s (Suc k)) = acc_point s k" + using False acc_point_def assms lessI less_imp_le nat.distinct(1) res_seq_res_1 res_seq_res + Zp_defs(3) closed_seqs_memE prime by (metis padic_set_res_coherent) +qed + +lemma acc_point_closed: + assumes "s \ closed_seqs Zp" + shows "acc_point s \ carrier Zp" +proof- + have "acc_point s \ padic_set p" + proof(rule padic_set_memI) + show "\m. acc_point s m \ carrier (residue_ring (p ^ m))" + proof- + fix m + show "acc_point s m \ carrier (residue_ring (p ^ m))" + proof(cases "m = 0") + case True + then show ?thesis + by (simp add: acc_point_def residue_ring_def) + next + case False + assume "m \ 0" + then have "acc_point s m = res_seq s m 0 m" (*"res_seq s (Suc k) = const_res_subseq (Suc k) (res_seq s k)"*) + by (simp add: acc_point_def) + then show ?thesis using const_res_range[of "(const_res_subseq (m-1) s)" m] acc_point_def[of s m] + by (metis False Suc_pred acc_point_cres assms const_res_range neq0_conv res_seq_res) + qed + qed + show "\m n. m < n \ residue (p ^ m) (acc_point s n) = acc_point s m" + proof- + fix m n::nat + assume A: "m < n" + show "residue (p ^ m) (acc_point s n) = acc_point s m" + proof- + obtain l where l_def: "l = n - m - 1" + by simp + have "residue (p ^ m) (acc_point s (Suc (m + l))) = acc_point s m" + proof(induction l) + case 0 + then show ?case + by (simp add: acc_point_res assms) + next + case (Suc l) + then show ?case + using Zp_defs(3) acc_point_def add_Suc_right assms le_add1 closed_seqs_memE nat.distinct(1) + padic_integers.prime padic_integers_axioms res_seq_res res_seq_res_1 + by (metis padic_set_res_coherent) + qed + then show ?thesis + by (metis A Suc_diff_Suc Suc_eq_plus1 add_Suc_right add_diff_inverse_nat diff_diff_left + l_def le_less_trans less_not_refl order_less_imp_le) + qed + qed + qed + then show ?thesis + by (simp add: Zp_defs(3)) +qed + +text\Choice function for a subsequence of s which converges to a, if it exists\ +fun convergent_subseq_fun :: "padic_int_seq \ padic_int \ (nat \ nat)" where +"convergent_subseq_fun s a 0 = 0"| +"convergent_subseq_fun s a (Suc n) = (SOME k. k > (convergent_subseq_fun s a n) + \ (s k (Suc n)) = a (Suc n))" + +definition convergent_subseq :: "padic_int_seq \ padic_int_seq" where +"convergent_subseq s = take_subseq s (convergent_subseq_fun s (acc_point s))" + +lemma increasing_conv_induction_0_pre: + assumes "s \ closed_seqs Zp" + assumes "a = acc_point s" + shows "\k > convergent_subseq_fun s a n. (s k (Suc n)) = a (Suc n)" +proof- + obtain l::nat where "l > 0 " by blast + have "is_subseq_of s (res_seq s (Suc n))" + using assms(1) res_seq_subseq' by blast + then obtain m where "s m = res_seq s (Suc n) l \ m \ l" + by (metis is_increasing_imp_geq_id is_subseq_of_def take_subseq_def ) + have "a (Suc n) = res_seq s (Suc n) 0 (Suc n)" + by (simp add: acc_point_def assms(2)) + have "s m (Suc n) = a (Suc n)" + by (metis \a (Suc n) = res_seq s (Suc n) 0 (Suc n)\ \s m = res_seq s (Suc n) l \ l \ m\ assms(1) res_seq_res') + thus ?thesis + using \0 < l\ \s m = res_seq s (Suc n) l \ l \ m\ less_le_trans \s m (Suc n) = a (Suc n)\ + by (metis \a (Suc n) = res_seq s (Suc n) 0 (Suc n)\ \is_subseq_of s (res_seq s (Suc n))\ + assms(1) lessI is_subseq_ofE res_seq_res' ) +qed + +lemma increasing_conv_subseq_fun_0: + assumes "s \ closed_seqs Zp" + assumes "\s'. s' = convergent_subseq s" + assumes "a = acc_point s" + shows "convergent_subseq_fun s a (Suc n) > convergent_subseq_fun s a n" + apply(auto) +proof(induction n) + case 0 + have "convergent_subseq_fun s a 0 = 0" by simp + then show ?case + by (smt assms(1) assms(3) less_Suc_eq less_Suc_eq_0_disj increasing_conv_induction_0_pre padic_integers_axioms someI_ex) +next + case (Suc k) + then show ?case + by (metis (mono_tags, lifting) assms(1) assms(3) increasing_conv_induction_0_pre someI_ex) +qed + +lemma increasing_conv_subseq_fun: + assumes "s \ closed_seqs Zp" + assumes "a = acc_point s" + assumes "\s'. s' = convergent_subseq s" + shows "is_increasing (convergent_subseq_fun s a)" + by (metis assms(1) assms(2) increasing_conv_subseq_fun_0 is_increasingI lift_Suc_mono_less) + +lemma convergent_subseq_is_subseq: + assumes "s \ closed_seqs Zp" + shows "is_subseq_of s (convergent_subseq s)" + using assms convergent_subseq_def increasing_conv_subseq_fun is_subseqI by blast + +lemma is_closed_seq_conv_subseq: + assumes "s \ closed_seqs Zp" + shows "(convergent_subseq s) \ closed_seqs Zp" + by (simp add: assms convergent_subseq_def closed_seqs_memI closed_seqs_memE take_subseq_def) + +lemma convergent_subseq_res: + assumes "s \ closed_seqs Zp" + assumes "a = acc_point s" + shows "convergent_subseq s l l = residue (p ^ l) (acc_point s l)" +proof- + have "\k. convergent_subseq s l = s k \ s k l = a l" + proof- + have "convergent_subseq s l = s (convergent_subseq_fun s a l)" + by (simp add: assms(2) convergent_subseq_def take_subseq_def) + obtain k where kdef: "(convergent_subseq_fun s a l) = k" + by simp + have "convergent_subseq s l = s k" + by (simp add: \convergent_subseq s l = s (convergent_subseq_fun s a l)\ kdef) + have "s k l = a l" + proof(cases "l = 0") + case True + then show ?thesis + using acc_point_def assms(1) assms(2) + by (metis closed_seqs_memE p_res_ring_0' residues_closed) + next + case False + have "0 < l" + using False by blast + then have "k > convergent_subseq_fun s a (l-1)" + by (metis One_nat_def Suc_pred assms(1) assms(2) increasing_conv_subseq_fun_0 kdef) + then have "s k l = a l" using kdef + assms(1) assms(2) convergent_subseq_fun.simps(2) increasing_conv_induction_0_pre + padic_integers_axioms someI_ex One_nat_def \0 < l\ increasing_conv_induction_0_pre + by (smt Suc_pred) + then show ?thesis + by simp + qed + then have "convergent_subseq s l = s k \ s k l = a l" + using \convergent_subseq s l = s k\ by blast + thus ?thesis + by blast + qed + thus ?thesis + using acc_point_closed assms(1) assms(2) Zp_defs(3) prime padic_set_res_coherent by force +qed + +lemma convergent_subseq_res': + assumes "s \ closed_seqs Zp" + assumes "n > l" + shows "convergent_subseq s n l = convergent_subseq s l l " +proof- + have 0: "convergent_subseq s l l = residue (p ^ l) (acc_point s l)" + using assms(1) convergent_subseq_res by auto + have 1: "convergent_subseq s n n = residue (p ^ n) (acc_point s n)" + by (simp add: assms(1) convergent_subseq_res) + have 2: "convergent_subseq s n l = residue (p ^ l) (convergent_subseq s l l)" + using 0 assms 1 Zp_defs(3) acc_point_closed is_closed_seq_conv_subseq + closed_seqs_memE le_refl less_imp_le_nat prime + by (metis padic_set_res_coherent) + show ?thesis using 0 1 2 Zp_defs(3) assms(1) is_closed_seq_conv_subseq closed_seqs_memE le_refl prime + by (metis padic_set_res_coherent) +qed + +lemma convergent_subsequence_is_convergent: + assumes "s \ closed_seqs Zp" + assumes "a = acc_point s" + shows "Zp_converges_to (convergent_subseq s) (acc_point s)" (*\n. \N. \k > N. s k n = a n"*) +proof(rule Zp_converges_toI) + show "acc_point s \ carrier Zp" + using acc_point_closed assms by blast + show "convergent_subseq s \ carrier (Zp\<^bsup>\\<^esup>)" + using is_closed_seq_conv_subseq assms by simp + show "\n. \N. \k>N. convergent_subseq s k n = acc_point s n" + proof- + fix n + show "\N. \k>N. convergent_subseq s k n = acc_point s n" + proof(induction n) + case 0 + then show ?case + using acc_point_closed[of s] assms convergent_subseq_def closed_seqs_memE of_nat_0 + ord_pos take_subseq_def zero_below_ord is_closed_seq_conv_subseq[of s] + by (metis residue_of_zero(2)) + next + case (Suc n) + have "acc_point s (Suc n) = res_seq s (Suc n) 0 (Suc n)" + by (simp add: acc_point_def) + obtain k where kdef: "convergent_subseq_fun s a (Suc n) = k" by simp + have "Suc n > 0" by simp + then have "k > (convergent_subseq_fun s a n)" + using assms(1) assms(2) increasing_conv_subseq_fun_0 kdef by blast + then have " k > (convergent_subseq_fun s a n) \ (s k (Suc n)) = a (Suc n)" using kdef + by (metis (mono_tags, lifting) assms(1) assms(2) convergent_subseq_fun.simps(2) increasing_conv_induction_0_pre someI_ex) + have "s k (Suc n) = a (Suc n)" + using \convergent_subseq_fun s a n < k \ s k (Suc n) = a (Suc n)\ by blast + then have "convergent_subseq s (Suc n) (Suc n) = a (Suc n)" + by (metis assms(2) convergent_subseq_def kdef take_subseq_def) + then have "\l > n. convergent_subseq s l (Suc n) = a (Suc n)" + using convergent_subseq_res' + by (metis Suc_lessI assms(1)) + then show ?case + using assms(2) by blast + qed + qed +qed + +theorem Zp_is_compact: + assumes "s \ closed_seqs Zp" + shows "\s'. is_subseq_of s s' \ (Zp_converges_to s' (acc_point s))" + using assms convergent_subseq_is_subseq convergent_subsequence_is_convergent + by blast + +end +end diff --git a/thys/Padic_Ints/document/root.bib b/thys/Padic_Ints/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Padic_Ints/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_Ints/document/root.tex b/thys/Padic_Ints/document/root.tex new file mode 100755 --- /dev/null +++ b/thys/Padic_Ints/document/root.tex @@ -0,0 +1,62 @@ +\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 Hensel's Lemma} +\author{Aaron Crighton} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +\begin{abstract} +We formalize the ring of $p$-adic integers within the framework of the HOL-Algebra library. The carrier of the ring $\mathbb{Z}_p$ is formalized as the inverse limit of the residue rings $\mathbb{Z}/p^n\mathbb{Z}$ for a fixed prime $p$. We define a locale for reasoning about $\mathbb{Z}_p$ for a fixed prime $p$, and define an integer-valued valuation, as well as an extended-integer valued valuation on $\mathbb{Z}_p$ (where $0 \in \mathbb{Z}_p$ is the unique ring element mapped to $\infty$). Basic topological facts about the $p$-adic integers are formalized, including the completeness and sequential compactness of $\mathbb{Z}_p$. Taylor expansions of polynomials over a commutative ring are defined, culminating in the formalization of Hensel's Lemma based on a proof due to Keith Conrad \cite{keithconrad}. +\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,599 +1,600 @@ 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 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 BNF_CC BNF_Operations BTree Banach_Steinhaus 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 Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK 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 ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties CSP_RefTK DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DOM_Components DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games 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 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 IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaGeoCoq Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knuth_Bendix_Order 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 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 Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem 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_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_arithmetic_LLL_and_HNF_algorithms Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference 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 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 PCF PLM POPLmark-deBruijn PSemigroupsConvolution +Padic_Ints 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 Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate 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 QHLProver QR_Decomposition Quantales 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 Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regression_Test_Selection Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Shadow_DOM Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem 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 Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL