# HG changeset patch # User wenzelm # Date 1483652286 -3600 # Node ID a0e1f64be67cb1f1b5e3bc0787d8975bfc9eeda7 # Parent 8e7db8df16a041d7496610dfc38be577b06f16aa# Parent 81a5473e6d047295caf670984c36c6facafb85b0 merged diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Analysis.thy Thu Jan 05 22:38:06 2017 +0100 @@ -12,6 +12,7 @@ Weierstrass_Theorems Polytope Further_Topology + Arcwise_Connected Poly_Roots Conformal_Mappings Generalised_Binomial_Theorem diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Arcwise_Connected.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Arcwise_Connected.thy Thu Jan 05 22:38:06 2017 +0100 @@ -0,0 +1,2213 @@ +(* Title: HOL/Analysis/Arcwise_Connected.thy + Authors: LC Paulson, based on material from HOL Light +*) + +section \Arcwise-connected sets\ + +theory Arcwise_Connected + imports Path_Connected Ordered_Euclidean_Space "~~/src/HOL/Number_Theory/Primes" + +begin + +subsection\The Brouwer reduction theorem\ + +theorem Brouwer_reduction_theorem_gen: + fixes S :: "'a::euclidean_space set" + assumes "closed S" "\ S" + and \: "\F. \\n. closed(F n); \n. \(F n); \n. F(Suc n) \ F n\ \ \(\range F)" + obtains T where "T \ S" "closed T" "\ T" "\U. \U \ S; closed U; \ U\ \ \ (U \ T)" +proof - + obtain B :: "nat \ 'a set" + where "inj B" "\n. open(B n)" and open_cov: "\S. open S \ \K. S = \(B ` K)" + by (metis Setcompr_eq_image that univ_second_countable_sequence) + define A where "A \ rec_nat S (\n a. if \U. U \ a \ closed U \ \ U \ U \ (B n) = {} + then @U. U \ a \ closed U \ \ U \ U \ (B n) = {} + else a)" + have [simp]: "A 0 = S" + by (simp add: A_def) + have ASuc: "A(Suc n) = (if \U. U \ A n \ closed U \ \ U \ U \ (B n) = {} + then @U. U \ A n \ closed U \ \ U \ U \ (B n) = {} + else A n)" for n + by (auto simp: A_def) + have sub: "\n. A(Suc n) \ A n" + by (auto simp: ASuc dest!: someI_ex) + have subS: "A n \ S" for n + by (induction n) (use sub in auto) + have clo: "closed (A n) \ \ (A n)" for n + by (induction n) (auto simp: assms ASuc dest!: someI_ex) + show ?thesis + proof + show "\range A \ S" + using \\n. A n \ S\ by blast + show "closed (INTER UNIV A)" + using clo by blast + show "\ (INTER UNIV A)" + by (simp add: clo \ sub) + show "\ U \ INTER UNIV A" if "U \ S" "closed U" "\ U" for U + proof - + have "\y. x \ A y" if "x \ U" and Usub: "U \ (\x. A x)" for x + proof - + obtain e where "e > 0" and e: "ball x e \ -U" + using \closed U\ \x \ U\ openE [of "-U"] by blast + moreover obtain K where K: "ball x e = UNION K B" + using open_cov [of "ball x e"] by auto + ultimately have "UNION K B \ -U" + by blast + have "K \ {}" + using \0 < e\ \ball x e = UNION K B\ by auto + then obtain n where "n \ K" "x \ B n" + by (metis K UN_E \0 < e\ centre_in_ball) + then have "U \ B n = {}" + using K e by auto + show ?thesis + proof (cases "\U\A n. closed U \ \ U \ U \ B n = {}") + case True + then show ?thesis + apply (rule_tac x="Suc n" in exI) + apply (simp add: ASuc) + apply (erule someI2_ex) + using \x \ B n\ by blast + next + case False + then show ?thesis + by (meson Inf_lower Usub \U \ B n = {}\ \\ U\ \closed U\ range_eqI subset_trans) + qed + qed + with that show ?thesis + by (meson Inter_iff psubsetE rangeI subsetI) + qed + qed +qed + +corollary Brouwer_reduction_theorem: + fixes S :: "'a::euclidean_space set" + assumes "compact S" "\ S" "S \ {}" + and \: "\F. \\n. compact(F n); \n. F n \ {}; \n. \(F n); \n. F(Suc n) \ F n\ \ \(\range F)" + obtains T where "T \ S" "compact T" "T \ {}" "\ T" + "\U. \U \ S; closed U; U \ {}; \ U\ \ \ (U \ T)" +proof (rule Brouwer_reduction_theorem_gen [of S "\T. T \ {} \ T \ S \ \ T"]) + fix F + assume cloF: "\n. closed (F n)" + and F: "\n. F n \ {} \ F n \ S \ \ (F n)" and Fsub: "\n. F (Suc n) \ F n" + show "INTER UNIV F \ {} \ INTER UNIV F \ S \ \ (INTER UNIV F)" + proof (intro conjI) + show "INTER UNIV F \ {}" + apply (rule compact_nest) + apply (meson F cloF \compact S\ seq_compact_closed_subset seq_compact_eq_compact) + apply (simp add: F) + by (meson Fsub lift_Suc_antimono_le) + show " INTER UNIV F \ S" + using F by blast + show "\ (INTER UNIV F)" + by (metis F Fsub \ \compact S\ cloF closed_Int_compact inf.orderE) + qed +next + show "S \ {} \ S \ S \ \ S" + by (simp add: assms) +qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+ + + +subsection\Arcwise Connections\ + +subsection\Density of points with dyadic rational coordinates.\ + +proposition closure_dyadic_rationals: + "closure (\k. \f \ Basis \ \. + { \i :: 'a :: euclidean_space \ Basis. (f i / 2^k) *\<^sub>R i }) = UNIV" +proof - + have "x \ closure (\k. \f \ Basis \ \. {\i \ Basis. (f i / 2^k) *\<^sub>R i})" for x::'a + proof (clarsimp simp: closure_approachable) + fix e::real + assume "e > 0" + then obtain k where k: "(1/2)^k < e/DIM('a)" + by (meson DIM_positive divide_less_eq_1_pos of_nat_0_less_iff one_less_numeral_iff real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral) + have "dist (\i\Basis. (real_of_int \2^k*(x \ i)\ / 2^k) *\<^sub>R i) x = + dist (\i\Basis. (real_of_int \2^k*(x \ i)\ / 2^k) *\<^sub>R i) (\i\Basis. (x \ i) *\<^sub>R i)" + by (simp add: euclidean_representation) + also have "... = norm ((\i\Basis. (real_of_int \2^k*(x \ i)\ / 2^k) *\<^sub>R i - (x \ i) *\<^sub>R i))" + by (simp add: dist_norm sum_subtractf) + also have "... \ DIM('a)*((1/2)^k)" + proof (rule sum_norm_bound, simp add: algebra_simps) + fix i::'a + assume "i \ Basis" + then have "norm ((real_of_int \x \ i*2^k\ / 2^k) *\<^sub>R i - (x \ i) *\<^sub>R i) = + \real_of_int \x \ i*2^k\ / 2^k - x \ i\" + by (simp add: scaleR_left_diff_distrib [symmetric]) + also have "... \ (1/2) ^ k" + by (simp add: divide_simps) linarith + finally show "norm ((real_of_int \x \ i*2^k\ / 2^k) *\<^sub>R i - (x \ i) *\<^sub>R i) \ (1/2) ^ k" . + qed + also have "... < DIM('a)*(e/DIM('a))" + using DIM_positive k linordered_comm_semiring_strict_class.comm_mult_strict_left_mono of_nat_0_less_iff by blast + also have "... = e" + by simp + finally have "dist (\i\Basis. (\2^k*(x \ i)\ / 2^k) *\<^sub>R i) x < e" . + then + show "\k. \f \ Basis \ \. dist (\b\Basis. (f b / 2^k) *\<^sub>R b) x < e" + apply (rule_tac x=k in exI) + apply (rule_tac x="\i. of_int (floor (2^k*(x \ i)))" in bexI) + apply auto + done + qed + then show ?thesis by auto +qed + +corollary closure_rational_coordinates: + "closure (\f \ Basis \ \. { \i :: 'a :: euclidean_space \ Basis. f i *\<^sub>R i }) = UNIV" +proof - + have *: "(\k. \f \ Basis \ \. { \i::'a \ Basis. (f i / 2^k) *\<^sub>R i }) + \ (\f \ Basis \ \. { \i \ Basis. f i *\<^sub>R i })" + proof clarsimp + fix k and f :: "'a \ real" + assume f: "f \ Basis \ \" + show "\x \ Basis \ \. (\i \ Basis. (f i / 2^k) *\<^sub>R i) = (\i \ Basis. x i *\<^sub>R i)" + apply (rule_tac x="\i. f i / 2^k" in bexI) + using Ints_subset_Rats f by auto + qed + show ?thesis + using closure_dyadic_rationals closure_mono [OF *] by blast +qed + +lemma closure_dyadic_rationals_in_convex_set: + "\convex S; interior S \ {}\ + \ closure(S \ + (\k. \f \ Basis \ \. + { \i :: 'a :: euclidean_space \ Basis. (f i / 2^k) *\<^sub>R i })) = + closure S" + by (simp add: closure_dyadic_rationals closure_convex_Int_superset) + +lemma closure_rationals_in_convex_set: + "\convex S; interior S \ {}\ + \ closure(S \ (\f \ Basis \ \. { \i :: 'a :: euclidean_space \ Basis. f i *\<^sub>R i })) = + closure S" + by (simp add: closure_rational_coordinates closure_convex_Int_superset) + + +text\ Every path between distinct points contains an arc, and hence +path connection is equivalent to arcwise connection for distinct points. +The proof is based on Whyburn's "Topological Analysis".\ + +lemma closure_dyadic_rationals_in_convex_set_pos_1: + fixes S :: "real set" + assumes "convex S" and intnz: "interior S \ {}" and pos: "\x. x \ S \ 0 \ x" + shows "closure(S \ (\k m. {of_nat m / 2^k})) = closure S" +proof - + have "\m. f 1/2^k = real m / 2^k" if "(f 1) / 2^k \ S" "f 1 \ \" for k and f :: "real \ real" + using that by (force simp: Ints_def zero_le_divide_iff power_le_zero_eq dest: pos zero_le_imp_eq_int) + then have "S \ (\k m. {real m / 2^k}) = S \ + (\k. \f\Basis \ \. {\i\Basis. (f i / 2^k) *\<^sub>R i})" + by force + then show ?thesis + using closure_dyadic_rationals_in_convex_set [OF \convex S\ intnz] by simp +qed + + +definition dyadics :: "'a::field_char_0 set" where "dyadics \ \k m. {of_nat m / 2^k}" + +lemma real_in_dyadics [simp]: "real m \ dyadics" + apply (simp add: dyadics_def) + by (metis divide_numeral_1 numeral_One power_0) + +lemma nat_neq_4k1: "of_nat m \ (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)" +proof + assume "of_nat m = (4 * of_nat k + 1) / (2 * 2^n :: 'a)" + then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)" + by (simp add: divide_simps) + then have "m * (2 * 2^n) = Suc (4 * k)" + using of_nat_eq_iff by blast + then have "odd (m * (2 * 2^n))" + by simp + then show False + by simp +qed + +lemma nat_neq_4k3: "of_nat m \ (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)" +proof + assume "of_nat m = (4 * of_nat k + 3) / (2 * 2^n :: 'a)" + then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)" + by (simp add: divide_simps) + then have "m * (2 * 2^n) = (4 * k) + 3" + using of_nat_eq_iff by blast + then have "odd (m * (2 * 2^n))" + by simp + then show False + by simp +qed + +lemma iff_4k: + assumes "r = real k" "odd k" + shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \ m=m' \ n=n'" +proof - + { assume "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n')" + then have "real ((4 * m + k) * (2 * 2 ^ n')) = real ((4 * m' + k) * (2 * 2^n))" + using assms by (auto simp: field_simps) + then have "(4 * m + k) * (2 * 2 ^ n') = (4 * m' + k) * (2 * 2^n)" + using of_nat_eq_iff by blast + then have "(4 * m + k) * (2 ^ n') = (4 * m' + k) * (2^n)" + by linarith + then obtain "4*m + k = 4*m' + k" "n=n'" + apply (rule prime_power_cancel2 [OF two_is_prime_nat]) + using assms by auto + then have "m=m'" "n=n'" + by auto + } + then show ?thesis by blast +qed + +lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \ (4 * real m' + 3) / (2 * 2 ^ n')" +proof + assume "(4 * real m + 1) / (2 * 2^n) = (4 * real m' + 3) / (2 * 2 ^ n')" + then have "real (Suc (4 * m) * (2 * 2 ^ n')) = real ((4 * m' + 3) * (2 * 2^n))" + by (auto simp: field_simps) + then have "Suc (4 * m) * (2 * 2 ^ n') = (4 * m' + 3) * (2 * 2^n)" + using of_nat_eq_iff by blast + then have "Suc (4 * m) * (2 ^ n') = (4 * m' + 3) * (2^n)" + by linarith + then have "Suc (4 * m) = (4 * m' + 3)" + by (rule prime_power_cancel2 [OF two_is_prime_nat]) auto + then have "1 + 2 * m' = 2 * m" + using \Suc (4 * m) = 4 * m' + 3\ by linarith + then show False + using even_Suc by presburger +qed + +lemma dyadic_413_cases: + obtains "(of_nat m::'a::field_char_0) / 2^k \ Nats" + | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 1) / 2^Suc k'" + | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'" +proof (cases "m>0") + case False + then have "m=0" by simp + with that show ?thesis by auto +next + case True + obtain k' m' where m': "odd m'" and k': "m = m' * 2^k'" + using prime_power_canonical [OF two_is_prime_nat True] by blast + then obtain q r where q: "m' = 4*q + r" and r: "r < 4" + by (metis not_add_less2 split_div zero_neq_numeral) + show ?thesis + proof (cases "k \ k'") + case True + have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)" + using k' by (simp add: field_simps) + also have "... = (of_nat m'::'a) * 2 ^ (k'-k)" + using k' True by (simp add: power_diff) + also have "... \ \" + by (metis Nats_mult of_nat_in_Nats of_nat_numeral of_nat_power) + finally show ?thesis by (auto simp: that) + next + case False + then obtain kd where kd: "Suc kd = k - k'" + using Suc_diff_Suc not_less by blast + have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)" + using k' by (simp add: field_simps) + also have "... = (of_nat m'::'a) / 2 ^ (k-k')" + using k' False by (simp add: power_diff) + also have "... = ((of_nat r + 4 * of_nat q)::'a) / 2 ^ (k-k')" + using q by force + finally have meq: "(of_nat m:: 'a) / 2^k = (of_nat r + 4 * of_nat q) / 2 ^ (k - k')" . + have "r \ 0" "r \ 2" + using q m' by presburger+ + with r consider "r = 1" | "r = 3" + by linarith + then show ?thesis + proof cases + assume "r = 1" + with meq kd that(2) [of kd q] show ?thesis + by simp + next + assume "r = 3" + with meq kd that(3) [of kd q] show ?thesis + by simp + qed + qed +qed + + +lemma dyadics_iff: + "(dyadics :: 'a::field_char_0 set) = + Nats \ (\k m. {of_nat (4*m + 1) / 2^Suc k}) \ (\k m. {of_nat (4*m + 3) / 2^Suc k})" + (is "_ = ?rhs") +proof + show "dyadics \ ?rhs" + unfolding dyadics_def + apply clarify + apply (rule dyadic_413_cases, force+) + done +next + show "?rhs \ dyadics" + apply (clarsimp simp: dyadics_def Nats_def simp del: power_Suc) + apply (intro conjI subsetI) + apply (auto simp del: power_Suc) + apply (metis divide_numeral_1 numeral_One power_0) + apply (metis of_nat_Suc of_nat_mult of_nat_numeral) + by (metis of_nat_add of_nat_mult of_nat_numeral) +qed + + +function (domintros) dyad_rec :: "[nat \ 'a, 'a\'a, 'a\'a, real] \ 'a" where + "dyad_rec b l r (real m) = b m" + | "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))" + | "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))" + | "x \ dyadics \ dyad_rec b l r x = undefined" + using iff_4k [of _ 1] iff_4k [of _ 3] + apply (simp_all add: nat_neq_4k1 nat_neq_4k3 neq_4k1_k43, atomize_elim) + apply (fastforce simp add: dyadics_iff Nats_def field_simps)+ + done + +lemma dyadics_levels: "dyadics = (\K. \k m. {of_nat m / 2^k})" + unfolding dyadics_def by auto + +lemma dyad_rec_level_termination: + assumes "k < K" + shows "dyad_rec_dom(b, l, r, real m / 2^k)" + using assms +proof (induction K arbitrary: k m) + case 0 + then show ?case by auto +next + case (Suc K) + then consider "k = K" | "k < K" + using less_antisym by blast + then show ?case + proof cases + assume "k = K" + show ?case + proof (rule dyadic_413_cases [of m k, where 'a=real]) + show "real m / 2^k \ \ \ dyad_rec_dom (b, l, r, real m / 2^k)" + by (force simp: Nats_def nat_neq_4k1 nat_neq_4k3 intro: dyad_rec.domintros) + show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 1) / 2^Suc k'" for m' k' + proof - + have "dyad_rec_dom (b, l, r, (4 * real m' + 1) / 2^Suc k')" + proof (rule dyad_rec.domintros) + fix m n + assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)" + then have "m' = m" "k' = n" using iff_4k [of _ 1] + by auto + have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')" + using Suc.IH \k = K\ \k' < k\ by blast + then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" + using \k' = n\ by (auto simp: algebra_simps) + next + fix m n + assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)" + then have "False" + by (metis neq_4k1_k43) + then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" .. + qed + then show ?case by (simp add: eq add_ac) + qed + show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 3) / 2^Suc k'" for m' k' + proof - + have "dyad_rec_dom (b, l, r, (4 * real m' + 3) / 2^Suc k')" + proof (rule dyad_rec.domintros) + fix m n + assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)" + then have "False" + by (metis neq_4k1_k43) + then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" .. + next + fix m n + assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)" + then have "m' = m" "k' = n" using iff_4k [of _ 3] + by auto + have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')" + using Suc.IH \k = K\ \k' < k\ by blast + then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" + using \k' = n\ by (auto simp: algebra_simps) + qed + then show ?case by (simp add: eq add_ac) + qed + qed + next + assume "k < K" + then show ?case + using Suc.IH by blast + qed +qed + + +lemma dyad_rec_termination: "x \ dyadics \ dyad_rec_dom(b,l,r,x)" + by (auto simp: dyadics_levels intro: dyad_rec_level_termination) + +lemma dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m" + by (simp add: dyad_rec.psimps dyad_rec_termination) + +lemma dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))" + apply (rule dyad_rec.psimps) + by (metis dyad_rec_level_termination lessI add.commute of_nat_Suc of_nat_mult of_nat_numeral) + + +lemma dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))" + apply (rule dyad_rec.psimps) + by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral) + +lemma dyad_rec_41_times2: + assumes "n > 0" + shows "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" +proof - + obtain n' where n': "n = Suc n'" + using assms not0_implies_Suc by blast + have "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 1)) / (2 * 2^n))" + by auto + also have "... = dyad_rec b l r ((4 * real m + 1) / 2^n)" + by (subst mult_divide_mult_cancel_left) auto + also have "... = l (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))" + by (simp add: add.commute [of 1] n' del: power_Suc) + also have "... = l (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))" + by (subst mult_divide_mult_cancel_left) auto + also have "... = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" + by (simp add: add.commute n') + finally show ?thesis . +qed + +lemma dyad_rec_43_times2: + assumes "n > 0" + shows "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" +proof - + obtain n' where n': "n = Suc n'" + using assms not0_implies_Suc by blast + have "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 3)) / (2 * 2^n))" + by auto + also have "... = dyad_rec b l r ((4 * real m + 3) / 2^n)" + by (subst mult_divide_mult_cancel_left) auto + also have "... = r (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))" + by (simp add: n' del: power_Suc) + also have "... = r (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))" + by (subst mult_divide_mult_cancel_left) auto + also have "... = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" + by (simp add: n') + finally show ?thesis . +qed + +definition dyad_rec2 + where "dyad_rec2 u v lc rc x = + dyad_rec (\z. (u,v)) (\(a,b). (a, lc a b (midpoint a b))) (\(a,b). (rc a b (midpoint a b), b)) (2*x)" + +abbreviation leftrec where "leftrec u v lc rc x \ fst (dyad_rec2 u v lc rc x)" +abbreviation rightrec where "rightrec u v lc rc x \ snd (dyad_rec2 u v lc rc x)" + +lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u" + by (simp add: dyad_rec2_def) + +lemma leftrec_41: "n > 0 \ leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)" + apply (simp only: dyad_rec2_def dyad_rec_41_times2) + apply (simp add: case_prod_beta) + done + +lemma leftrec_43: "n > 0 \ + leftrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = + rc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)) + (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))" + apply (simp only: dyad_rec2_def dyad_rec_43_times2) + apply (simp add: case_prod_beta) + done + +lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v" + by (simp add: dyad_rec2_def) + +lemma rightrec_41: "n > 0 \ + rightrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = + lc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)) + (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))" + apply (simp only: dyad_rec2_def dyad_rec_41_times2) + apply (simp add: case_prod_beta) + done + +lemma rightrec_43: "n > 0 \ rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)" + apply (simp only: dyad_rec2_def dyad_rec_43_times2) + apply (simp add: case_prod_beta) + done + +lemma dyadics_in_open_unit_interval: + "{0<..<1} \ (\k m. {real m / 2^k}) = (\k. \m \ {0<..<2^k}. {real m / 2^k})" + by (auto simp: divide_simps) + + + +theorem homeomorphic_monotone_image_interval: + fixes f :: "real \ 'a::{real_normed_vector,complete_space}" + assumes cont_f: "continuous_on {0..1} f" + and conn: "\y. connected ({0..1} \ f -` {y})" + and f_1not0: "f 1 \ f 0" + shows "(f ` {0..1}) homeomorphic {0..1::real}" +proof - + have "\c d. a \ c \ c \ m \ m \ d \ d \ b \ + (\x \ {c..d}. f x = f m) \ + (\x \ {a.. f m)) \ + (\x \ {d<..b}. (f x \ f m)) \ + (\x \ {a..y \ {d<..b}. f x \ f y)" + if m: "m \ {a..b}" and ab01: "{a..b} \ {0..1}" for a b m + proof - + have comp: "compact (f -` {f m} \ {0..1})" + by (simp add: compact_eq_bounded_closed bounded_Int closed_vimage_Int cont_f) + obtain c0 d0 where cd0: "{0..1} \ f -` {f m} = {c0..d0}" + using connected_compact_interval_1 [of "{0..1} \ f -` {f m}"] conn comp + by (metis Int_commute) + with that have "m \ cbox c0 d0" + by auto + obtain c d where cd: "{a..b} \ f -` {f m} = {c..d}" + apply (rule_tac c="max a c0" and d="min b d0" in that) + using ab01 cd0 by auto + then have cdab: "{c..d} \ {a..b}" + by blast + show ?thesis + proof (intro exI conjI ballI) + show "a \ c" "d \ b" + using cdab cd m by auto + show "c \ m" "m \ d" + using cd m by auto + show "\x. x \ {c..d} \ f x = f m" + using cd by blast + show "f x \ f m" if "x \ {a..c \ m\ by force + show "f x \ f m" if "x \ {d<..b}" for x + using that m cd [THEN equalityD1, THEN subsetD, of x] \m \ d\ by force + show "f x \ f y" if "x \ {a.. {d<..b}" for x y + proof (cases "f x = f m \ f y = f m") + case True + then show ?thesis + using \\x. x \ {a.. f x \ f m\ that by auto + next + case False + have False if "f x = f y" + proof - + have "x \ m" "m \ y" + using \c \ m\ \x \ {a.. \m \ d\ \y \ {d<..b}\ by auto + then have "x \ ({0..1} \ f -` {f y})" "y \ ({0..1} \ f -` {f y})" + using \x \ {a.. \y \ {d<..b}\ ab01 by (auto simp: that) + then have "m \ ({0..1} \ f -` {f y})" + by (meson \m \ y\ \x \ m\ is_interval_connected_1 conn [of "f y"] is_interval_1) + with False show False by auto + qed + then show ?thesis by auto + qed + qed + qed + then obtain leftcut rightcut where LR: + "\a b m. \m \ {a..b}; {a..b} \ {0..1}\ \ + (a \ leftcut a b m \ leftcut a b m \ m \ m \ rightcut a b m \ rightcut a b m \ b \ + (\x \ {leftcut a b m..rightcut a b m}. f x = f m) \ + (\x \ {a.. f m) \ + (\x \ {rightcut a b m<..b}. f x \ f m) \ + (\x \ {a..y \ {rightcut a b m<..b}. f x \ f y))" + apply atomize + apply (clarsimp simp only: imp_conjL [symmetric] choice_iff choice_iff') + apply (rule that, blast) + done + then have left_right: "\a b m. \m \ {a..b}; {a..b} \ {0..1}\ \ a \ leftcut a b m \ rightcut a b m \ b" + and left_right_m: "\a b m. \m \ {a..b}; {a..b} \ {0..1}\ \ leftcut a b m \ m \ m \ rightcut a b m" + by auto + have left_neq: "\a \ x; x < leftcut a b m; a \ m; m \ b; {a..b} \ {0..1}\ \ f x \ f m" + and right_neq: "\rightcut a b m < x; x \ b; a \ m; m \ b; {a..b} \ {0..1}\ \ f x \ f m" + and left_right_neq: "\a \ x; x < leftcut a b m; rightcut a b m < y; y \ b; a \ m; m \ b; {a..b} \ {0..1}\ \ f x \ f m" + and feqm: "\leftcut a b m \ x; x \ rightcut a b m; a \ m; m \ b; {a..b} \ {0..1}\ + \ f x = f m" for a b m x y + by (meson atLeastAtMost_iff greaterThanAtMost_iff atLeastLessThan_iff LR)+ + have f_eqI: "\a b m x y. \leftcut a b m \ x; x \ rightcut a b m; leftcut a b m \ y; y \ rightcut a b m; + a \ m; m \ b; {a..b} \ {0..1}\ \ f x = f y" + by (metis feqm) + define u where "u \ rightcut 0 1 0" + have lc[simp]: "leftcut 0 1 0 = 0" and u01: "0 \ u" "u \ 1" + using LR [of 0 0 1] by (auto simp: u_def) + have f0u: "\x. x \ {0..u} \ f x = f 0" + using LR [of 0 0 1] unfolding u_def [symmetric] + by (metis \leftcut 0 1 0 = 0\ atLeastAtMost_iff order_refl zero_le_one) + have fu1: "\x. x \ {u<..1} \ f x \ f 0" + using LR [of 0 0 1] unfolding u_def [symmetric] by fastforce + define v where "v \ leftcut u 1 1" + have rc[simp]: "rightcut u 1 1 = 1" and v01: "u \ v" "v \ 1" + using LR [of 1 u 1] u01 by (auto simp: v_def) + have fuv: "\x. x \ {u.. f x \ f 1" + using LR [of 1 u 1] u01 v_def by fastforce + have f0v: "\x. x \ {0.. f x \ f 1" + by (metis f_1not0 atLeastAtMost_iff atLeastLessThan_iff f0u fuv linear) + have fv1: "\x. x \ {v..1} \ f x = f 1" + using LR [of 1 u 1] u01 v_def by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl rc) + define a where "a \ leftrec u v leftcut rightcut" + define b where "b \ rightrec u v leftcut rightcut" + define c where "c \ \x. midpoint (a x) (b x)" + have a_real [simp]: "a (real j) = u" for j + using a_def leftrec_base + by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral) + have b_real [simp]: "b (real j) = v" for j + using b_def rightrec_base + by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral) + have a41: "a ((4 * real m + 1) / 2^Suc n) = a ((2 * real m + 1) / 2^n)" if "n > 0" for m n + using that a_def leftrec_41 by blast + have b41: "b ((4 * real m + 1) / 2^Suc n) = + leftcut (a ((2 * real m + 1) / 2^n)) + (b ((2 * real m + 1) / 2^n)) + (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n + using that a_def b_def c_def rightrec_41 by blast + have a43: "a ((4 * real m + 3) / 2^Suc n) = + rightcut (a ((2 * real m + 1) / 2^n)) + (b ((2 * real m + 1) / 2^n)) + (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n + using that a_def b_def c_def leftrec_43 by blast + have b43: "b ((4 * real m + 3) / 2^Suc n) = b ((2 * real m + 1) / 2^n)" if "n > 0" for m n + using that b_def rightrec_43 by blast + have uabv: "u \ a (real m / 2 ^ n) \ a (real m / 2 ^ n) \ b (real m / 2 ^ n) \ b (real m / 2 ^ n) \ v" for m n + proof (induction n arbitrary: m) + case 0 + then show ?case by (simp add: v01) + next + case (Suc n p) + show ?case + proof (cases "even p") + case True + then obtain m where "p = 2*m" by (metis evenE) + then show ?thesis + by (simp add: Suc.IH) + next + case False + then obtain m where m: "p = 2*m + 1" by (metis oddE) + show ?thesis + proof (cases n) + case 0 + then show ?thesis + by (simp add: a_def b_def leftrec_base rightrec_base v01) + next + case (Suc n') + then have "n > 0" by simp + have a_le_c: "a (real m / 2^n) \ c (real m / 2^n)" for m + unfolding c_def by (metis Suc.IH ge_midpoint_1) + have c_le_b: "c (real m / 2^n) \ b (real m / 2^n)" for m + unfolding c_def by (metis Suc.IH le_midpoint_1) + have c_ge_u: "c (real m / 2^n) \ u" for m + using Suc.IH a_le_c order_trans by blast + have c_le_v: "c (real m / 2^n) \ v" for m + using Suc.IH c_le_b order_trans by blast + have a_ge_0: "0 \ a (real m / 2^n)" for m + using Suc.IH order_trans u01(1) by blast + have b_le_1: "b (real m / 2^n) \ 1" for m + using Suc.IH order_trans v01(2) by blast + have left_le: "leftcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \ c ((real m) / 2^n)" for m + by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b) + have right_ge: "rightcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \ c ((real m) / 2^n)" for m + by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b) + show ?thesis + proof (cases "even m") + case True + then obtain r where r: "m = 2*r" by (metis evenE) + show ?thesis + using a_le_c [of "m+1"] c_le_b [of "m+1"] a_ge_0 [of "m+1"] b_le_1 [of "m+1"] + Suc.IH [of "m+1"] + apply (simp add: r m add.commute [of 1] \n > 0\ a41 b41 del: power_Suc) + apply (auto simp: left_right [THEN conjunct1]) + using order_trans [OF left_le c_le_v] + by (metis (no_types, hide_lams) add.commute mult_2 of_nat_Suc of_nat_add) + next + case False + then obtain r where r: "m = 2*r + 1" by (metis oddE) + show ?thesis + using a_le_c [of "m"] c_le_b [of "m"] a_ge_0 [of "m"] b_le_1 [of "m"] + Suc.IH [of "m+1"] + apply (simp add: r m add.commute [of 3] \n > 0\ a43 b43 del: power_Suc) + apply (auto simp: add.commute left_right [THEN conjunct2]) + using order_trans [OF c_ge_u right_ge] + apply (metis (no_types, hide_lams) mult_2 numeral_One of_nat_add of_nat_numeral) + apply (metis Suc.IH mult_2 of_nat_1 of_nat_add) + done + qed + qed + qed + qed + have a_ge_0 [simp]: "0 \ a(m / 2^n)" and b_le_1 [simp]: "b(m / 2^n) \ 1" for m::nat and n + using uabv order_trans u01 v01 by blast+ + then have b_ge_0 [simp]: "0 \ b(m / 2^n)" and a_le_1 [simp]: "a(m / 2^n) \ 1" for m::nat and n + using uabv order_trans by blast+ + have alec [simp]: "a(m / 2^n) \ c(m / 2^n)" and cleb [simp]: "c(m / 2^n) \ b(m / 2^n)" for m::nat and n + by (auto simp: c_def ge_midpoint_1 le_midpoint_1 uabv) + have c_ge_0 [simp]: "0 \ c(m / 2^n)" and c_le_1 [simp]: "c(m / 2^n) \ 1" for m::nat and n + using a_ge_0 alec order_trans apply blast + by (meson b_le_1 cleb order_trans) + have "\d = m-n; odd j; \real i / 2^m - real j / 2^n\ < 1/2 ^ n\ + \ (a(j / 2^n)) \ (c(i / 2^m)) \ (c(i / 2^m)) \ (b(j / 2^n))" for d i j m n + proof (induction d arbitrary: j n rule: less_induct) + case (less d j n) + show ?case + proof (cases "m \ n") + case True + have "\2^n\ * \real i / 2^m - real j / 2^n\ = 0" + proof (rule Ints_nonzero_abs_less1) + have "(real i * 2^n - real j * 2^m) / 2^m = (real i * 2^n) / 2^m - (real j * 2^m) / 2^m" + using diff_divide_distrib by blast + also have "... = (real i * 2 ^ (n-m)) - (real j)" + using True by (auto simp: power_diff field_simps) + also have "... \ \" + by simp + finally have "(real i * 2^n - real j * 2^m) / 2^m \ \" . + with True Ints_abs show "\2^n\ * \real i / 2^m - real j / 2^n\ \ \" + by (fastforce simp: divide_simps) + show "\\2^n\ * \real i / 2^m - real j / 2^n\\ < 1" + using less.prems by (auto simp: divide_simps) + qed + then have "real i / 2^m = real j / 2^n" + by auto + then show ?thesis + by auto + next + case False + then have "n < m" by auto + obtain k where k: "j = Suc (2*k)" + using \odd j\ oddE by fastforce + show ?thesis + proof (cases "n > 0") + case False + then have "a (real j / 2^n) = u" + by simp + also have "... \ c (real i / 2^m)" + using alec uabv by (blast intro: order_trans) + finally have ac: "a (real j / 2^n) \ c (real i / 2^m)" . + have "c (real i / 2^m) \ v" + using cleb uabv by (blast intro: order_trans) + also have "... = b (real j / 2^n)" + using False by simp + finally show ?thesis + by (auto simp: ac) + next + case True show ?thesis + proof (cases "real i / 2^m" "real j / 2^n" rule: linorder_cases) + case less + moreover have "real (4 * k + 1) / 2 ^ Suc n + 1 / (2 ^ Suc n) = real j / 2 ^ n" + using k by (force simp: divide_simps) + moreover have "\real i / 2 ^ m - real j / 2 ^ n\ < 2 / (2 ^ Suc n)" + using less.prems by simp + ultimately have closer: "\real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\ < 1 / (2 ^ Suc n)" + using less.prems by linarith + have *: "a (real (4 * k + 1) / 2 ^ Suc n) \ c (real i / 2 ^ m) \ + c (real i / 2 ^ m) \ b (real (4 * k + 1) / 2 ^ Suc n)" + apply (rule less.IH [OF _ refl]) + using closer \n < m\ \d = m - n\ apply (auto simp: divide_simps \n < m\ diff_less_mono2) + done + show ?thesis + using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] + using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"] + using k a41 b41 * \0 < n\ + apply (simp add: add.commute) + done + next + case equal then show ?thesis by simp + next + case greater + moreover have "real (4 * k + 3) / 2 ^ Suc n - 1 / (2 ^ Suc n) = real j / 2 ^ n" + using k by (force simp: divide_simps) + moreover have "\real i / 2 ^ m - real j / 2 ^ n\ < 2 * 1 / (2 ^ Suc n)" + using less.prems by simp + ultimately have closer: "\real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\ < 1 / (2 ^ Suc n)" + using less.prems by linarith + have *: "a (real (4 * k + 3) / 2 ^ Suc n) \ c (real i / 2 ^ m) \ + c (real i / 2 ^ m) \ b (real (4 * k + 3) / 2 ^ Suc n)" + apply (rule less.IH [OF _ refl]) + using closer \n < m\ \d = m - n\ apply (auto simp: divide_simps \n < m\ diff_less_mono2) + done + show ?thesis + using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] + using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"] + using k a43 b43 * \0 < n\ + apply (simp add: add.commute) + done + qed + qed + qed + qed + then have aj_le_ci: "a (real j / 2 ^ n) \ c (real i / 2 ^ m)" + and ci_le_bj: "c (real i / 2 ^ m) \ b (real j / 2 ^ n)" if "odd j" "\real i / 2^m - real j / 2^n\ < 1/2 ^ n" for i j m n + using that by blast+ + have close_ab: "odd m \ \a (real m / 2 ^ n) - b (real m / 2 ^ n)\ \ 2 / 2^n" for m n + proof (induction n arbitrary: m) + case 0 + with u01 v01 show ?case by auto + next + case (Suc n m) + with oddE obtain k where k: "m = Suc (2*k)" by fastforce + show ?case + proof (cases "n > 0") + case False + with u01 v01 show ?thesis + by (simp add: a_def b_def leftrec_base rightrec_base) + next + case True + show ?thesis + proof (cases "even k") + case True + then obtain j where j: "k = 2*j" by (metis evenE) + have "\a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\ \ 2/2 ^ n" + proof - + have "odd (Suc k)" + using True by auto + then show ?thesis + by (metis (no_types) Groups.add_ac(2) Suc.IH j of_nat_Suc of_nat_mult of_nat_numeral) + qed + moreover have "a ((2 * real j + 1) / 2 ^ n) \ + leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))" + using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] + by (auto simp: add.commute left_right) + moreover have "leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \ + c ((2 * real j + 1) / 2 ^ n)" + using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] + by (auto simp: add.commute left_right_m) + ultimately have "\a ((2 * real j + 1) / 2 ^ n) - + leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))\ + \ 2/2 ^ Suc n" + by (simp add: c_def midpoint_def) + with j k \n > 0\ show ?thesis + by (simp add: add.commute [of 1] a41 b41 del: power_Suc) + next + case False + then obtain j where j: "k = 2*j + 1" by (metis oddE) + have "\a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\ \ 2/2 ^ n" + using Suc.IH [OF False] j by (auto simp: algebra_simps) + moreover have "c ((2 * real j + 1) / 2 ^ n) \ + rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))" + using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] + by (auto simp: add.commute left_right_m) + moreover have "rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \ + b ((2 * real j + 1) / 2 ^ n)" + using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] + by (auto simp: add.commute left_right) + ultimately have "\rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) - + b ((2 * real j + 1) / 2 ^ n)\ \ 2/2 ^ Suc n" + by (simp add: c_def midpoint_def) + with j k \n > 0\ show ?thesis + by (simp add: add.commute [of 3] a43 b43 del: power_Suc) + qed + qed + qed + have m1_to_3: "4 * real k - 1 = real (4 * (k-1)) + 3" if "0 < k" for k + using that by auto + have fb_eq_fa: "\0 < j; 2*j < 2 ^ n\ \ f(b((2 * real j - 1) / 2^n)) = f(a((2 * real j + 1) / 2^n))" for n j + proof (induction n arbitrary: j) + case 0 + then show ?case by auto + next + case (Suc n j) show ?case + proof (cases "n > 0") + case False + with Suc.prems show ?thesis by auto + next + case True + show ?thesis proof (cases "even j") + case True + then obtain k where k: "j = 2*k" by (metis evenE) + with \0 < j\ have "k > 0" "2 * k < 2 ^ n" + using Suc.prems(2) k by auto + with k \0 < n\ Suc.IH [of k] show ?thesis + apply (simp add: m1_to_3 a41 b43 del: power_Suc) + apply (subst of_nat_diff, auto) + done + next + case False + then obtain k where k: "j = 2*k + 1" by (metis oddE) + have "f (leftcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n))) + = f (c ((2 * k + 1) / 2^n))" + "f (c ((2 * k + 1) / 2^n)) + = f (rightcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))" + using alec [of "2*k+1" n] cleb [of "2*k+1" n] a_ge_0 [of "2*k+1" n] b_le_1 [of "2*k+1" n] k + using left_right_m [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] + apply (auto simp: add.commute feqm [OF order_refl] feqm [OF _ order_refl, symmetric]) + done + then + show ?thesis + by (simp add: k add.commute [of 1] add.commute [of 3] a43 b41\0 < n\ del: power_Suc) + qed + qed + qed + have f_eq_fc: "\0 < j; j < 2 ^ n\ + \ f(b((2*j - 1) / 2 ^ (Suc n))) = f(c(j / 2^n)) \ + f(a((2*j + 1) / 2 ^ (Suc n))) = f(c(j / 2^n))" for n and j::nat + proof (induction n arbitrary: j) + case 0 + then show ?case by auto + next + case (Suc n) + show ?case + proof (cases "even j") + case True + then obtain k where k: "j = 2*k" by (metis evenE) + then have less2n: "k < 2 ^ n" + using Suc.prems(2) by auto + have "0 < k" using \0 < j\ k by linarith + then have m1_to_3: "real (4 * k - Suc 0) = real (4 * (k-1)) + 3" + by auto + then show ?thesis + using Suc.IH [of k] k \0 < k\ + apply (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc) + apply (auto simp: of_nat_diff) + done + next + case False + then obtain k where k: "j = 2*k + 1" by (metis oddE) + with Suc.prems have "k < 2^n" by auto + show ?thesis + using alec [of "2*k+1" "Suc n"] cleb [of "2*k+1" "Suc n"] a_ge_0 [of "2*k+1" "Suc n"] b_le_1 [of "2*k+1" "Suc n"] k + using left_right_m [of "c((2*k + 1) / 2 ^ Suc n)" "a((2*k + 1) / 2 ^ Suc n)" "b((2*k + 1) / 2 ^ Suc n)"] + apply (simp add: add.commute [of 1] add.commute [of 3] m1_to_3 b41 a43 del: power_Suc) + apply (force intro: feqm) + done + qed + qed + define D01 where "D01 \ {0<..<1} \ (\k m. {real m / 2^k})" + have cloD01 [simp]: "closure D01 = {0..1}" + unfolding D01_def + by (subst closure_dyadic_rationals_in_convex_set_pos_1) auto + have "uniformly_continuous_on D01 (f \ c)" + proof (clarsimp simp: uniformly_continuous_on_def) + fix e::real + assume "0 < e" + have ucontf: "uniformly_continuous_on {0..1} f" + by (simp add: compact_uniformly_continuous [OF cont_f]) + then obtain d where "0 < d" and d: "\x x'. \x \ {0..1}; x' \ {0..1}; norm (x' - x) < d\ \ norm (f x' - f x) < e/2" + unfolding uniformly_continuous_on_def dist_norm + by (metis \0 < e\ less_divide_eq_numeral1(1) mult_zero_left) + obtain n where n: "1/2^n < min d 1" + by (metis \0 < d\ divide_less_eq_1 less_numeral_extra(1) min_def one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_numeral) + with gr0I have "n > 0" + by (force simp: divide_simps) + show "\d>0. \x\D01. \x'\D01. dist x' x < d \ dist (f (c x')) (f (c x)) < e" + proof (intro exI ballI impI conjI) + show "(0::real) < 1/2^n" by auto + next + have dist_fc_close: "dist (f(c(real i / 2^m))) (f(c(real j / 2^n))) < e/2" + if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and clo: "abs(i / 2^m - j / 2^n) < 1/2 ^ n" for i j m + proof - + have abs3: "\x - a\ < e \ x = a \ \x - (a - e/2)\ < e/2 \ \x - (a + e/2)\ < e/2" for x a e::real + by linarith + consider "i / 2 ^ m = j / 2 ^ n" + | "\i / 2 ^ m - (2 * j - 1) / 2 ^ Suc n\ < 1/2 ^ Suc n" + | "\i / 2 ^ m - (2 * j + 1) / 2 ^ Suc n\ < 1/2 ^ Suc n" + using abs3 [OF clo] j by (auto simp: field_simps of_nat_diff) + then show ?thesis + proof cases + case 1 with \0 < e\ show ?thesis by auto + next + case 2 + have *: "abs(a - b) \ 1/2 ^ n \ 1/2 ^ n < d \ a \ c \ c \ b \ b - c < d" for a b c + by auto + have "norm (c (real i / 2 ^ m) - b (real (2 * j - 1) / 2 ^ Suc n)) < d" + using 2 j n close_ab [of "2*j-1" "Suc n"] + using b_ge_0 [of "2*j-1" "Suc n"] b_le_1 [of "2*j-1" "Suc n"] + using aj_le_ci [of "2*j-1" i m "Suc n"] + using ci_le_bj [of "2*j-1" i m "Suc n"] + apply (simp add: divide_simps of_nat_diff del: power_Suc) + apply (auto simp: divide_simps intro!: *) + done + moreover have "f(c(j / 2^n)) = f(b ((2*j - 1) / 2 ^ (Suc n)))" + using f_eq_fc [OF j] by metis + ultimately show ?thesis + by (metis dist_norm atLeastAtMost_iff b_ge_0 b_le_1 c_ge_0 c_le_1 d) + next + case 3 + have *: "abs(a - b) \ 1/2 ^ n \ 1/2 ^ n < d \ a \ c \ c \ b \ c - a < d" for a b c + by auto + have "norm (c (real i / 2 ^ m) - a (real (2 * j + 1) / 2 ^ Suc n)) < d" + using 3 j n close_ab [of "2*j+1" "Suc n"] + using b_ge_0 [of "2*j+1" "Suc n"] b_le_1 [of "2*j+1" "Suc n"] + using aj_le_ci [of "2*j+1" i m "Suc n"] + using ci_le_bj [of "2*j+1" i m "Suc n"] + apply (simp add: divide_simps of_nat_diff del: power_Suc) + apply (auto simp: divide_simps intro!: *) + done + moreover have "f(c(j / 2^n)) = f(a ((2*j + 1) / 2 ^ (Suc n)))" + using f_eq_fc [OF j] by metis + ultimately show ?thesis + by (metis dist_norm a_ge_0 atLeastAtMost_iff a_ge_0 a_le_1 c_ge_0 c_le_1 d) + qed + qed + show "dist (f (c x')) (f (c x)) < e" + if "x \ D01" "x' \ D01" "dist x' x < 1/2^n" for x x' + using that unfolding D01_def dyadics_in_open_unit_interval + proof clarsimp + fix i k::nat and m p + assume i: "0 < i" "i < 2 ^ m" and k: "0 0" + by (auto simp: le_max_iff_disj) + then obtain j where "floor (max (2^n*i / 2^m) (2^n*k / 2^p)) = int j" + using zero_le_floor zero_le_imp_eq_int by blast + then have j_le: "real j \ max (2^n * i / 2^m) (2^n * k / 2^p)" + and less_j1: "max (2^n * i / 2^m) (2^n * k / 2^p) < real j + 1" + using floor_correct [of "max (2^n * i / 2^m) (2^n * k / 2^p)"] by linarith+ + show thesis + proof (cases "j = 0") + case True + show thesis + proof + show "(1::nat) < 2 ^ n" + apply (subst one_less_power) + using \n > 0\ by auto + show "\real i / 2 ^ m - real 1/2 ^ n\ < 1/2 ^ n" + using i less_j1 by (simp add: dist_norm field_simps True) + show "\real k / 2 ^ p - real 1/2 ^ n\ < 1/2 ^ n" + using k less_j1 by (simp add: dist_norm field_simps True) + qed simp + next + case False + have 1: "real j * 2 ^ m < real i * 2 ^ n" + if j: "real j * 2 ^ p \ real k * 2 ^ n" and k: "real k * 2 ^ m < real i * 2 ^ p" + for i k m p + proof - + have "real j * 2 ^ p * 2 ^ m \ real k * 2 ^ n * 2 ^ m" + using j by simp + moreover have "real k * 2 ^ m * 2 ^ n < real i * 2 ^ p * 2 ^ n" + using k by simp + ultimately have "real j * 2 ^ p * 2 ^ m < real i * 2 ^ p * 2 ^ n" + by (simp only: mult_ac) + then show ?thesis + by simp + qed + have 2: "real j * 2 ^ m < 2 ^ m + real i * 2 ^ n" + if j: "real j * 2 ^ p \ real k * 2 ^ n" and k: "real k * (2 ^ m * 2 ^ n) < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)" + for i k m p + proof - + have "real j * 2 ^ p * 2 ^ m \ real k * (2 ^ m * 2 ^ n)" + using j by simp + also have "... < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)" + by (rule k) + finally have "(real j * 2 ^ m) * 2 ^ p < (2 ^ m + real i * 2 ^ n) * 2 ^ p" + by (simp add: algebra_simps) + then show ?thesis + by simp + qed + have 3: "real j * 2 ^ p < 2 ^ p + real k * 2 ^ n" + if j: "real j * 2 ^ m \ real i * 2 ^ n" and i: "real i * 2 ^ p \ real k * 2 ^ m" + proof - + have "real j * 2 ^ m * 2 ^ p \ real i * 2 ^ n * 2 ^ p" + using j by simp + moreover have "real i * 2 ^ p * 2 ^ n \ real k * 2 ^ m * 2 ^ n" + using i by simp + ultimately have "real j * 2 ^ m * 2 ^ p \ real k * 2 ^ m * 2 ^ n" + by (simp only: mult_ac) + then have "real j * 2 ^ p \ real k * 2 ^ n" + by simp + also have "... < 2 ^ p + real k * 2 ^ n" + by auto + finally show ?thesis by simp + qed + show ?thesis + proof + have "real j < 2 ^ n" + using j_le i k + apply (auto simp: le_max_iff_disj simp del: real_of_nat_less_numeral_power_cancel_iff elim!: le_less_trans) + apply (auto simp: field_simps) + done + then show "j < 2 ^ n" + by auto + show "\real i / 2 ^ m - real j / 2 ^ n\ < 1/2 ^ n" + using clo less_j1 j_le + apply (auto simp: le_max_iff_disj divide_simps dist_norm) + apply (auto simp: algebra_simps abs_if split: if_split_asm dest: 1 2) + done + show "\real k / 2 ^ p - real j / 2 ^ n\ < 1/2 ^ n" + using clo less_j1 j_le + apply (auto simp: le_max_iff_disj divide_simps dist_norm) + apply (auto simp: algebra_simps not_less abs_if split: if_split_asm dest: 3 2) + done + qed (use False in simp) + qed + qed + show "dist (f (c (real k / 2 ^ p))) (f (c (real i / 2 ^ m))) < e" + proof (rule dist_triangle_half_l) + show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2" + apply (rule dist_fc_close) + using \0 < j\ \j < 2 ^ n\ k clo_kj by auto + show "dist (f (c (real i / 2 ^ m))) (f (c (real j / 2 ^ n))) < e/2" + apply (rule dist_fc_close) + using \0 < j\ \j < 2 ^ n\ i clo_ij by auto + qed + qed + qed + qed + then obtain h where ucont_h: "uniformly_continuous_on {0..1} h" + and fc_eq: "\x. x \ D01 \ (f \ c) x = h x" + proof (rule uniformly_continuous_on_extension_on_closure [of D01 "f \ c"]) + qed (use closure_subset [of D01] in \auto intro!: that\) + then have cont_h: "continuous_on {0..1} h" + using uniformly_continuous_imp_continuous by blast + have h_eq: "h (real k / 2 ^ m) = f (c (real k / 2 ^ m))" if "0 < k" "k < 2^m" for k m + using fc_eq that by (force simp: D01_def) + have "h ` {0..1} = f ` {0..1}" + proof + have "h ` (closure D01) \ f ` {0..1}" + proof (rule image_closure_subset) + show "continuous_on (closure D01) h" + using cont_h by simp + show "closed (f ` {0..1})" + using compact_continuous_image [OF cont_f] compact_imp_closed by blast + show "h ` D01 \ f ` {0..1}" + by (force simp: dyadics_in_open_unit_interval D01_def h_eq) + qed + with cloD01 show "h ` {0..1} \ f ` {0..1}" by simp + have a12 [simp]: "a (1/2) = u" + by (metis a_def leftrec_base numeral_One of_nat_numeral) + have b12 [simp]: "b (1/2) = v" + by (metis b_def rightrec_base numeral_One of_nat_numeral) + have "f ` {0..1} \ closure(h ` D01)" + proof (clarsimp simp: closure_approachable dyadics_in_open_unit_interval D01_def) + fix x e::real + assume "0 \ x" "x \ 1" "0 < e" + have ucont_f: "uniformly_continuous_on {0..1} f" + using compact_uniformly_continuous cont_f by blast + then obtain \ where "\ > 0" + and \: "\x x'. \x \ {0..1}; x' \ {0..1}; dist x' x < \\ \ norm (f x' - f x) < e" + using \0 < e\ by (auto simp: uniformly_continuous_on_def dist_norm) + have *: "\m::nat. \y. odd m \ 0 < m \ m < 2 ^ n \ y \ {a(m / 2^n) .. b(m / 2^n)} \ f y = f x" + if "n \ 0" for n + using that + proof (induction n) + case 0 then show ?case by auto + next + case (Suc n) + show ?case + proof (cases "n=0") + case True + consider "x \ {0..u}" | "x \ {u..v}" | "x \ {v..1}" + using \0 \ x\ \x \ 1\ by force + then have "\y\a (real 1/2). y \ b (real 1/2) \ f y = f x" + proof cases + case 1 + then show ?thesis + apply (rule_tac x=u in exI) + using uabv [of 1 1] f0u [of u] f0u [of x] by auto + next + case 2 + then show ?thesis + by (rule_tac x=x in exI) auto + next + case 3 + then show ?thesis + apply (rule_tac x=v in exI) + using uabv [of 1 1] fv1 [of v] fv1 [of x] by auto + qed + with \n=0\ show ?thesis + by (rule_tac x=1 in exI) auto + next + case False + with Suc obtain m y + where "odd m" "0 < m" and mless: "m < 2 ^ n" + and y: "y \ {a (real m / 2 ^ n)..b (real m / 2 ^ n)}" and feq: "f y = f x" + by metis + then obtain j where j: "m = 2*j + 1" by (metis oddE) + consider "y \ {a((2*j + 1) / 2^n) .. b((4*j + 1) / 2 ^ (Suc n))}" + | "y \ {b((4*j + 1) / 2 ^ (Suc n)) .. a((4*j + 3) / 2 ^ (Suc n))}" + | "y \ {a((4*j + 3) / 2 ^ (Suc n)) .. b((2*j + 1) / 2^n)}" + using y j by force + then show ?thesis + proof cases + case 1 + then show ?thesis + apply (rule_tac x="4*j + 1" in exI) + apply (rule_tac x=y in exI) + using mless j \n \ 0\ + apply (simp add: feq a41 b41 add.commute [of 1] del: power_Suc) + apply (simp add: algebra_simps) + done + next + case 2 + show ?thesis + apply (rule_tac x="4*j + 1" in exI) + apply (rule_tac x="b((4*j + 1) / 2 ^ (Suc n))" in exI) + using mless \n \ 0\ 2 j + using alec [of "2*j+1" n] cleb [of "2*j+1" n] a_ge_0 [of "2*j+1" n] b_le_1 [of "2*j+1" n] + using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^n)" "b((2*j + 1) / 2^n)"] + apply (simp add: a41 b41 a43 b43 add.commute [of 1] add.commute [of 3] del: power_Suc) + apply (auto simp: feq [symmetric] intro: f_eqI) + done + next + case 3 + then show ?thesis + apply (rule_tac x="4*j + 3" in exI) + apply (rule_tac x=y in exI) + using mless j \n \ 0\ + apply (simp add: feq a43 b43 del: power_Suc) + apply (simp add: algebra_simps) + done + qed + qed + qed + obtain n where n: "1/2^n < min (\ / 2) 1" + by (metis \0 < \\ divide_less_eq_1 less_numeral_extra(1) min_less_iff_conj one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral) + with gr0I have "n \ 0" + by fastforce + with * obtain m::nat and y + where "odd m" "0 < m" and mless: "m < 2 ^ n" + and y: "y \ {a(m / 2^n) .. b(m / 2^n)}" and feq: "f x = f y" + by metis + then have "0 \ y" "y \ 1" + by (metis atLeastAtMost_iff a_ge_0 b_le_1 order.trans)+ + moreover have "y < \ + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \ + y" + using y apply simp_all + using alec [of m n] cleb [of m n] n real_sum_of_halves close_ab [OF \odd m\, of n] + by linarith+ + moreover note \0 < m\ mless \0 \ x\ \x \ 1\ + ultimately show "\k. \m\{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e" + apply (rule_tac x=n in exI) + apply (rule_tac x=m in bexI) + apply (auto simp: dist_norm h_eq feq \) + done + qed + also have "... \ h ` {0..1}" + apply (rule closure_minimal) + using compact_continuous_image [OF cont_h] compact_imp_closed by (auto simp: D01_def) + finally show "f ` {0..1} \ h ` {0..1}" . + qed + moreover have "inj_on h {0..1}" + proof - + have "u < v" + by (metis atLeastAtMost_iff f0u f_1not0 fv1 order.not_eq_order_implies_strict u01(1) u01(2) v01(1)) + have f_not_fu: "\x. \u < x; x \ v\ \ f x \ f u" + by (metis atLeastAtMost_iff f0u fu1 greaterThanAtMost_iff order_refl order_trans u01(1) v01(2)) + have f_not_fv: "\x. \u \ x; x < v\ \ f x \ f v" + by (metis atLeastAtMost_iff order_refl order_trans v01(2) atLeastLessThan_iff fuv fv1) + have a_less_b: + "a(j / 2^n) < b(j / 2^n) \ + (\x. a(j / 2^n) < x \ x \ b(j / 2^n) \ f x \ f(a(j / 2^n))) \ + (\x. a(j / 2^n) \ x \ x < b(j / 2^n) \ f x \ f(b(j / 2^n)))" for n and j::nat + proof (induction n arbitrary: j) + case 0 then show ?case + by (simp add: \u < v\ f_not_fu f_not_fv) + next + case (Suc n j) show ?case + proof (cases "n > 0") + case False then show ?thesis + by (auto simp: a_def b_def leftrec_base rightrec_base \u < v\ f_not_fu f_not_fv) + next + case True show ?thesis + proof (cases "even j") + case True + with \0 < n\ Suc.IH show ?thesis + by (auto elim!: evenE) + next + case False + then obtain k where k: "j = 2*k + 1" by (metis oddE) + then show ?thesis + proof (cases "even k") + case True + then obtain m where m: "k = 2*m" by (metis evenE) + have fleft: "f (leftcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = + f (c((2*m + 1) / 2^n))" + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto intro: f_eqI) + show ?thesis + proof (intro conjI impI notI allI) + have False if "b (real j / 2 ^ Suc n) \ a (real j / 2 ^ Suc n)" + proof - + have "f (c ((1 + real m * 2) / 2 ^ n)) = f (a ((1 + real m * 2) / 2 ^ n))" + using k m \0 < n\ fleft that a41 [of n m] b41 [of n m] + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto simp: algebra_simps) + moreover have "a (real (1 + m * 2) / 2 ^ n) < c (real (1 + m * 2) / 2 ^ n)" + using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def) + moreover have "c (real (1 + m * 2) / 2 ^ n) \ b (real (1 + m * 2) / 2 ^ n)" + using cleb by blast + ultimately show ?thesis + using Suc.IH [of "1 + m * 2"] by force + qed + then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force + next + fix x + assume "a (real j / 2 ^ Suc n) < x" "x \ b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))" + then show False + using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct1] + using k m \0 < n\ a41 [of n m] b41 [of n m] + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto simp: algebra_simps) + next + fix x + assume "a (real j / 2 ^ Suc n) \ x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))" + then show False + using k m \0 < n\ a41 [of n m] b41 [of n m] fleft left_neq + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + by (auto simp: algebra_simps) + qed + next + case False + with oddE obtain m where m: "k = Suc (2*m)" by fastforce + have fright: "f (rightcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = f (c((2*m + 1) / 2^n))" + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto intro: f_eqI [OF _ order_refl]) + show ?thesis + proof (intro conjI impI notI allI) + have False if "b (real j / 2 ^ Suc n) \ a (real j / 2 ^ Suc n)" + proof - + have "f (c ((1 + real m * 2) / 2 ^ n)) = f (b ((1 + real m * 2) / 2 ^ n))" + using k m \0 < n\ fright that a43 [of n m] b43 [of n m] + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto simp: algebra_simps) + moreover have "a (real (1 + m * 2) / 2 ^ n) \ c (real (1 + m * 2) / 2 ^ n)" + using alec by blast + moreover have "c (real (1 + m * 2) / 2 ^ n) < b (real (1 + m * 2) / 2 ^ n)" + using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def) + ultimately show ?thesis + using Suc.IH [of "1 + m * 2"] by force + qed + then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force + next + fix x + assume "a (real j / 2 ^ Suc n) < x" "x \ b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))" + then show False + using k m \0 < n\ a43 [of n m] b43 [of n m] fright right_neq + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + by (auto simp: algebra_simps) + next + fix x + assume "a (real j / 2 ^ Suc n) \ x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))" + then show False + using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct2] + using k m \0 < n\ a43 [of n m] b43 [of n m] + using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] + using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] + by (auto simp: algebra_simps fright simp del: power_Suc) + qed + qed + qed + qed + qed + have c_gt_0 [simp]: "0 < c(m / 2^n)" and c_less_1 [simp]: "c(m / 2^n) < 1" for m::nat and n + using a_less_b [of m n] apply (simp_all add: c_def midpoint_def) + using a_ge_0 [of m n] b_le_1 [of m n] apply linarith+ + done + have approx: "\j n. odd j \ n \ 0 \ + real i / 2^m \ real j / 2^n \ + real j / 2^n \ real k / 2^p \ + \real i / 2 ^ m - real j / 2 ^ n\ < 1/2^n \ + \real k / 2 ^ p - real j / 2 ^ n\ < 1/2^n" + if "0 < i" "i < 2 ^ m" "0 < k" "k < 2 ^ p" "i / 2^m < k / 2^p" "m + p = N" for N m p i k + using that + proof (induction N arbitrary: m p i k rule: less_induct) + case (less N) + then consider "i / 2^m \ 1/2" "1/2 \ k / 2^p" | "k / 2^p < 1/2" | "k / 2^p \ 1/2" "1/2 < i / 2^m" + by linarith + then show ?case + proof cases + case 1 + with less.prems show ?thesis + by (rule_tac x=1 in exI)+ (fastforce simp: divide_simps) + next + case 2 show ?thesis + proof (cases m) + case 0 with less.prems show ?thesis + by auto + next + case (Suc m') show ?thesis + proof (cases p) + case 0 with less.prems show ?thesis by auto + next + case (Suc p') + have False if "real i * 2 ^ p' < real k * 2 ^ m'" "k < 2 ^ p'" "2 ^ m' \ i" + proof - + have "real k * 2 ^ m' < 2 ^ p' * 2 ^ m'" + using that by simp + then have "real i * 2 ^ p' < 2 ^ p' * 2 ^ m'" + using that by linarith + with that show ?thesis by simp + qed + then show ?thesis + using less.IH [of "m'+p'" i m' k p'] less.prems \m = Suc m'\ 2 Suc + apply atomize + apply (force simp: divide_simps) + done + qed + qed + next + case 3 show ?thesis + proof (cases m) + case 0 with less.prems show ?thesis + by auto + next + case (Suc m') show ?thesis + proof (cases p) + case 0 with less.prems show ?thesis by auto + next + case (Suc p') + then show ?thesis + using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \m = Suc m'\ Suc 3 + apply atomize + apply (auto simp: field_simps of_nat_diff) + apply (rule_tac x="2 ^ n + j" in exI, simp) + apply (rule_tac x="Suc n" in exI) + apply (auto simp: field_simps) + done + qed + qed + qed + qed + have clec: "c(real i / 2^m) \ c(real j / 2^n)" + if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and ij: "i / 2^m < j / 2^n" for m i n j + proof - + obtain j' n' where "odd j'" "n' \ 0" + and i_le_j: "real i / 2 ^ m \ real j' / 2 ^ n'" + and j_le_j: "real j' / 2 ^ n' \ real j / 2 ^ n" + and clo_ij: "\real i / 2 ^ m - real j' / 2 ^ n'\ < 1/2 ^ n'" + and clo_jj: "\real j / 2 ^ n - real j' / 2 ^ n'\ < 1/2 ^ n'" + using approx [of i m j n "m+n"] that i j ij by auto + with oddE obtain q where q: "j' = Suc (2*q)" by fastforce + have "c (real i / 2 ^ m) \ c((2*q + 1) / 2^n')" + proof (cases "i / 2^m = (2*q + 1) / 2^n'") + case True then show ?thesis by simp + next + case False + with i_le_j q have less: "i / 2^m < (2*q + 1) / 2^n'" + by auto + have *: "\i < q; abs(i - q) < s*2; q = r + s\ \ abs(i - r) < s" for i q s r::real + by auto + have "c(i / 2^m) \ b(real(4 * q + 1) / 2 ^ (Suc n'))" + apply (rule ci_le_bj, force) + apply (rule * [OF less]) + using i_le_j clo_ij q apply (auto simp: divide_simps) + done + then show ?thesis + using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] b41 [of n' q] \n' \ 0\ + using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"] + by (auto simp: algebra_simps) + qed + also have "... \ c(real j / 2^n)" + proof (cases "j / 2^n = (2*q + 1) / 2^n'") + case True + then show ?thesis by simp + next + case False + with j_le_j q have less: "(2*q + 1) / 2^n' < j / 2^n" + by auto + have *: "\q < i; abs(i - q) < s*2; r = q + s\ \ abs(i - r) < s" for i q s r::real + by auto + have "a(real(4*q + 3) / 2 ^ (Suc n')) \ c(j / 2^n)" + apply (rule aj_le_ci, force) + apply (rule * [OF less]) + using j_le_j clo_jj q apply (auto simp: divide_simps) + done + then show ?thesis + using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] a43 [of n' q] \n' \ 0\ + using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"] + by (auto simp: algebra_simps) + qed + finally show ?thesis . + qed + have "x = y" if "0 \ x" "x \ 1" "0 \ y" "y \ 1" "h x = h y" for x y + using that + proof (induction x y rule: linorder_class.linorder_less_wlog) + case (less x1 x2) + obtain m n where m: "0 < m" "m < 2 ^ n" + and x12: "x1 < m / 2^n" "m / 2^n < x2" + and neq: "h x1 \ h (real m / 2^n)" + proof - + have "(x1 + x2) / 2 \ closure D01" + using cloD01 less.hyps less.prems by auto + with less obtain y where "y \ D01" and dist_y: "dist y ((x1 + x2) / 2) < (x2 - x1) / 64" + unfolding closure_approachable + by (metis diff_gt_0_iff_gt less_divide_eq_numeral1(1) mult_zero_left) + obtain m n where m: "0 < m" "m < 2 ^ n" + and clo: "\real m / 2 ^ n - (x1 + x2) / 2\ < (x2 - x1) / 64" + and n: "1/2^n < (x2 - x1) / 128" + proof - + have "min 1 ((x2 - x1) / 128) > 0" "1/2 < (1::real)" + using less by auto + then obtain N where N: "1/2^N < min 1 ((x2 - x1) / 128)" + by (metis power_one_over real_arch_pow_inv) + then have "N > 0" + using less_divide_eq_1 by force + obtain p q where p: "p < 2 ^ q" "p \ 0" and yeq: "y = real p / 2 ^ q" + using \y \ D01\ by (auto simp: zero_less_divide_iff D01_def) + show ?thesis + proof + show "0 < 2^N * p" + using p by auto + show "2 ^ N * p < 2 ^ (N+q)" + by (simp add: p power_add) + have "\real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\ = \real p / 2 ^ q - (x1 + x2) / 2\" + by (simp add: power_add) + also have "... = \y - (x1 + x2) / 2\" + by (simp add: yeq) + also have "... < (x2 - x1) / 64" + using dist_y by (simp add: dist_norm) + finally show "\real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\ < (x2 - x1) / 64" . + have "(1::real) / 2 ^ (N + q) \ 1/2^N" + by (simp add: field_simps) + also have "... < (x2 - x1) / 128" + using N by force + finally show "1/2 ^ (N + q) < (x2 - x1) / 128" . + qed + qed + obtain m' n' m'' n'' where "0 < m'" "m' < 2 ^ n'" "x1 < m' / 2^n'" "m' / 2^n' < x2" + and "0 < m''" "m'' < 2 ^ n''" "x1 < m'' / 2^n''" "m'' / 2^n'' < x2" + and neq: "h (real m'' / 2^n'') \ h (real m' / 2^n')" + proof + show "0 < Suc (2*m)" + by simp + show m21: "Suc (2*m) < 2 ^ Suc n" + using m by auto + show "x1 < real (Suc (2 * m)) / 2 ^ Suc n" + using clo by (simp add: field_simps abs_if split: if_split_asm) + show "real (Suc (2 * m)) / 2 ^ Suc n < x2" + using n clo by (simp add: field_simps abs_if split: if_split_asm) + show "0 < 4*m + 3" + by simp + have "m+1 \ 2 ^ n" + using m by simp + then have "4 * (m+1) \ 4 * (2 ^ n)" + by simp + then show m43: "4*m + 3 < 2 ^ (n+2)" + by (simp add: algebra_simps) + show "x1 < real (4 * m + 3) / 2 ^ (n + 2)" + using clo by (simp add: field_simps abs_if split: if_split_asm) + show "real (4 * m + 3) / 2 ^ (n + 2) < x2" + using n clo by (simp add: field_simps abs_if split: if_split_asm) + have c_fold: "midpoint (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) = c ((2 * real m + 1) / 2 ^ Suc n)" + by (simp add: c_def) + define R where "R \ rightcut (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) (c ((2 * real m + 1) / 2 ^ Suc n))" + have "R < b ((2 * real m + 1) / 2 ^ Suc n)" + unfolding R_def using a_less_b [of "4*m + 3" "n+2"] a43 [of "Suc n" m] b43 [of "Suc n" m] + by simp + then have Rless: "R < midpoint R (b ((2 * real m + 1) / 2 ^ Suc n))" + by (simp add: midpoint_def) + have midR_le: "midpoint R (b ((2 * real m + 1) / 2 ^ Suc n)) \ b ((2 * real m + 1) / (2 * 2 ^ n))" + using \R < b ((2 * real m + 1) / 2 ^ Suc n)\ + by (simp add: midpoint_def) + have "(real (Suc (2 * m)) / 2 ^ Suc n) \ D01" "real (4 * m + 3) / 2 ^ (n + 2) \ D01" + by (simp_all add: D01_def m21 m43 del: power_Suc of_nat_Suc of_nat_add add_2_eq_Suc') blast+ + then show "h (real (4 * m + 3) / 2 ^ (n + 2)) \ h (real (Suc (2 * m)) / 2 ^ Suc n)" + using a_less_b [of "4*m + 3" "n+2", THEN conjunct1] + using a43 [of "Suc n" m] b43 [of "Suc n" m] + using alec [of "2*m+1" "Suc n"] cleb [of "2*m+1" "Suc n"] a_ge_0 [of "2*m+1" "Suc n"] b_le_1 [of "2*m+1" "Suc n"] + apply (simp add: fc_eq [symmetric] c_def del: power_Suc) + apply (simp only: add.commute [of 1] c_fold R_def [symmetric]) + apply (rule right_neq) + using Rless apply (simp add: R_def) + apply (rule midR_le, auto) + done + qed + then show ?thesis by (metis that) + qed + have m_div: "0 < m / 2^n" "m / 2^n < 1" + using m by (auto simp: divide_simps) + have closure0m: "{0..m / 2^n} = closure ({0<..< m / 2^n} \ (\k m. {real m / 2 ^ k}))" + by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m) + have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 1} \ (\k m. {real m / 2 ^ k}))" + apply (subst closure_dyadic_rationals_in_convex_set_pos_1; simp add: not_le m) + using \0 < real m / 2 ^ n\ by linarith + have cont_h': "continuous_on (closure ({u<.. (\k m. {real m / 2 ^ k}))) h" + if "0 \ u" "v \ 1" for u v + apply (rule continuous_on_subset [OF cont_h]) + apply (rule closure_minimal [OF subsetI]) + using that apply auto + done + have closed_f': "closed (f ` {u..v})" if "0 \ u" "v \ 1" for u v + by (metis compact_continuous_image cont_f compact_interval atLeastatMost_subset_iff + compact_imp_closed continuous_on_subset that) + have less_2I: "\k i. real i / 2 ^ k < 1 \ i < 2 ^ k" + by simp + have "h ` ({0<.. (\q p. {real p / 2 ^ q})) \ f ` {0..c (m / 2 ^ n)}" + proof clarsimp + fix p q + assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n" + then have [simp]: "0 < p" "p < 2 ^ q" + apply (simp add: divide_simps) + apply (blast intro: p less_2I m_div less_trans) + done + have "f (c (real p / 2 ^ q)) \ f ` {0..c (real m / 2 ^ n)}" + by (auto simp: clec p m) + then show "h (real p / 2 ^ q) \ f ` {0..c (real m / 2 ^ n)}" + by (simp add: h_eq) + qed + then have "h ` {0 .. m / 2^n} \ f ` {0 .. c(m / 2^n)}" + apply (subst closure0m) + apply (rule image_closure_subset [OF cont_h' closed_f']) + using m_div apply auto + done + then have hx1: "h x1 \ f ` {0 .. c(m / 2^n)}" + using x12 less.prems(1) by auto + then obtain t1 where t1: "h x1 = f t1" "0 \ t1" "t1 \ c (m / 2 ^ n)" + by auto + have "h ` ({m / 2 ^ n<..<1} \ (\q p. {real p / 2 ^ q})) \ f ` {c (m / 2 ^ n)..1}" + proof clarsimp + fix p q + assume p: "real m / 2 ^ n < real p / 2 ^ q" and [simp]: "p < 2 ^ q" + then have [simp]: "0 < p" + using gr_zeroI m_div by fastforce + have "f (c (real p / 2 ^ q)) \ f ` {c (m / 2 ^ n)..1}" + by (auto simp: clec p m) + then show "h (real p / 2 ^ q) \ f ` {c (real m / 2 ^ n)..1}" + by (simp add: h_eq) + qed + then have "h ` {m / 2^n .. 1} \ f ` {c(m / 2^n) .. 1}" + apply (subst closurem1) + apply (rule image_closure_subset [OF cont_h' closed_f']) + using m apply auto + done + then have hx2: "h x2 \ f ` {c(m / 2^n)..1}" + using x12 less.prems by auto + then obtain t2 where t2: "h x2 = f t2" "c (m / 2 ^ n) \ t2" "t2 \ 1" + by auto + with t1 less neq have False + using conn [of "h x2", unfolded is_interval_connected_1 [symmetric] is_interval_1, rule_format, of t1 t2 "c(m / 2^n)"] + by (simp add: h_eq m) + then show ?case by blast + qed auto + then show ?thesis + by (auto simp: inj_on_def) + qed + ultimately have "{0..1::real} homeomorphic f ` {0..1}" + using homeomorphic_compact [OF _ cont_h] by blast + then show ?thesis + using homeomorphic_sym by blast +qed + + +theorem path_contains_arc: + fixes p :: "real \ 'a::{complete_space,real_normed_vector}" + assumes "path p" and a: "pathstart p = a" and b: "pathfinish p = b" and "a \ b" + obtains q where "arc q" "path_image q \ path_image p" "pathstart q = a" "pathfinish q = b" +proof - + have ucont_p: "uniformly_continuous_on {0..1} p" + using \path p\ unfolding path_def + by (metis compact_Icc compact_uniformly_continuous) + define \ where "\ \ \S. S \ {0..1} \ 0 \ S \ 1 \ S \ + (\x \ S. \y \ S. open_segment x y \ S = {} \ p x = p y)" + obtain T where "closed T" "\ T" and T: "\U. \closed U; \ U\ \ \ (U \ T)" + proof (rule Brouwer_reduction_theorem_gen [of "{0..1}" \]) + have *: "{x<.. {0..1} = {x<.. x" "y \ 1" "x \ y" for x y::real + using that by auto + show "\ {0..1}" + by (auto simp: \_def open_segment_eq_real_ivl *) + show "\ (INTER UNIV F)" + if "\n. closed (F n)" and \: "\n. \ (F n)" and Fsub: "\n. F (Suc n) \ F n" for F + proof - + have F01: "\n. F n \ {0..1} \ 0 \ F n \ 1 \ F n" + and peq: "\n x y. \x \ F n; y \ F n; open_segment x y \ F n = {}\ \ p x = p y" + by (metis \ \_def)+ + have pqF: False if "\u. x \ F u" "\x. y \ F x" "open_segment x y \ (\x. F x) = {}" and neg: "p x \ p y" + for x y + using that + proof (induction x y rule: linorder_class.linorder_less_wlog) + case (less x y) + have xy: "x \ {0..1}" "y \ {0..1}" + by (metis less.prems subsetCE F01)+ + have "norm(p x - p y) / 2 > 0" + using less by auto + then obtain e where "e > 0" + and e: "\u v. \u \ {0..1}; v \ {0..1}; dist v u < e\ \ dist (p v) (p u) < norm(p x - p y) / 2" + by (metis uniformly_continuous_onE [OF ucont_p]) + have minxy: "min e (y - x) < (y - x) * (3 / 2)" + by (subst min_less_iff_disj) (simp add: less) + obtain w z where "w < z" and w: "w \ {x<.. {x<..0 < e\ less by simp_all + have Fclo: "\T. T \ range F \ closed T" + by (metis \\n. closed (F n)\ image_iff) + have eq: "{w..z} \ INTER UNIV F = {}" + using less w z apply (auto simp: open_segment_eq_real_ivl) + by (metis (no_types, hide_lams) INT_I IntI empty_iff greaterThanLessThan_iff not_le order.trans) + then obtain K where "finite K" and K: "{w..z} \ (\ (F ` K)) = {}" + by (metis finite_subset_image compact_imp_fip [OF compact_interval Fclo]) + then have "K \ {}" + using \w < z\ \{w..z} \ INTER K F = {}\ by auto + define n where "n \ Max K" + have "n \ K" unfolding n_def by (metis \K \ {}\ \finite K\ Max_in) + have "F n \ \ (F ` K)" + unfolding n_def by (metis Fsub Max_ge \K \ {}\ \finite K\ cINF_greatest lift_Suc_antimono_le) + with K have wzF_null: "{w..z} \ F n = {}" + by (metis disjoint_iff_not_equal subset_eq) + obtain u where u: "u \ F n" "u \ {x..w}" "({u..w} - {u}) \ F n = {}" + proof (cases "w \ F n") + case True + then show ?thesis + by (metis wzF_null \w < z\ atLeastAtMost_iff disjoint_iff_not_equal less_eq_real_def) + next + case False + obtain u where "u \ F n" "u \ {x..w}" "{u<.. F n = {}" + proof (rule segment_to_point_exists [of "F n \ {x..w}" w]) + show "closed (F n \ {x..w})" + by (metis \\n. closed (F n)\ closed_Int closed_real_atLeastAtMost) + show "F n \ {x..w} \ {}" + by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(1) less_eq_real_def w) + qed (auto simp: open_segment_eq_real_ivl intro!: that) + with False show thesis + apply (auto simp: disjoint_iff_not_equal intro!: that) + by (metis greaterThanLessThan_iff less_eq_real_def) + qed + obtain v where v: "v \ F n" "v \ {z..y}" "({z..v} - {v}) \ F n = {}" + proof (cases "z \ F n") + case True + have "z \ {w..z}" + using \w < z\ by auto + then show ?thesis + by (metis wzF_null Int_iff True empty_iff) + next + case False + show ?thesis + proof (rule segment_to_point_exists [of "F n \ {z..y}" z]) + show "closed (F n \ {z..y})" + by (metis \\n. closed (F n)\ closed_Int closed_atLeastAtMost) + show "F n \ {z..y} \ {}" + by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(2) less_eq_real_def z) + show "\b. \b \ F n \ {z..y}; open_segment z b \ (F n \ {z..y}) = {}\ \ thesis" + apply (rule that) + apply (auto simp: open_segment_eq_real_ivl) + by (metis DiffI Int_iff atLeastAtMost_diff_ends atLeastAtMost_iff atLeastatMost_empty_iff empty_iff insert_iff False) + qed + qed + obtain u v where "u \ {0..1}" "v \ {0..1}" "norm(u - x) < e" "norm(v - y) < e" "p u = p v" + proof + show "u \ {0..1}" "v \ {0..1}" + by (metis F01 \u \ F n\ \v \ F n\ subsetD)+ + show "norm(u - x) < e" "norm (v - y) < e" + using \u \ {x..w}\ \v \ {z..y}\ atLeastAtMost_iff real_norm_def wxe zye by auto + show "p u = p v" + proof (rule peq) + show "u \ F n" "v \ F n" + by (auto simp: u v) + have "False" if "\ \ F n" "u < \" "\ < v" for \ + proof - + have "\ \ {z..v}" + by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that v(3)) + moreover have "\ \ {w..z} \ F n" + by (metis equals0D wzF_null) + ultimately have "\ \ {u..w}" + using that by auto + then show ?thesis + by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that u(3)) + qed + moreover + have "\\ \ F n; v < \; \ < u\ \ False" for \ + using \u \ {x..w}\ \v \ {z..y}\ \w < z\ by simp + ultimately + show "open_segment u v \ F n = {}" + by (force simp: open_segment_eq_real_ivl) + qed + qed + then show ?case + using e [of x u] e [of y v] xy + apply (simp add: open_segment_eq_real_ivl dist_norm del: divide_const_simps) + by (metis dist_norm dist_triangle_half_r less_irrefl) + qed (auto simp: open_segment_commute) + show ?thesis + unfolding \_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that F01 subsetCE pqF) + qed + show "closed {0..1::real}" by auto + qed (meson \_def) + then have "T \ {0..1}" "0 \ T" "1 \ T" + and peq: "\x y. \x \ T; y \ T; open_segment x y \ T = {}\ \ p x = p y" + unfolding \_def by metis+ + then have "T \ {}" by auto + define h where "h \ \x. p(@y. y \ T \ open_segment x y \ T = {})" + have "p y = p z" if "y \ T" "z \ T" and xyT: "open_segment x y \ T = {}" and xzT: "open_segment x z \ T = {}" + for x y z + proof (cases "x \ T") + case True + with that show ?thesis by (metis \\ T\ \_def) + next + case False + have "insert x (open_segment x y \ open_segment x z) \ T = {}" + by (metis False Int_Un_distrib2 Int_insert_left Un_empty_right xyT xzT) + moreover have "open_segment y z \ T \ insert x (open_segment x y \ open_segment x z) \ T" + apply auto + by (metis greaterThanLessThan_iff less_eq_real_def less_le_trans linorder_neqE_linordered_idom open_segment_eq_real_ivl) + ultimately have "open_segment y z \ T = {}" + by blast + with that peq show ?thesis by metis + qed + then have h_eq_p_gen: "h x = p y" if "y \ T" "open_segment x y \ T = {}" for x y + using that unfolding h_def + by (metis (mono_tags, lifting) some_eq_ex) + then have h_eq_p: "\x. x \ T \ h x = p x" + by simp + have disjoint: "\x. \y. y \ T \ open_segment x y \ T = {}" + by (meson \T \ {}\ \closed T\ segment_to_point_exists) + have heq: "h x = h x'" if "open_segment x x' \ T = {}" for x x' + proof (cases "x \ T \ x' \ T") + case True + then show ?thesis + by (metis h_eq_p h_eq_p_gen open_segment_commute that) + next + case False + obtain y y' where "y \ T" "open_segment x y \ T = {}" "h x = p y" + "y' \ T" "open_segment x' y' \ T = {}" "h x' = p y'" + by (meson disjoint h_eq_p_gen) + moreover have "open_segment y y' \ (insert x (insert x' (open_segment x y \ open_segment x' y' \ open_segment x x')))" + by (auto simp: open_segment_eq_real_ivl) + ultimately show ?thesis + using False that by (fastforce simp add: h_eq_p intro!: peq) + qed + have "h ` {0..1} homeomorphic {0..1::real}" + proof (rule homeomorphic_monotone_image_interval) + show "continuous_on {0..1} h" + proof (clarsimp simp add: continuous_on_iff) + fix u \::real + assume "0 < \" "0 \ u" "u \ 1" + then obtain \ where "\ > 0" and \: "\v. v \ {0..1} \ dist v u < \ \ dist (p v) (p u) < \ / 2" + using ucont_p [unfolded uniformly_continuous_on_def] + by (metis atLeastAtMost_iff half_gt_zero_iff) + then have "dist (h v) (h u) < \" if "v \ {0..1}" "dist v u < \" for v + proof (cases "open_segment u v \ T = {}") + case True + then show ?thesis + using \0 < \\ heq by auto + next + case False + have uvT: "closed (closed_segment u v \ T)" "closed_segment u v \ T \ {}" + using False open_closed_segment by (auto simp: \closed T\ closed_Int) + obtain w where "w \ T" and w: "w \ closed_segment u v" "open_segment u w \ T = {}" + apply (rule segment_to_point_exists [OF uvT, of u]) + by (metis IntD1 Int_commute Int_left_commute ends_in_segment(1) inf.orderE subset_oc_segment) + then have puw: "dist (p u) (p w) < \ / 2" + by (metis (no_types) \T \ {0..1}\ \dist v u < \\ \ dist_commute dist_in_closed_segment le_less_trans subsetCE) + obtain z where "z \ T" and z: "z \ closed_segment u v" "open_segment v z \ T = {}" + apply (rule segment_to_point_exists [OF uvT, of v]) + by (metis IntD2 Int_commute Int_left_commute ends_in_segment(2) inf.orderE subset_oc_segment) + then have "dist (p u) (p z) < \ / 2" + by (metis \T \ {0..1}\ \dist v u < \\ \ dist_commute dist_in_closed_segment le_less_trans subsetCE) + then show ?thesis + using puw by (metis (no_types) \w \ T\ \z \ T\ dist_commute dist_triangle_half_l h_eq_p_gen w(2) z(2)) + qed + with \0 < \\ show "\\>0. \v\{0..1}. dist v u < \ \ dist (h v) (h u) < \" by blast + qed + show "connected ({0..1} \ h -` {z})" for z + proof (clarsimp simp add: connected_iff_connected_component) + fix u v + assume huv_eq: "h v = h u" and uv: "0 \ u" "u \ 1" "0 \ v" "v \ 1" + have "\T. connected T \ T \ {0..1} \ T \ h -` {h u} \ u \ T \ v \ T" + proof (intro exI conjI) + show "connected (closed_segment u v)" + by simp + show "closed_segment u v \ {0..1}" + by (simp add: uv closed_segment_eq_real_ivl) + have pxy: "p x = p y" + if "T \ {0..1}" "0 \ T" "1 \ T" "x \ T" "y \ T" + and disjT: "open_segment x y \ (T - open_segment u v) = {}" + and xynot: "x \ open_segment u v" "y \ open_segment u v" + for x y + proof (cases "open_segment x y \ open_segment u v = {}") + case True + then show ?thesis + by (metis Diff_Int_distrib Diff_empty peq disjT \x \ T\ \y \ T\) + next + case False + then have "open_segment x u \ open_segment y v \ open_segment x y - open_segment u v \ + open_segment y u \ open_segment x v \ open_segment x y - open_segment u v" (is "?xuyv \ ?yuxv") + using xynot by (fastforce simp add: open_segment_eq_real_ivl not_le not_less split: if_split_asm) + then show "p x = p y" + proof + assume "?xuyv" + then have "open_segment x u \ T = {}" "open_segment y v \ T = {}" + using disjT by auto + then have "h x = h y" + using heq huv_eq by auto + then show ?thesis + using h_eq_p \x \ T\ \y \ T\ by auto + next + assume "?yuxv" + then have "open_segment y u \ T = {}" "open_segment x v \ T = {}" + using disjT by auto + then have "h x = h y" + using heq [of y u] heq [of x v] huv_eq by auto + then show ?thesis + using h_eq_p \x \ T\ \y \ T\ by auto + qed + qed + have "\ T - open_segment u v \ T" + proof (rule T) + show "closed (T - open_segment u v)" + by (simp add: closed_Diff [OF \closed T\] open_segment_eq_real_ivl) + have "0 \ open_segment u v" "1 \ open_segment u v" + using open_segment_eq_real_ivl uv by auto + then show "\ (T - open_segment u v)" + using \T \ {0..1}\ \0 \ T\ \1 \ T\ + by (auto simp: \_def) (meson peq pxy) + qed + then have "open_segment u v \ T = {}" + by blast + then show "closed_segment u v \ h -` {h u}" + by (force intro: heq simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl split: if_split_asm)+ + qed auto + then show "connected_component ({0..1} \ h -` {h u}) u v" + by (simp add: connected_component_def) + qed + show "h 1 \ h 0" + by (metis \\ T\ \_def a \a \ b\ b h_eq_p pathfinish_def pathstart_def) + qed + then obtain f and g :: "real \ 'a" + where gfeq: "(\x\h ` {0..1}. (g(f x) = x))" and fhim: "f ` h ` {0..1} = {0..1}" and contf: "continuous_on (h ` {0..1}) f" + and fgeq: "(\y\{0..1}. (f(g y) = y))" and pag: "path_image g = h ` {0..1}" and contg: "continuous_on {0..1} g" + by (auto simp: homeomorphic_def homeomorphism_def path_image_def) + then have "arc g" + by (metis arc_def path_def inj_on_def) + obtain u v where "u \ {0..1}" "a = g u" "v \ {0..1}" "b = g v" + by (metis (mono_tags, hide_lams) \\ T\ \_def a b fhim gfeq h_eq_p imageI path_image_def pathfinish_def pathfinish_in_path_image pathstart_def pathstart_in_path_image) + then have "a \ path_image g" "b \ path_image g" + using path_image_def by blast+ + have ph: "path_image h \ path_image p" + by (metis image_mono image_subset_iff path_image_def disjoint h_eq_p_gen \T \ {0..1}\) + show ?thesis + proof + show "pathstart (subpath u v g) = a" "pathfinish (subpath u v g) = b" + by (simp_all add: \a = g u\ \b = g v\) + show "path_image (subpath u v g) \ path_image p" + by (metis \arc g\ \u \ {0..1}\ \v \ {0..1}\ arc_imp_path order_trans pag path_image_def path_image_subpath_subset ph) + show "arc (subpath u v g)" + using \arc g\ \a = g u\ \b = g v\ \u \ {0..1}\ \v \ {0..1}\ arc_subpath_arc \a \ b\ by blast + qed +qed + + +corollary path_connected_arcwise: + fixes S :: "'a::{complete_space,real_normed_vector} set" + shows "path_connected S \ + (\x \ S. \y \ S. x \ y \ (\g. arc g \ path_image g \ S \ pathstart g = x \ pathfinish g = y))" + (is "?lhs = ?rhs") +proof (intro iffI impI ballI) + fix x y + assume "path_connected S" "x \ S" "y \ S" "x \ y" + then obtain p where p: "path p" "path_image p \ S" "pathstart p = x" "pathfinish p = y" + by (force simp: path_connected_def) + then show "\g. arc g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" + by (metis \x \ y\ order_trans path_contains_arc) +next + assume R [rule_format]: ?rhs + show ?lhs + unfolding path_connected_def + proof (intro ballI) + fix x y + assume "x \ S" "y \ S" + show "\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" + proof (cases "x = y") + case True with \x \ S\ path_component_def path_component_refl show ?thesis + by blast + next + case False with R [OF \x \ S\ \y \ S\] show ?thesis + by (auto intro: arc_imp_path) + qed + qed +qed + + +corollary arc_connected_trans: + fixes g :: "real \ 'a::{complete_space,real_normed_vector}" + assumes "arc g" "arc h" "pathfinish g = pathstart h" "pathstart g \ pathfinish h" + obtains i where "arc i" "path_image i \ path_image g \ path_image h" + "pathstart i = pathstart g" "pathfinish i = pathfinish h" + by (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join) + + + + +subsection\Accessibility of frontier points\ + +lemma dense_accessible_frontier_points: + fixes S :: "'a::{complete_space,real_normed_vector} set" + assumes "open S" and opeSV: "openin (subtopology euclidean (frontier S)) V" and "V \ {}" + obtains g where "arc g" "g ` {0..<1} \ S" "pathstart g \ S" "pathfinish g \ V" +proof - + obtain z where "z \ V" + using \V \ {}\ by auto + then obtain r where "r > 0" and r: "ball z r \ frontier S \ V" + by (metis openin_contains_ball opeSV) + then have "z \ frontier S" + using \z \ V\ opeSV openin_contains_ball by blast + then have "z \ closure S" "z \ S" + by (simp_all add: frontier_def assms interior_open) + with \r > 0\ have "infinite (S \ ball z r)" + by (auto simp: closure_def islimpt_eq_infinite_ball) + then obtain y where "y \ S" and y: "y \ ball z r" + using infinite_imp_nonempty by force + then have "y \ frontier S" + by (meson \open S\ disjoint_iff_not_equal frontier_disjoint_eq) + have "y \ z" + using \y \ S\ \z \ S\ by blast + have "path_connected(ball z r)" + by (simp add: convex_imp_path_connected) + with y \r > 0\ obtain g where "arc g" and pig: "path_image g \ ball z r" + and g: "pathstart g = y" "pathfinish g = z" + using \y \ z\ by (force simp: path_connected_arcwise) + have "compact (g -` frontier S \ {0..1})" + apply (simp add: compact_eq_bounded_closed bounded_Int bounded_closed_interval) + apply (rule closed_vimage_Int) + using \arc g\ apply (auto simp: arc_def path_def) + done + moreover have "g -` frontier S \ {0..1} \ {}" + proof - + have "\r. r \ g -` frontier S \ r \ {0..1}" + by (metis \z \ frontier S\ g(2) imageE path_image_def pathfinish_in_path_image vimageI2) + then show ?thesis + by blast + qed + ultimately obtain t where gt: "g t \ frontier S" and "0 \ t" "t \ 1" + and t: "\u. \g u \ frontier S; 0 \ u; u \ 1\ \ t \ u" + by (force simp: dest!: compact_attains_inf) + moreover have "t \ 0" + by (metis \y \ frontier S\ g(1) gt pathstart_def) + ultimately have t01: "0 < t" "t \ 1" + by auto + have "V \ frontier S" + using opeSV openin_contains_ball by blast + show ?thesis + proof + show "arc (subpath 0 t g)" + by (simp add: \0 \ t\ \t \ 1\ \arc g\ \t \ 0\ arc_subpath_arc) + have "g 0 \ S" + by (metis \y \ S\ g(1) pathstart_def) + then show "pathstart (subpath 0 t g) \ S" + by auto + have "g t \ V" + by (metis IntI atLeastAtMost_iff gt image_eqI path_image_def pig r subsetCE \0 \ t\ \t \ 1\) + then show "pathfinish (subpath 0 t g) \ V" + by auto + then have "inj_on (subpath 0 t g) {0..1}" + using t01 + apply (clarsimp simp: inj_on_def subpath_def) + apply (drule inj_onD [OF arc_imp_inj_on [OF \arc g\]]) + using mult_le_one apply auto + done + then have "subpath 0 t g ` {0..<1} \ subpath 0 t g ` {0..1} - {subpath 0 t g 1}" + by (force simp: dest: inj_onD) + moreover have False if "subpath 0 t g ` ({0..<1}) - S \ {}" + proof - + have contg: "continuous_on {0..1} g" + using \arc g\ by (auto simp: arc_def path_def) + have "subpath 0 t g ` {0..<1} \ frontier S \ {}" + proof (rule connected_Int_frontier [OF _ _ that]) + show "connected (subpath 0 t g ` {0..<1})" + apply (rule connected_continuous_image) + apply (simp add: subpath_def) + apply (intro continuous_intros continuous_on_compose2 [OF contg]) + apply (auto simp: \0 \ t\ \t \ 1\ mult_le_one) + done + show "subpath 0 t g ` {0..<1} \ S \ {}" + using \y \ S\ g(1) by (force simp: subpath_def image_def pathstart_def) + qed + then obtain x where "x \ subpath 0 t g ` {0..<1}" "x \ frontier S" + by blast + with t01 \0 \ t\ mult_le_one t show False + by (fastforce simp: subpath_def) + qed + then have "subpath 0 t g ` {0..1} - {subpath 0 t g 1} \ S" + using subsetD by fastforce + ultimately show "subpath 0 t g ` {0..<1} \ S" + by auto + qed +qed + + +lemma dense_accessible_frontier_points_connected: + fixes S :: "'a::{complete_space,real_normed_vector} set" + assumes "open S" "connected S" "x \ S" "V \ {}" + and ope: "openin (subtopology euclidean (frontier S)) V" + obtains g where "arc g" "g ` {0..<1} \ S" "pathstart g = x" "pathfinish g \ V" +proof - + have "V \ frontier S" + using ope openin_imp_subset by blast + with \open S\ \x \ S\ have "x \ V" + using interior_open by (auto simp: frontier_def) + obtain g where "arc g" and g: "g ` {0..<1} \ S" "pathstart g \ S" "pathfinish g \ V" + by (metis dense_accessible_frontier_points [OF \open S\ ope \V \ {}\]) + then have "path_connected S" + by (simp add: assms connected_open_path_connected) + with \pathstart g \ S\ \x \ S\ have "path_component S x (pathstart g)" + by (simp add: path_connected_component) + then obtain f where "path f" and f: "path_image f \ S" "pathstart f = x" "pathfinish f = pathstart g" + by (auto simp: path_component_def) + then have "path (f +++ g)" + by (simp add: \arc g\ arc_imp_path) + then obtain h where "arc h" + and h: "path_image h \ path_image (f +++ g)" "pathstart h = x" "pathfinish h = pathfinish g" + apply (rule path_contains_arc [of "f +++ g" x "pathfinish g"]) + using f \x \ V\ \pathfinish g \ V\ by auto + have "h ` {0..1} - {h 1} \ S" + using f g h apply (clarsimp simp: path_image_join) + apply (simp add: path_image_def pathfinish_def subset_iff image_def Bex_def) + by (metis le_less) + then have "h ` {0..<1} \ S" + using \arc h\ by (force simp: arc_def dest: inj_onD) + then show thesis + apply (rule that [OF \arc h\]) + using h \pathfinish g \ V\ by auto +qed + +lemma dense_access_fp_aux: + fixes S :: "'a::{complete_space,real_normed_vector} set" + assumes S: "open S" "connected S" + and opeSU: "openin (subtopology euclidean (frontier S)) U" + and opeSV: "openin (subtopology euclidean (frontier S)) V" + and "V \ {}" "\ U \ V" + obtains g where "arc g" "pathstart g \ U" "pathfinish g \ V" "g ` {0<..<1} \ S" +proof - + have "S \ {}" + using opeSV \V \ {}\ by (metis frontier_empty openin_subtopology_empty) + then obtain x where "x \ S" by auto + obtain g where "arc g" and g: "g ` {0..<1} \ S" "pathstart g = x" "pathfinish g \ V" + using dense_accessible_frontier_points_connected [OF S \x \ S\ \V \ {}\ opeSV] by blast + obtain h where "arc h" and h: "h ` {0..<1} \ S" "pathstart h = x" "pathfinish h \ U - {pathfinish g}" + proof (rule dense_accessible_frontier_points_connected [OF S \x \ S\]) + show "U - {pathfinish g} \ {}" + using \pathfinish g \ V\ \\ U \ V\ by blast + show "openin (subtopology euclidean (frontier S)) (U - {pathfinish g})" + by (simp add: opeSU openin_delete) + qed auto + obtain \ where "arc \" + and \: "path_image \ \ path_image (reversepath h +++ g)" + "pathstart \ = pathfinish h" "pathfinish \ = pathfinish g" + proof (rule path_contains_arc [of "(reversepath h +++ g)" "pathfinish h" "pathfinish g"]) + show "path (reversepath h +++ g)" + by (simp add: \arc g\ \arc h\ \pathstart g = x\ \pathstart h = x\ arc_imp_path) + show "pathstart (reversepath h +++ g) = pathfinish h" + "pathfinish (reversepath h +++ g) = pathfinish g" + by auto + show "pathfinish h \ pathfinish g" + using \pathfinish h \ U - {pathfinish g}\ by auto + qed auto + show ?thesis + proof + show "arc \" "pathstart \ \ U" "pathfinish \ \ V" + using \ \arc \\ \pathfinish h \ U - {pathfinish g}\ \pathfinish g \ V\ by auto + have "\ ` {0..1} - {\ 0, \ 1} \ S" + using \ g h + apply (simp add: path_image_join) + apply (simp add: path_image_def pathstart_def pathfinish_def subset_iff image_def Bex_def) + by (metis linorder_neqE_linordered_idom not_less) + then show "\ ` {0<..<1} \ S" + using \arc h\ \arc \\ + by (metis arc_imp_simple_path path_image_def pathfinish_def pathstart_def simple_path_endless) + qed +qed + +lemma dense_accessible_frontier_point_pairs: + fixes S :: "'a::{complete_space,real_normed_vector} set" + assumes S: "open S" "connected S" + and opeSU: "openin (subtopology euclidean (frontier S)) U" + and opeSV: "openin (subtopology euclidean (frontier S)) V" + and "U \ {}" "V \ {}" "U \ V" + obtains g where "arc g" "pathstart g \ U" "pathfinish g \ V" "g ` {0<..<1} \ S" +proof - + consider "\ U \ V" | "\ V \ U" + using \U \ V\ by blast + then show ?thesis + proof cases + case 1 then show ?thesis + using assms dense_access_fp_aux [OF S opeSU opeSV] that by blast + next + case 2 + obtain g where "arc g" and g: "pathstart g \ V" "pathfinish g \ U" "g ` {0<..<1} \ S" + using assms dense_access_fp_aux [OF S opeSV opeSU] "2" by blast + show ?thesis + proof + show "arc (reversepath g)" + by (simp add: \arc g\ arc_reversepath) + show "pathstart (reversepath g) \ U" "pathfinish (reversepath g) \ V" + using g by auto + show "reversepath g ` {0<..<1} \ S" + using g by (auto simp: reversepath_def) + qed + qed +qed + +end diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Jan 05 22:38:06 2017 +0100 @@ -2021,6 +2021,55 @@ done qed +lemma connected_sphere_eq: + fixes a :: "'a :: euclidean_space" + shows "connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" + (is "?lhs = ?rhs") +proof (cases r "0::real" rule: linorder_cases) + case less + then show ?thesis by auto +next + case equal + then show ?thesis by auto +next + case greater + show ?thesis + proof + assume L: ?lhs + have "False" if 1: "DIM('a) = 1" + proof - + obtain x y where xy: "sphere a r = {x,y}" "x \ y" + using sphere_1D_doubleton [OF 1 greater] + by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) + then have "finite (sphere a r)" + by auto + with L \r > 0\ show "False" + apply (auto simp: connected_finite_iff_sing) + using xy by auto + qed + with greater show ?rhs + by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) + next + assume ?rhs + then show ?lhs + using connected_sphere greater by auto + qed +qed + +lemma path_connected_sphere_eq: + fixes a :: "'a :: euclidean_space" + shows "path_connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + using connected_sphere_eq path_connected_imp_connected by blast +next + assume R: ?rhs + then show ?lhs + by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) +qed + proposition frontier_subset_retraction: fixes S :: "'a::euclidean_space set" assumes "bounded S" and fros: "frontier S \ T" @@ -2073,7 +2122,7 @@ ultimately show False by simp qed -subsection\Retractions\ +subsection\More Properties of Retractions\ lemma retraction: "retraction s t r \ @@ -2272,6 +2321,60 @@ using assms by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) +subsubsection\A few simple lemmas about deformation retracts\ + +lemma deformation_retract_imp_homotopy_eqv: + fixes S :: "'a::euclidean_space set" + assumes "homotopic_with (\x. True) S S id r" "retraction S T r" + shows "S homotopy_eqv T" + apply (simp add: homotopy_eqv_def) + apply (rule_tac x=r in exI) + using assms apply (simp add: retraction_def) + apply (rule_tac x=id in exI) + apply (auto simp: continuous_on_id) + apply (metis homotopic_with_symD) + by (metis continuous_on_id' homotopic_with_equal homotopic_with_symD id_apply image_id subset_refl) + +lemma deformation_retract: + fixes S :: "'a::euclidean_space set" + shows "(\r. homotopic_with (\x. True) S S id r \ retraction S T r) \ + T retract_of S \ (\f. homotopic_with (\x. True) S S id f \ f ` S \ T)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (auto simp: retract_of_def retraction_def) +next + assume ?rhs + then show ?lhs + apply (clarsimp simp add: retract_of_def retraction_def) + apply (rule_tac x=r in exI, simp) + apply (rule homotopic_with_trans, assumption) + apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) + apply (rule_tac Y=S in homotopic_compose_continuous_left) + apply (auto simp: homotopic_with_sym) + done +qed + +lemma deformation_retract_of_contractible_sing: + fixes S :: "'a::euclidean_space set" + assumes "contractible S" "a \ S" + obtains r where "homotopic_with (\x. True) S S id r" "retraction S {a} r" +proof - + have "{a} retract_of S" + by (simp add: \a \ S\) + moreover have "homotopic_with (\x. True) S S id (\x. a)" + using assms + apply (clarsimp simp add: contractible_def) + apply (rule homotopic_with_trans, assumption) + by (metis assms(1) contractible_imp_path_connected homotopic_constant_maps homotopic_with_sym homotopic_with_trans insert_absorb insert_not_empty path_component_mem(1) path_connected_component) + moreover have "(\x. a) ` S \ {a}" + by (simp add: image_subsetI) + ultimately show ?thesis + using that deformation_retract by metis +qed + + subsection\Punctured affine hulls, etc.\ lemma continuous_on_compact_surface_projection_aux: @@ -2493,6 +2596,23 @@ by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def rel_frontier_retract_of_punctured_affine_hull) +lemma homotopy_eqv_rel_frontier_punctured_convex: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "bounded S" "a \ rel_interior S" "convex T" "rel_frontier S \ T" "T \ affine hull S" + shows "(rel_frontier S) homotopy_eqv (T - {a})" + apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T]) + using assms + apply auto + apply (subst homotopy_eqv_sym) + using deformation_retract_imp_homotopy_eqv by blast + +lemma homotopy_eqv_rel_frontier_punctured_affine_hull: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "bounded S" "a \ rel_interior S" + shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})" +apply (rule homotopy_eqv_rel_frontier_punctured_convex) + using assms rel_frontier_affine_hull by force+ + lemma path_connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "path_connected(rel_frontier S)" @@ -3931,7 +4051,6 @@ assumes "AR S" "AR T" shows "AR(S \ T)" using assms by (simp add: AR_ANR ANR_Times contractible_Times) - lemma ENR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" @@ -3972,6 +4091,297 @@ assumes "bounded S" "convex S" shows "ANR(rel_frontier S)" by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms) + +lemma ENR_closedin_Un_local: + fixes S :: "'a::euclidean_space set" + shows "\ENR S; ENR T; ENR(S \ T); + closedin (subtopology euclidean (S \ T)) S; closedin (subtopology euclidean (S \ T)) T\ + \ ENR(S \ T)" +by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un) + +lemma ENR_closed_Un: + fixes S :: "'a::euclidean_space set" + shows "\closed S; closed T; ENR S; ENR T; ENR(S \ T)\ \ ENR(S \ T)" +by (auto simp: closed_subset ENR_closedin_Un_local) + +lemma absolute_retract_Un: + fixes S :: "'a::euclidean_space set" + shows "\S retract_of UNIV; T retract_of UNIV; (S \ T) retract_of UNIV\ + \ (S \ T) retract_of UNIV" + by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset) + +lemma retract_from_Un_Int: + fixes S :: "'a::euclidean_space set" + assumes clS: "closedin (subtopology euclidean (S \ T)) S" + and clT: "closedin (subtopology euclidean (S \ T)) T" + and Un: "(S \ T) retract_of U" and Int: "(S \ T) retract_of T" + shows "S retract_of U" +proof - + obtain r where r: "continuous_on T r" "r ` T \ S \ T" "\x\S \ T. r x = x" + using Int by (auto simp: retraction_def retract_of_def) + have "S retract_of S \ T" + unfolding retraction_def retract_of_def + proof (intro exI conjI) + show "continuous_on (S \ T) (\x. if x \ S then x else r x)" + apply (rule continuous_on_cases_local [OF clS clT]) + using r by (auto simp: continuous_on_id) + qed (use r in auto) + also have "... retract_of U" + by (rule Un) + finally show ?thesis . +qed + +lemma AR_from_Un_Int_local: + fixes S :: "'a::euclidean_space set" + assumes clS: "closedin (subtopology euclidean (S \ T)) S" + and clT: "closedin (subtopology euclidean (S \ T)) T" + and Un: "AR(S \ T)" and Int: "AR(S \ T)" + shows "AR S" + apply (rule AR_retract_of_AR [OF Un]) + by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2) + +lemma AR_from_Un_Int_local': + fixes S :: "'a::euclidean_space set" + assumes "closedin (subtopology euclidean (S \ T)) S" + and "closedin (subtopology euclidean (S \ T)) T" + and "AR(S \ T)" "AR(S \ T)" + shows "AR T" + using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute) + +lemma AR_from_Un_Int: + fixes S :: "'a::euclidean_space set" + assumes clo: "closed S" "closed T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" + shows "AR S" + by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) + +lemma ANR_from_Un_Int_local: + fixes S :: "'a::euclidean_space set" + assumes clS: "closedin (subtopology euclidean (S \ T)) S" + and clT: "closedin (subtopology euclidean (S \ T)) T" + and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" + shows "ANR S" +proof - + obtain V where clo: "closedin (subtopology euclidean (S \ T)) (S \ T)" + and ope: "openin (subtopology euclidean (S \ T)) V" + and ret: "S \ T retract_of V" + using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int) + then obtain r where r: "continuous_on V r" and rim: "r ` V \ S \ T" and req: "\x\S \ T. r x = x" + by (auto simp: retraction_def retract_of_def) + have Vsub: "V \ S \ T" + by (meson ope openin_contains_cball) + have Vsup: "S \ T \ V" + by (simp add: retract_of_imp_subset ret) + then have eq: "S \ V = ((S \ T) - T) \ V" + by auto + have eq': "S \ V = S \ (V \ T)" + using Vsub by blast + have "continuous_on (S \ V \ T) (\x. if x \ S then x else r x)" + proof (rule continuous_on_cases_local) + show "closedin (subtopology euclidean (S \ V \ T)) S" + using clS closedin_subset_trans inf.boundedE by blast + show "closedin (subtopology euclidean (S \ V \ T)) (V \ T)" + using clT Vsup by (auto simp: closedin_closed) + show "continuous_on (V \ T) r" + by (meson Int_lower1 continuous_on_subset r) + qed (use req continuous_on_id in auto) + with rim have "S retract_of S \ V" + unfolding retraction_def retract_of_def + apply (rule_tac x="\x. if x \ S then x else r x" in exI) + apply (auto simp: eq') + done + then show ?thesis + using ANR_neighborhood_retract [OF Un] + using \S \ V = S \ T - T \ V\ clT ope by fastforce +qed + +lemma ANR_from_Un_Int: + fixes S :: "'a::euclidean_space set" + assumes clo: "closed S" "closed T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" + shows "ANR S" + by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) + +proposition ANR_finite_Union_convex_closed: + fixes \ :: "'a::euclidean_space set set" + assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" + shows "ANR(\\)" +proof - + have "ANR(\\)" if "card \ < n" for n + using assms that + proof (induction n arbitrary: \) + case 0 then show ?case by simp + next + case (Suc n) + have "ANR(\\)" if "finite \" "\ \ \" for \ + using that + proof (induction \) + case empty + then show ?case by simp + next + case (insert C \) + have "ANR (C \ \\)" + proof (rule ANR_closed_Un) + show "ANR (C \ \\)" + unfolding Int_Union + proof (rule Suc) + show "finite (op \ C ` \)" + by (simp add: insert.hyps(1)) + show "\Ca. Ca \ op \ C ` \ \ closed Ca" + by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2) + show "\Ca. Ca \ op \ C ` \ \ convex Ca" + by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE) + show "card (op \ C ` \) < n" + proof - + have "card \ \ n" + by (meson Suc.prems(4) not_less not_less_eq) + then show ?thesis + by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less) + qed + qed + show "closed (\\)" + using Suc.prems(2) insert.hyps(1) insert.prems by blast + qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto) + then show ?case + by simp + qed + then show ?case + using Suc.prems(1) by blast + qed + then show ?thesis + by blast +qed + + +lemma finite_imp_ANR: + fixes S :: "'a::euclidean_space set" + assumes "finite S" + shows "ANR S" +proof - + have "ANR(\x \ S. {x})" + by (blast intro: ANR_finite_Union_convex_closed assms) + then show ?thesis + by simp +qed + +lemma ANR_insert: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "closed S" + shows "ANR(insert a S)" + by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un) + +lemma ANR_path_component_ANR: + fixes S :: "'a::euclidean_space set" + shows "ANR S \ ANR(path_component_set S x)" + using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast + +lemma ANR_connected_component_ANR: + fixes S :: "'a::euclidean_space set" + shows "ANR S \ ANR(connected_component_set S x)" + by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected) + +lemma ANR_component_ANR: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "c \ components S" + shows "ANR c" + by (metis ANR_connected_component_ANR assms componentsE) + +subsection\Original ANR material, now for ENRs.\ + +lemma ENR_bounded: + fixes S :: "'a::euclidean_space set" + assumes "bounded S" + shows "ENR S \ (\U. open U \ bounded U \ S retract_of U)" + (is "?lhs = ?rhs") +proof + obtain r where "0 < r" and r: "S \ ball 0 r" + using bounded_subset_ballD assms by blast + assume ?lhs + then show ?rhs + apply (clarsimp simp: ENR_def) + apply (rule_tac x="ball 0 r \ U" in exI, auto) + using r retract_of_imp_subset retract_of_subset by fastforce +next + assume ?rhs + then show ?lhs + using ENR_def by blast +qed + +lemma absolute_retract_imp_AR_gen: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "S retract_of T" "convex T" "T \ {}" "S homeomorphic S'" "closedin (subtopology euclidean U) S'" + shows "S' retract_of U" +proof - + have "AR T" + by (simp add: assms convex_imp_AR) + then have "AR S" + using AR_retract_of_AR assms by auto + then show ?thesis + using assms AR_imp_absolute_retract by metis +qed + +lemma absolute_retract_imp_AR: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'" + shows "S' retract_of UNIV" + using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast + +lemma homeomorphic_compact_arness: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "S homeomorphic S'" + shows "compact S \ S retract_of UNIV \ compact S' \ S' retract_of UNIV" + using assms homeomorphic_compactness + apply auto + apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+ + done + +lemma absolute_retract_from_Un_Int: + fixes S :: "'a::euclidean_space set" + assumes "(S \ T) retract_of UNIV" "(S \ T) retract_of UNIV" "closed S" "closed T" + shows "S retract_of UNIV" + using AR_from_Un_Int assms retract_of_UNIV by auto + +lemma ENR_from_Un_Int_gen: + fixes S :: "'a::euclidean_space set" + assumes "closedin (subtopology euclidean (S \ T)) S" "closedin (subtopology euclidean (S \ T)) T" "ENR(S \ T)" "ENR(S \ T)" + shows "ENR S" + apply (simp add: ENR_ANR) + using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast + + +lemma ENR_from_Un_Int: + fixes S :: "'a::euclidean_space set" + assumes "closed S" "closed T" "ENR(S \ T)" "ENR(S \ T)" + shows "ENR S" + by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2) + + +lemma ENR_finite_Union_convex_closed: + fixes \ :: "'a::euclidean_space set set" + assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" + shows "ENR(\ \)" + by (simp add: ENR_ANR ANR_finite_Union_convex_closed \ clo closed_Union closed_imp_locally_compact con) + +lemma finite_imp_ENR: + fixes S :: "'a::euclidean_space set" + shows "finite S \ ENR S" + by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact) + +lemma ENR_insert: + fixes S :: "'a::euclidean_space set" + assumes "closed S" "ENR S" + shows "ENR(insert a S)" +proof - + have "ENR ({a} \ S)" + by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right) + then show ?thesis + by auto +qed + +lemma ENR_path_component_ENR: + fixes S :: "'a::euclidean_space set" + assumes "ENR S" + shows "ENR(path_component_set S x)" + by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms + locally_path_connected_2 openin_subtopology_self path_component_eq_empty) (*UNUSED lemma ENR_Times: @@ -3983,6 +4393,61 @@ SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; *) +subsection\Finally, spheres are ANRs and ENRs\ + +lemma absolute_retract_homeomorphic_convex_compact: + fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set" + assumes "S homeomorphic U" "S \ {}" "S \ T" "convex U" "compact U" + shows "S retract_of T" + by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI) + +lemma frontier_retract_of_punctured_universe: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "bounded S" "a \ interior S" + shows "(frontier S) retract_of (- {a})" + using rel_frontier_retract_of_punctured_affine_hull + by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior) + +lemma sphere_retract_of_punctured_universe_gen: + fixes a :: "'a::euclidean_space" + assumes "b \ ball a r" + shows "sphere a r retract_of (- {b})" +proof - + have "frontier (cball a r) retract_of (- {b})" + apply (rule frontier_retract_of_punctured_universe) + using assms by auto + then show ?thesis + by simp +qed + +lemma sphere_retract_of_punctured_universe: + fixes a :: "'a::euclidean_space" + assumes "0 < r" + shows "sphere a r retract_of (- {a})" + by (simp add: assms sphere_retract_of_punctured_universe_gen) + +proposition ENR_sphere: + fixes a :: "'a::euclidean_space" + shows "ENR(sphere a r)" +proof (cases "0 < r") + case True + then have "sphere a r retract_of -{a}" + by (simp add: sphere_retract_of_punctured_universe) + with open_delete show ?thesis + by (auto simp: ENR_def) +next + case False + then show ?thesis + using finite_imp_ENR + by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) +qed + +corollary ANR_sphere: + fixes a :: "'a::euclidean_space" + shows "ANR(sphere a r)" + by (simp add: ENR_imp_ANR ENR_sphere) + + subsection\Borsuk homotopy extension theorem\ text\It's only this late so we can use the concept of retraction, @@ -4259,6 +4724,78 @@ then show ?thesis by blast qed +lemma homotopic_Borsuk_maps_in_bounded_component: + fixes a :: "'a :: euclidean_space" + assumes "compact S" and "a \ S"and "b \ S" + and boc: "bounded (connected_component_set (- S) a)" + and hom: "homotopic_with (\x. True) S (sphere 0 1) + (\x. (x - a) /\<^sub>R norm (x - a)) + (\x. (x - b) /\<^sub>R norm (x - b))" + shows "connected_component (- S) a b" +proof (rule ccontr) + assume notcc: "\ connected_component (- S) a b" + let ?T = "S \ connected_component_set (- S) a" + have "\g. continuous_on (S \ connected_component_set (- S) a) g \ + g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ + (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" + by (simp add: \a \ S\ componentsI non_extensible_Borsuk_map [OF \compact S\ _ boc]) + moreover obtain g where "continuous_on (S \ connected_component_set (- S) a) g" + "g ` (S \ connected_component_set (- S) a) \ sphere 0 1" + "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" + proof (rule Borsuk_homotopy_extension_homotopic) + show "closedin (subtopology euclidean ?T) S" + by (simp add: \compact S\ closed_subset compact_imp_closed) + show "continuous_on ?T (\x. (x - b) /\<^sub>R norm (x - b))" + by (simp add: \b \ S\ notcc continuous_on_Borsuk_map) + show "(\x. (x - b) /\<^sub>R norm (x - b)) ` ?T \ sphere 0 1" + by (simp add: \b \ S\ notcc Borsuk_map_into_sphere) + show "homotopic_with (\x. True) S (sphere 0 1) + (\x. (x - b) /\<^sub>R norm (x - b)) (\x. (x - a) /\<^sub>R norm (x - a))" + by (simp add: hom homotopic_with_symD) + qed (auto simp: ANR_sphere intro: that) + ultimately show False by blast +qed + + +lemma Borsuk_maps_homotopic_in_connected_component_eq: + fixes a :: "'a :: euclidean_space" + assumes S: "compact S" "a \ S" "b \ S" and 2: "2 \ DIM('a)" + shows "(homotopic_with (\x. True) S (sphere 0 1) + (\x. (x - a) /\<^sub>R norm (x - a)) + (\x. (x - b) /\<^sub>R norm (x - b)) \ + connected_component (- S) a b)" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + show ?rhs + proof (cases "bounded(connected_component_set (- S) a)") + case True + show ?thesis + by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L]) + next + case not_bo_a: False + show ?thesis + proof (cases "bounded(connected_component_set (- S) b)") + case True + show ?thesis + using homotopic_Borsuk_maps_in_bounded_component [OF S] + by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym) + next + case False + then show ?thesis + using cobounded_unique_unbounded_component [of "-S" a b] \compact S\ not_bo_a + by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq) + qed + qed +next + assume R: ?rhs + then have "path_component (- S) a b" + using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce + then show ?lhs + by (simp add: Borsuk_maps_homotopic_in_path_component) +qed + + subsection\The complement of a set and path-connectedness\ text\Complement in dimension N > 1 of set homeomorphic to any interval in diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Thu Jan 05 22:38:06 2017 +0100 @@ -577,9 +577,6 @@ subsection \Valid paths, and their start and finish\ -lemma Diff_Un_eq: "A - (B \ C) = A - B - C" - by blast - definition valid_path :: "(real \ 'a :: real_normed_vector) \ bool" where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Thu Jan 05 22:38:06 2017 +0100 @@ -3559,4 +3559,70 @@ shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) + +subsection\Homeomorphism of simple closed curves to circles\ + +proposition homeomorphic_simple_path_image_circle: + fixes a :: complex and \ :: "real \ 'a::t2_space" + assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" + shows "(path_image \) homeomorphic sphere a r" +proof - + have "homotopic_loops (path_image \) \ \" + by (simp add: assms homotopic_loops_refl simple_path_imp_path) + then have hom: "homotopic_with (\h. True) (sphere 0 1) (path_image \) + (\ \ (\z. Arg z / (2*pi))) (\ \ (\z. Arg z / (2*pi)))" + by (rule homotopic_loops_imp_homotopic_circlemaps) + have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg z / (2*pi))) g" + proof (rule homeomorphism_compact) + show "continuous_on (sphere 0 1) (\ \ (\z. Arg z / (2*pi)))" + using hom homotopic_with_imp_continuous by blast + show "inj_on (\ \ (\z. Arg z / (2*pi))) (sphere 0 1)" + proof + fix x y + assume xy: "x \ sphere 0 1" "y \ sphere 0 1" + and eq: "(\ \ (\z. Arg z / (2*pi))) x = (\ \ (\z. Arg z / (2*pi))) y" + then have "(Arg x / (2*pi)) = (Arg y / (2*pi))" + proof - + have "(Arg x / (2*pi)) \ {0..1}" "(Arg y / (2*pi)) \ {0..1}" + using Arg_ge_0 Arg_lt_2pi dual_order.strict_iff_order by fastforce+ + with eq show ?thesis + using \simple_path \\ Arg_lt_2pi unfolding simple_path_def o_def + by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) + qed + with xy show "x = y" + by (metis Arg Arg_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) + qed + have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg z / (2*pi)) = \ x" + by (metis Arg_ge_0 Arg_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) + moreover have "\z\sphere 0 1. \ x = \ (Arg z / (2*pi))" if "0 \ x" "x \ 1" for x + proof (cases "x=1") + case True + then show ?thesis + apply (rule_tac x=1 in bexI) + apply (metis loop Arg_of_real divide_eq_0_iff of_real_1 pathfinish_def pathstart_def \0 \ x\, auto) + done + next + case False + then have *: "(Arg (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" + using that by (auto simp: Arg_exp divide_simps) + show ?thesis + by (rule_tac x="exp(ii* of_real(2*pi*x))" in bexI) (auto simp: *) + qed + ultimately show "(\ \ (\z. Arg z / (2*pi))) ` sphere 0 1 = path_image \" + by (auto simp: path_image_def image_iff) + qed auto + then have "path_image \ homeomorphic sphere (0::complex) 1" + using homeomorphic_def homeomorphic_sym by blast + also have "... homeomorphic sphere a r" + by (simp add: assms homeomorphic_spheres) + finally show ?thesis . +qed + +lemma homeomorphic_simple_path_images: + fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" + assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" + assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" + shows "(path_image \1) homeomorphic (path_image \2)" + by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) + end diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Jan 05 22:38:06 2017 +0100 @@ -1314,21 +1314,6 @@ lemma fst_snd_linear: "linear (\(x,y). x + y)" unfolding linear_iff by (simp add: algebra_simps) -lemma scaleR_2: - fixes x :: "'a::real_vector" - shows "scaleR 2 x = x + x" - unfolding one_add_one [symmetric] scaleR_left_distrib by simp - -lemma scaleR_half_double [simp]: - fixes a :: "'a::real_normed_vector" - shows "(1 / 2) *\<^sub>R (a + a) = a" -proof - - have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" - by (metis scaleR_2 scaleR_scaleR) - then show ?thesis - by simp -qed - lemma vector_choose_size: assumes "0 \ c" obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c" diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Further_Topology.thy Thu Jan 05 22:38:06 2017 +0100 @@ -1134,11 +1134,6 @@ (*Up to extend_map_affine_to_sphere_cofinite_gen*) -lemma closedin_closed_subset: - "\closedin (subtopology euclidean U) V; T \ U; S = V \ T\ - \ closedin (subtopology euclidean T) S" - by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) - lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" @@ -2976,6 +2971,133 @@ by (simp add: aff_dim_cball) qed +lemma simply_connected_sphere_eq: + fixes a :: "'a::euclidean_space" + shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") +proof (cases "r \ 0") + case True + have "simply_connected (sphere a r)" + apply (rule convex_imp_simply_connected) + using True less_eq_real_def by auto + with True show ?thesis by auto +next + case False + show ?thesis + proof + assume L: ?lhs + have "False" if "DIM('a) = 1 \ DIM('a) = 2" + using that + proof + assume "DIM('a) = 1" + with L show False + using connected_sphere_eq simply_connected_imp_connected + by (metis False Suc_1 not_less_eq_eq order_refl) + next + assume "DIM('a) = 2" + then have "sphere a r homeomorphic sphere (0::complex) 1" + by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) + then have "simply_connected(sphere (0::complex) 1)" + using L homeomorphic_simply_connected_eq by blast + then obtain a::complex where "homotopic_with (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" + apply (simp add: simply_connected_eq_contractible_circlemap) + by (metis continuous_on_id' id_apply image_id subset_refl) + then show False + using contractible_sphere contractible_def not_one_le_zero by blast + qed + with False show ?rhs + apply simp + by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) + next + assume ?rhs + with False show ?lhs by (simp add: simply_connected_sphere) + qed +qed + + +lemma simply_connected_punctured_universe_eq: + fixes a :: "'a::euclidean_space" + shows "simply_connected(- {a}) \ 3 \ DIM('a)" +proof - + have [simp]: "a \ rel_interior (cball a 1)" + by (simp add: rel_interior_nonempty_interior) + have [simp]: "affine hull cball a 1 - {a} = -{a}" + by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) + have "simply_connected(- {a}) \ simply_connected(sphere a 1)" + apply (rule sym) + apply (rule homotopy_eqv_simple_connectedness) + using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto + done + also have "... \ 3 \ DIM('a)" + by (simp add: simply_connected_sphere_eq) + finally show ?thesis . +qed + +lemma not_simply_connected_circle: + fixes a :: complex + shows "0 < r \ \ simply_connected(sphere a r)" +by (simp add: simply_connected_sphere_eq) + + +proposition simply_connected_punctured_convex: + fixes a :: "'a::euclidean_space" + assumes "convex S" and 3: "3 \ aff_dim S" + shows "simply_connected(S - {a})" +proof (cases "a \ rel_interior S") + case True + then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" + by (auto simp: rel_interior_cball) + have con: "convex (cball a e \ affine hull S)" + by (simp add: convex_Int) + have bo: "bounded (cball a e \ affine hull S)" + by (simp add: bounded_Int) + have "affine hull S \ interior (cball a e) \ {}" + using \0 < e\ \a \ S\ hull_subset by fastforce + then have "3 \ aff_dim (affine hull S \ cball a e)" + by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) + also have "... = aff_dim (cball a e \ affine hull S)" + by (simp add: Int_commute) + finally have "3 \ aff_dim (cball a e \ affine hull S)" . + moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" + proof (rule homotopy_eqv_rel_frontier_punctured_convex) + show "a \ rel_interior (cball a e \ affine hull S)" + by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) + have "closed (cball a e \ affine hull S)" + by blast + then show "rel_frontier (cball a e \ affine hull S) \ S" + apply (simp add: rel_frontier_def) + using e by blast + show "S \ affine hull (cball a e \ affine hull S)" + by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) + qed (auto simp: assms con bo) + ultimately show ?thesis + using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] + by blast +next + case False + show ?thesis + apply (rule contractible_imp_simply_connected) + apply (rule contractible_convex_tweak_boundary_points [OF \convex S\]) + apply (simp add: False rel_interior_subset subset_Diff_insert) + by (meson Diff_subset closure_subset subset_trans) +qed + +corollary simply_connected_punctured_universe: + fixes a :: "'a::euclidean_space" + assumes "3 \ DIM('a)" + shows "simply_connected(- {a})" +proof - + have [simp]: "affine hull cball a 1 = UNIV" + apply auto + by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq) + have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" + apply (rule homotopy_eqv_simple_connectedness) + apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull) + apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+ + done + then show ?thesis + using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) +qed + subsection\The power, squaring and exponential functions as covering maps\ diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Homeomorphism.thy Thu Jan 05 22:38:06 2017 +0100 @@ -8,6 +8,34 @@ imports Path_Connected begin +lemma homeomorphic_spheres': + fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space" + assumes "0 < \" and dimeq: "DIM('a) = DIM('b)" + shows "(sphere a \) homeomorphic (sphere b \)" +proof - + obtain f :: "'a\'b" and g where "linear f" "linear g" + and fg: "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g(f x) = x" "\y. f(g y) = y" + by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq]) + then have "continuous_on UNIV f" "continuous_on UNIV g" + using linear_continuous_on linear_linear by blast+ + then show ?thesis + unfolding homeomorphic_minimal + apply(rule_tac x="\x. b + f(x - a)" in exI) + apply(rule_tac x="\x. a + g(x - b)" in exI) + using assms + apply (force intro: continuous_intros + continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg) + done +qed + +lemma homeomorphic_spheres_gen: + fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space" + assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)" + shows "(sphere a r homeomorphic sphere b s)" + apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres']) + using assms apply auto + done + subsection \Homeomorphism of all convex compact sets with nonempty interior\ proposition ray_to_rel_frontier: @@ -1360,4 +1388,1036 @@ shows "g1 x = g2 x" using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv by blast +lemma covering_space_locally: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes loc: "locally \ C" and cov: "covering_space C p S" + and pim: "\T. \T \ C; \ T\ \ \(p ` T)" + shows "locally \ S" +proof - + have "locally \ (p ` C)" + apply (rule locally_open_map_image [OF loc]) + using cov covering_space_imp_continuous apply blast + using cov covering_space_imp_surjective covering_space_open_map apply blast + by (simp add: pim) + then show ?thesis + using covering_space_imp_surjective [OF cov] by metis +qed + + +proposition covering_space_locally_eq: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes cov: "covering_space C p S" + and pim: "\T. \T \ C; \ T\ \ \(p ` T)" + and qim: "\q U. \U \ S; continuous_on U q; \ U\ \ \(q ` U)" + shows "locally \ S \ locally \ C" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + show ?rhs + proof (rule locallyI) + fix V x + assume V: "openin (subtopology euclidean C) V" and "x \ V" + have "p x \ p ` C" + by (metis IntE V \x \ V\ imageI openin_open) + then obtain T \ where "p x \ T" + and opeT: "openin (subtopology euclidean S) T" + and veq: "\\ = {x \ C. p x \ T}" + and ope: "\U\\. openin (subtopology euclidean C) U" + and hom: "\U\\. \q. homeomorphism U T p q" + using cov unfolding covering_space_def by (blast intro: that) + have "x \ \\" + using V veq \p x \ T\ \x \ V\ openin_imp_subset by fastforce + then obtain U where "x \ U" "U \ \" + by blast + then obtain q where opeU: "openin (subtopology euclidean C) U" and q: "homeomorphism U T p q" + using ope hom by blast + with V have "openin (subtopology euclidean C) (U \ V)" + by blast + then have UV: "openin (subtopology euclidean S) (p ` (U \ V))" + using cov covering_space_open_map by blast + obtain W W' where opeW: "openin (subtopology euclidean S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" + using locallyE [OF L UV] \x \ U\ \x \ V\ by blast + then have "W \ T" + by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) + show "\U Z. openin (subtopology euclidean C) U \ + \ Z \ x \ U \ U \ Z \ Z \ V" + proof (intro exI conjI) + have "openin (subtopology euclidean T) W" + by (meson opeW opeT openin_imp_subset openin_subset_trans \W \ T\) + then have "openin (subtopology euclidean U) (q ` W)" + by (meson homeomorphism_imp_open_map homeomorphism_symD q) + then show "openin (subtopology euclidean C) (q ` W)" + using opeU openin_trans by blast + show "\ (q ` W')" + by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \\ W'\ continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) + show "x \ q ` W" + by (metis \p x \ W\ \x \ U\ homeomorphism_def imageI q) + show "q ` W \ q ` W'" + using \W \ W'\ by blast + have "W' \ p ` V" + using W'sub by blast + then show "q ` W' \ V" + using W'sub homeomorphism_apply1 [OF q] by auto + qed + qed +next + assume ?rhs + then show ?lhs + using cov covering_space_locally pim by blast +qed + +lemma covering_space_locally_compact_eq: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "covering_space C p S" + shows "locally compact S \ locally compact C" + apply (rule covering_space_locally_eq [OF assms]) + apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous) + using compact_continuous_image by blast + +lemma covering_space_locally_connected_eq: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "covering_space C p S" + shows "locally connected S \ locally connected C" + apply (rule covering_space_locally_eq [OF assms]) + apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) + using connected_continuous_image by blast + +lemma covering_space_locally_path_connected_eq: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "covering_space C p S" + shows "locally path_connected S \ locally path_connected C" + apply (rule covering_space_locally_eq [OF assms]) + apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) + using path_connected_continuous_image by blast + + +lemma covering_space_locally_compact: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "locally compact C" "covering_space C p S" + shows "locally compact S" + using assms covering_space_locally_compact_eq by blast + + +lemma covering_space_locally_connected: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "locally connected C" "covering_space C p S" + shows "locally connected S" + using assms covering_space_locally_connected_eq by blast + +lemma covering_space_locally_path_connected: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes "locally path_connected C" "covering_space C p S" + shows "locally path_connected S" + using assms covering_space_locally_path_connected_eq by blast + +proposition covering_space_lift_homotopy: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and h :: "real \ 'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" + and conth: "continuous_on ({0..1} \ U) h" + and him: "h ` ({0..1} \ U) \ S" + and heq: "\y. y \ U \ h (0,y) = p(f y)" + and contf: "continuous_on U f" and fim: "f ` U \ C" + obtains k where "continuous_on ({0..1} \ U) k" + "k ` ({0..1} \ U) \ C" + "\y. y \ U \ k(0, y) = f y" + "\z. z \ {0..1} \ U \ h z = p(k z)" +proof - + have "\V k. openin (subtopology euclidean U) V \ y \ V \ + continuous_on ({0..1} \ V) k \ k ` ({0..1} \ V) \ C \ + (\z \ V. k(0, z) = f z) \ (\z \ {0..1} \ V. h z = p(k z))" + if "y \ U" for y + proof - + obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (subtopology euclidean S) (UU s) \ + (\\. \\ = {x. x \ C \ p x \ (UU s)} \ + (\U \ \. openin (subtopology euclidean C) U) \ + pairwise disjnt \ \ + (\U \ \. \q. homeomorphism U (UU s) p q))" + using cov unfolding covering_space_def by (metis (mono_tags)) + then have ope: "\s. s \ S \ s \ (UU s) \ openin (subtopology euclidean S) (UU s)" + by blast + have "\k n i. open k \ open n \ + t \ k \ y \ n \ i \ S \ h ` (({0..1} \ k) \ (U \ n)) \ UU i" if "t \ {0..1}" for t + proof - + have hinS: "h (t, y) \ S" + using \y \ U\ him that by blast + then have "(t,y) \ {z \ {0..1} \ U. h z \ UU (h (t, y))}" + using \y \ U\ \t \ {0..1}\ by (auto simp: ope) + moreover have ope_01U: "openin (subtopology euclidean ({0..1} \ U)) {z. z \ ({0..1} \ U) \ h z \ UU(h(t, y))}" + using hinS ope continuous_on_open_gen [OF him] conth by blast + ultimately obtain V W where opeV: "open V" and "t \ {0..1} \ V" "t \ {0..1} \ V" + and opeW: "open W" and "y \ U" "y \ W" + and VW: "({0..1} \ V) \ (U \ W) \ {z. z \ ({0..1} \ U) \ h z \ UU(h(t, y))}" + by (rule Times_in_interior_subtopology) (auto simp: openin_open) + then show ?thesis + using hinS by blast + qed + then obtain K NN X where + K: "\t. t \ {0..1} \ open (K t)" + and NN: "\t. t \ {0..1} \ open (NN t)" + and inUS: "\t. t \ {0..1} \ t \ K t \ y \ NN t \ X t \ S" + and him: "\t. t \ {0..1} \ h ` (({0..1} \ K t) \ (U \ NN t)) \ UU (X t)" + by (metis (mono_tags)) + obtain \ where "\ \ ((\i. K i \ NN i)) ` {0..1}" "finite \" "{0::real..1} \ {y} \ \\" + proof (rule compactE) + show "compact ({0::real..1} \ {y})" + by (simp add: compact_Times) + show "{0..1} \ {y} \ (\i\{0..1}. K i \ NN i)" + using K inUS by auto + show "\B. B \ (\i. K i \ NN i) ` {0..1} \ open B" + using K NN by (auto simp: open_Times) + qed blast + then obtain tk where "tk \ {0..1}" "finite tk" + and tk: "{0::real..1} \ {y} \ (\i \ tk. K i \ NN i)" + by (metis (no_types, lifting) finite_subset_image) + then have "tk \ {}" + by auto + define n where "n = INTER tk NN" + have "y \ n" "open n" + using inUS NN \tk \ {0..1}\ \finite tk\ + by (auto simp: n_def open_INT subset_iff) + obtain \ where "0 < \" and \: "\T. \T \ {0..1}; diameter T < \\ \ \B\K ` tk. T \ B" + proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"]) + show "K ` tk \ {}" + using \tk \ {}\ by auto + show "{0..1} \ UNION tk K" + using tk by auto + show "\B. B \ K ` tk \ open B" + using \tk \ {0..1}\ K by auto + qed auto + obtain N::nat where N: "N > 1 / \" + using reals_Archimedean2 by blast + then have "N > 0" + using \0 < \\ order.asym by force + have *: "\V k. openin (subtopology euclidean U) V \ y \ V \ + continuous_on ({0..of_nat n / N} \ V) k \ + k ` ({0..of_nat n / N} \ V) \ C \ + (\z\V. k (0, z) = f z) \ + (\z\{0..of_nat n / N} \ V. h z = p (k z))" if "n \ N" for n + using that + proof (induction n) + case 0 + show ?case + apply (rule_tac x=U in exI) + apply (rule_tac x="f \ snd" in exI) + apply (intro conjI \y \ U\ continuous_intros continuous_on_subset [OF contf]) + using fim apply (auto simp: heq) + done + next + case (Suc n) + then obtain V k where opeUV: "openin (subtopology euclidean U) V" + and "y \ V" + and contk: "continuous_on ({0..real n / real N} \ V) k" + and kim: "k ` ({0..real n / real N} \ V) \ C" + and keq: "\z. z \ V \ k (0, z) = f z" + and heq: "\z. z \ {0..real n / real N} \ V \ h z = p (k z)" + using Suc_leD by auto + have "n \ N" + using Suc.prems by auto + obtain t where "t \ tk" and t: "{real n / real N .. (1 + real n) / real N} \ K t" + proof (rule bexE [OF \]) + show "{real n / real N .. (1 + real n) / real N} \ {0..1}" + using Suc.prems by (auto simp: divide_simps algebra_simps) + show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \" + using \0 < \\ N by (auto simp: divide_simps algebra_simps) + qed blast + have t01: "t \ {0..1}" + using \t \ tk\ \tk \ {0..1}\ by blast + obtain \ where "\\ = {x. x \ C \ p x \ (UU (X t))}" + and opeC: "\U. U \ \ \ openin (subtopology euclidean C) U" + and "pairwise disjnt \" + and homuu: "\U. U \ \ \ \q. homeomorphism U (UU (X t)) p q" + using inUS [OF t01] UU by meson + have n_div_N_in: "real n / real N \ {real n / real N .. (1 + real n) / real N}" + using N by (auto simp: divide_simps algebra_simps) + with t have nN_in_kkt: "real n / real N \ K t" + by blast + have "k (real n / real N, y) \ {x. x \ C \ p x \ (UU (X t))}" + proof (simp, rule conjI) + show "k (real n / real N, y) \ C" + using \y \ V\ kim keq by force + have "p (k (real n / real N, y)) = h (real n / real N, y)" + by (simp add: \y \ V\ heq) + also have "... \ h ` (({0..1} \ K t) \ (U \ NN t))" + apply (rule imageI) + using \y \ V\ t01 \n \ N\ + apply (simp add: nN_in_kkt \y \ U\ inUS divide_simps) + done + also have "... \ UU (X t)" + using him t01 by blast + finally show "p (k (real n / real N, y)) \ UU (X t)" . + qed + then have "k (real n / real N, y) \ \\" + using \\\ = {x \ C. p x \ UU (X t)}\ by blast + then obtain W where W: "k (real n / real N, y) \ W" and "W \ \" + by blast + then obtain p' where opeC': "openin (subtopology euclidean C) W" + and hom': "homeomorphism W (UU (X t)) p p'" + using homuu opeC by blast + then have "W \ C" + using openin_imp_subset by blast + define W' where "W' = UU(X t)" + have opeVW: "openin (subtopology euclidean V) {z \ V. (k \ Pair (real n / real N)) z \ W}" + apply (rule continuous_openin_preimage [OF _ _ opeC']) + apply (intro continuous_intros continuous_on_subset [OF contk]) + using kim apply (auto simp: \y \ V\ W) + done + obtain N' where opeUN': "openin (subtopology euclidean U) N'" + and "y \ N'" and kimw: "k ` ({(real n / real N)} \ N') \ W" + apply (rule_tac N' = "{z \ V. (k \ Pair ((real n / real N))) z \ W}" in that) + apply (fastforce simp: \y \ V\ W intro!: openin_trans [OF opeVW opeUV])+ + done + obtain Q Q' where opeUQ: "openin (subtopology euclidean U) Q" + and cloUQ': "closedin (subtopology euclidean U) Q'" + and "y \ Q" "Q \ Q'" + and Q': "Q' \ (U \ NN(t)) \ N' \ V" + proof - + obtain VO VX where "open VO" "open VX" and VO: "V = U \ VO" and VX: "N' = U \ VX" + using opeUV opeUN' by (auto simp: openin_open) + then have "open (NN(t) \ VO \ VX)" + using NN t01 by blast + then obtain e where "e > 0" and e: "cball y e \ NN(t) \ VO \ VX" + by (metis Int_iff \N' = U \ VX\ \V = U \ VO\ \y \ N'\ \y \ V\ inUS open_contains_cball t01) + show ?thesis + proof + show "openin (subtopology euclidean U) (U \ ball y e)" + by blast + show "closedin (subtopology euclidean U) (U \ cball y e)" + using e by (auto simp: closedin_closed) + qed (use \y \ U\ \e > 0\ VO VX e in auto) + qed + then have "y \ Q'" "Q \ (U \ NN(t)) \ N' \ V" + by blast+ + have neq: "{0..real n / real N} \ {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}" + apply (auto simp: divide_simps) + by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1) + then have neqQ': "{0..real n / real N} \ Q' \ {real n / real N..(1 + real n) / real N} \ Q' = {0..(1 + real n) / real N} \ Q'" + by blast + have cont: "continuous_on ({0..(1 + real n) / real N} \ Q') + (\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" + unfolding neqQ' [symmetric] + proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) + show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \ Q')) + ({0..real n / real N} \ Q')" + apply (simp add: closedin_closed) + apply (rule_tac x="{0 .. real n / real N} \ UNIV" in exI) + using n_div_N_in apply (auto simp: closed_Times) + done + show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \ Q')) + ({real n / real N..(1 + real n) / real N} \ Q')" + apply (simp add: closedin_closed) + apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \ UNIV" in exI) + apply (auto simp: closed_Times) + by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) + show "continuous_on ({0..real n / real N} \ Q') k" + apply (rule continuous_on_subset [OF contk]) + using Q' by auto + have "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') h" + proof (rule continuous_on_subset [OF conth]) + show "{real n / real N..(1 + real n) / real N} \ Q' \ {0..1} \ U" + using \N > 0\ + apply auto + apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) + using Suc.prems order_trans apply fastforce + apply (metis IntE cloUQ' closedin_closed) + done + qed + moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \ Q')) p'" + proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']]) + have "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ h ` (({0..1} \ K t) \ (U \ NN t))" + apply (rule image_mono) + using \0 < \\ \N > 0\ Suc.prems apply auto + apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) + using Suc.prems order_trans apply fastforce + using t Q' apply auto + done + with him show "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ UU (X t)" + using t01 by blast + qed + ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') (p' \ h)" + by (rule continuous_on_compose) + have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \ Q'" for b + proof - + have "k (real n / real N, b) \ W" + using that Q' kimw by force + then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))" + by (simp add: homeomorphism_apply1 [OF hom']) + then show ?thesis + using Q' that by (force simp: heq) + qed + then show "\x. x \ {real n / real N..(1 + real n) / real N} \ Q' \ + x \ {0..real n / real N} \ Q' \ k x = (p' \ h) x" + by auto + qed + have h_in_UU: "h (x, y) \ UU (X t)" if "y \ Q" "\ x \ real n / real N" "0 \ x" "x \ (1 + real n) / real N" for x y + proof - + have "x \ 1" + using Suc.prems that order_trans by force + moreover have "x \ K t" + by (meson atLeastAtMost_iff le_less not_le subset_eq t that) + moreover have "y \ U" + using \y \ Q\ opeUQ openin_imp_subset by blast + moreover have "y \ NN t" + using Q' \Q \ Q'\ \y \ Q\ by auto + ultimately have "(x, y) \ (({0..1} \ K t) \ (U \ NN t))" + using that by auto + then have "h (x, y) \ h ` (({0..1} \ K t) \ (U \ NN t))" + by blast + also have "... \ UU (X t)" + by (metis him t01) + finally show ?thesis . + qed + let ?k = "(\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" + show ?case + proof (intro exI conjI) + show "continuous_on ({0..real (Suc n) / real N} \ Q) ?k" + apply (rule continuous_on_subset [OF cont]) + using \Q \ Q'\ by auto + have "\a b. \a \ real n / real N; b \ Q'; 0 \ a\ \ k (a, b) \ C" + using kim Q' by force + moreover have "\a b. \b \ Q; 0 \ a; a \ (1 + real n) / real N; \ a \ real n / real N\ \ p' (h (a, b)) \ C" + apply (rule \W \ C\ [THEN subsetD]) + using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \Q \ Q'\ \W \ C\ + apply auto + done + ultimately show "?k ` ({0..real (Suc n) / real N} \ Q) \ C" + using Q' \Q \ Q'\ by force + show "\z\Q. ?k (0, z) = f z" + using Q' keq \Q \ Q'\ by auto + show "\z \ {0..real (Suc n) / real N} \ Q. h z = p(?k z)" + using \Q \ U \ NN t \ N' \ V\ heq apply clarsimp + using h_in_UU Q' \Q \ Q'\ apply (auto simp: homeomorphism_apply2 [OF hom', symmetric]) + done + qed (auto simp: \y \ Q\ opeUQ) + qed + show ?thesis + using*[OF order_refl] N \0 < \\ by (simp add: split: if_split_asm) + qed + then obtain V fs where opeV: "\y. y \ U \ openin (subtopology euclidean U) (V y)" + and V: "\y. y \ U \ y \ V y" + and contfs: "\y. y \ U \ continuous_on ({0..1} \ V y) (fs y)" + and *: "\y. y \ U \ (fs y) ` ({0..1} \ V y) \ C \ + (\z \ V y. fs y (0, z) = f z) \ + (\z \ {0..1} \ V y. h z = p(fs y z))" + by (metis (mono_tags)) + then have VU: "\y. y \ U \ V y \ U" + by (meson openin_imp_subset) + obtain k where contk: "continuous_on ({0..1} \ U) k" + and k: "\x i. \i \ U; x \ {0..1} \ U \ {0..1} \ V i\ \ k x = fs i x" + proof (rule pasting_lemma_exists) + show "{0..1} \ U \ (\i\U. {0..1} \ V i)" + apply auto + using V by blast + show "\i. i \ U \ openin (subtopology euclidean ({0..1} \ U)) ({0..1} \ V i)" + by (simp add: opeV openin_Times) + show "\i. i \ U \ continuous_on ({0..1} \ V i) (fs i)" + using contfs by blast + show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ {0..1} \ U \ {0..1} \ V i \ {0..1} \ V j" + for i j x + proof - + obtain u y where "x = (u, y)" "y \ V i" "y \ V j" "0 \ u" "u \ 1" + using x by auto + show ?thesis + proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \ {y}" h]) + show "fs i (0, y) = fs j (0, y)" + using*V by (simp add: \y \ V i\ \y \ V j\ that) + show conth_y: "continuous_on ({0..1} \ {y}) h" + apply (rule continuous_on_subset [OF conth]) + using VU \y \ V j\ that by auto + show "h ` ({0..1} \ {y}) \ S" + using \y \ V i\ assms(3) VU that by fastforce + show "continuous_on ({0..1} \ {y}) (fs i)" + using continuous_on_subset [OF contfs] \i \ U\ + by (simp add: \y \ V i\ subset_iff) + show "fs i ` ({0..1} \ {y}) \ C" + using "*" \y \ V i\ \i \ U\ by fastforce + show "\x. x \ {0..1} \ {y} \ h x = p (fs i x)" + using "*" \y \ V i\ \i \ U\ by blast + show "continuous_on ({0..1} \ {y}) (fs j)" + using continuous_on_subset [OF contfs] \j \ U\ + by (simp add: \y \ V j\ subset_iff) + show "fs j ` ({0..1} \ {y}) \ C" + using "*" \y \ V j\ \j \ U\ by fastforce + show "\x. x \ {0..1} \ {y} \ h x = p (fs j x)" + using "*" \y \ V j\ \j \ U\ by blast + show "connected ({0..1::real} \ {y})" + using connected_Icc connected_Times connected_sing by blast + show "(0, y) \ {0..1::real} \ {y}" + by force + show "x \ {0..1} \ {y}" + using \x = (u, y)\ x by blast + qed + qed + qed blast + show ?thesis + proof + show "k ` ({0..1} \ U) \ C" + using V*k VU by fastforce + show "\y. y \ U \ k (0, y) = f y" + by (simp add: V*k) + show "\z. z \ {0..1} \ U \ h z = p (k z)" + using V*k by auto + qed (auto simp: contk) +qed + +corollary covering_space_lift_homotopy_alt: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and h :: "'c::real_normed_vector \ real \ 'b" + assumes cov: "covering_space C p S" + and conth: "continuous_on (U \ {0..1}) h" + and him: "h ` (U \ {0..1}) \ S" + and heq: "\y. y \ U \ h (y,0) = p(f y)" + and contf: "continuous_on U f" and fim: "f ` U \ C" + obtains k where "continuous_on (U \ {0..1}) k" + "k ` (U \ {0..1}) \ C" + "\y. y \ U \ k(y, 0) = f y" + "\z. z \ U \ {0..1} \ h z = p(k z)" +proof - + have "continuous_on ({0..1} \ U) (h \ (\z. (snd z, fst z)))" + by (intro continuous_intros continuous_on_subset [OF conth]) auto + then obtain k where contk: "continuous_on ({0..1} \ U) k" + and kim: "k ` ({0..1} \ U) \ C" + and k0: "\y. y \ U \ k(0, y) = f y" + and heqp: "\z. z \ {0..1} \ U \ (h \ (\z. Pair (snd z) (fst z))) z = p(k z)" + apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim]) + using him by (auto simp: contf heq) + show ?thesis + apply (rule_tac k="k \ (\z. Pair (snd z) (fst z))" in that) + apply (intro continuous_intros continuous_on_subset [OF contk]) + using kim heqp apply (auto simp: k0) + done +qed + +corollary covering_space_lift_homotopic_function: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and g:: "'c::real_normed_vector \ 'a" + assumes cov: "covering_space C p S" + and contg: "continuous_on U g" + and gim: "g ` U \ C" + and pgeq: "\y. y \ U \ p(g y) = f y" + and hom: "homotopic_with (\x. True) U S f f'" + obtains g' where "continuous_on U g'" "image g' U \ C" "\y. y \ U \ p(g' y) = f' y" +proof - + obtain h where conth: "continuous_on ({0..1::real} \ U) h" + and him: "h ` ({0..1} \ U) \ S" + and h0: "\x. h(0, x) = f x" + and h1: "\x. h(1, x) = f' x" + using hom by (auto simp: homotopic_with_def) + have "\y. y \ U \ h (0, y) = p (g y)" + by (simp add: h0 pgeq) + then obtain k where contk: "continuous_on ({0..1} \ U) k" + and kim: "k ` ({0..1} \ U) \ C" + and k0: "\y. y \ U \ k(0, y) = g y" + and heq: "\z. z \ {0..1} \ U \ h z = p(k z)" + using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis + show ?thesis + proof + show "continuous_on U (k \ Pair 1)" + by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one) + show "(k \ Pair 1) ` U \ C" + using kim by auto + show "\y. y \ U \ p ((k \ Pair 1) y) = f' y" + by (auto simp: h1 heq [symmetric]) + qed +qed + +corollary covering_space_lift_inessential_function: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and U :: "'c::real_normed_vector set" + assumes cov: "covering_space C p S" + and hom: "homotopic_with (\x. True) U S f (\x. a)" + obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" +proof (cases "U = {}") + case True + then show ?thesis + using that continuous_on_empty by blast +next + case False + then obtain b where b: "b \ C" "p b = a" + using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom] + by auto + then have gim: "(\y. b) ` U \ C" + by blast + show ?thesis + apply (rule covering_space_lift_homotopic_function + [OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]]) + using b that apply auto + done +qed + +subsection\ Lifting of general functions to covering space\ + +proposition covering_space_lift_path_strong: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and f :: "'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" and "a \ C" + and "path g" and pag: "path_image g \ S" and pas: "pathstart g = p a" + obtains h where "path h" "path_image h \ C" "pathstart h = a" + and "\t. t \ {0..1} \ p(h t) = g t" +proof - + obtain k:: "real \ 'c \ 'a" + where contk: "continuous_on ({0..1} \ {undefined}) k" + and kim: "k ` ({0..1} \ {undefined}) \ C" + and k0: "k (0, undefined) = a" + and pk: "\z. z \ {0..1} \ {undefined} \ p(k z) = (g \ fst) z" + proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \ fst"]) + show "continuous_on ({0..1::real} \ {undefined::'c}) (g \ fst)" + apply (intro continuous_intros) + using \path g\ by (simp add: path_def) + show "(g \ fst) ` ({0..1} \ {undefined}) \ S" + using pag by (auto simp: path_image_def) + show "(g \ fst) (0, y) = p a" if "y \ {undefined}" for y::'c + by (metis comp_def fst_conv pas pathstart_def) + qed (use assms in auto) + show ?thesis + proof + show "path (k \ (\t. Pair t undefined))" + unfolding path_def + by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto + show "path_image (k \ (\t. (t, undefined))) \ C" + using kim by (auto simp: path_image_def) + show "pathstart (k \ (\t. (t, undefined))) = a" + by (auto simp: pathstart_def k0) + show "\t. t \ {0..1} \ p ((k \ (\t. (t, undefined))) t) = g t" + by (auto simp: pk) + qed +qed + +corollary covering_space_lift_path: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \ S" + obtains h where "path h" "path_image h \ C" "\t. t \ {0..1} \ p(h t) = g t" +proof - + obtain a where "a \ C" "pathstart g = p a" + by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE) + show ?thesis + using covering_space_lift_path_strong [OF cov \a \ C\ \path g\ pig] + by (metis \pathstart g = p a\ that) +qed + + +proposition covering_space_lift_homotopic_paths: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes cov: "covering_space C p S" + and "path g1" and pig1: "path_image g1 \ S" + and "path g2" and pig2: "path_image g2 \ S" + and hom: "homotopic_paths S g1 g2" + and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" + and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" + and h1h2: "pathstart h1 = pathstart h2" + shows "homotopic_paths C h1 h2" +proof - + obtain h :: "real \ real \ 'b" + where conth: "continuous_on ({0..1} \ {0..1}) h" + and him: "h ` ({0..1} \ {0..1}) \ S" + and h0: "\x. h (0, x) = g1 x" and h1: "\x. h (1, x) = g2 x" + and heq0: "\t. t \ {0..1} \ h (t, 0) = g1 0" + and heq1: "\t. t \ {0..1} \ h (t, 1) = g1 1" + using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def) + obtain k where contk: "continuous_on ({0..1} \ {0..1}) k" + and kim: "k ` ({0..1} \ {0..1}) \ C" + and kh2: "\y. y \ {0..1} \ k (y, 0) = h2 0" + and hpk: "\z. z \ {0..1} \ {0..1} \ h z = p (k z)" + apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\x. h2 0"]) + using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def) + using path_image_def pih2 continuous_on_const by fastforce+ + have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2" + using \path g1\ \path g2\ path_def by blast+ + have g1im: "g1 ` {0..1} \ S" and g2im: "g2 ` {0..1} \ S" + using path_image_def pig1 pig2 by auto + have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2" + using \path h1\ \path h2\ path_def by blast+ + have h1im: "h1 ` {0..1} \ C" and h2im: "h2 ` {0..1} \ C" + using path_image_def pih1 pih2 by auto + show ?thesis + unfolding homotopic_paths pathstart_def pathfinish_def + proof (intro exI conjI ballI) + show keqh1: "k(0, x) = h1 x" if "x \ {0..1}" for x + proof (rule covering_space_lift_unique [OF cov _ contg1 g1im]) + show "k (0,0) = h1 0" + by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one) + show "continuous_on {0..1} (\a. k (0, a))" + by (intro continuous_intros continuous_on_compose2 [OF contk]) auto + show "\x. x \ {0..1} \ g1 x = p (k (0, x))" + by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl) + qed (use conth1 h1im kim that in \auto simp: ph1\) + show "k(1, x) = h2 x" if "x \ {0..1}" for x + proof (rule covering_space_lift_unique [OF cov _ contg2 g2im]) + show "k (1,0) = h2 0" + by (metis atLeastAtMost_iff kh2 order_refl zero_le_one) + show "continuous_on {0..1} (\a. k (1, a))" + by (intro continuous_intros continuous_on_compose2 [OF contk]) auto + show "\x. x \ {0..1} \ g2 x = p (k (1, x))" + by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one) + qed (use conth2 h2im kim that in \auto simp: ph2\) + show "\t. t \ {0..1} \ (k \ Pair t) 0 = h1 0" + by (metis comp_apply h1h2 kh2 pathstart_def) + show "(k \ Pair t) 1 = h1 1" if "t \ {0..1}" for t + proof (rule covering_space_lift_unique + [OF cov, of "\a. (k \ Pair a) 1" 0 "\a. h1 1" "{0..1}" "\x. g1 1"]) + show "(k \ Pair 0) 1 = h1 1" + using keqh1 by auto + show "continuous_on {0..1} (\a. (k \ Pair a) 1)" + apply simp + by (intro continuous_intros continuous_on_compose2 [OF contk]) auto + show "\x. x \ {0..1} \ g1 1 = p ((k \ Pair x) 1)" + using heq1 hpk by auto + qed (use contk kim g1im h1im that in \auto simp: ph1 continuous_on_const\) + qed (use contk kim in auto) +qed + + +corollary covering_space_monodromy: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes cov: "covering_space C p S" + and "path g1" and pig1: "path_image g1 \ S" + and "path g2" and pig2: "path_image g2 \ S" + and hom: "homotopic_paths S g1 g2" + and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" + and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" + and h1h2: "pathstart h1 = pathstart h2" + shows "pathfinish h1 = pathfinish h2" + using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish by blast + + +corollary covering_space_lift_homotopic_path: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes cov: "covering_space C p S" + and hom: "homotopic_paths S f f'" + and "path g" and pig: "path_image g \ C" + and a: "pathstart g = a" and b: "pathfinish g = b" + and pgeq: "\t. t \ {0..1} \ p(g t) = f t" + obtains g' where "path g'" "path_image g' \ C" + "pathstart g' = a" "pathfinish g' = b" "\t. t \ {0..1} \ p(g' t) = f' t" +proof (rule covering_space_lift_path_strong [OF cov, of a f']) + show "a \ C" + using a pig by auto + show "path f'" "path_image f' \ S" + using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+ + show "pathstart f' = p a" + by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one) +qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig) + + +proposition covering_space_lift_general: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and f :: "'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" and "a \ C" "z \ U" + and U: "path_connected U" "locally path_connected U" + and contf: "continuous_on U f" and fim: "f ` U \ S" + and feq: "f z = p a" + and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ + \ \q. path q \ path_image q \ C \ + pathstart q = a \ pathfinish q = a \ + homotopic_paths S (f \ r) (p \ q)" + obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" +proof - + have *: "\g h. path g \ path_image g \ U \ + pathstart g = z \ pathfinish g = y \ + path h \ path_image h \ C \ pathstart h = a \ + (\t \ {0..1}. p(h t) = f(g t))" + if "y \ U" for y + proof - + obtain g where "path g" "path_image g \ U" and pastg: "pathstart g = z" + and pafig: "pathfinish g = y" + using U \z \ U\ \y \ U\ by (force simp: path_connected_def) + obtain h where "path h" "path_image h \ C" "pathstart h = a" + and "\t. t \ {0..1} \ p(h t) = (f \ g) t" + proof (rule covering_space_lift_path_strong [OF cov \a \ C\]) + show "path (f \ g)" + using \path g\ \path_image g \ U\ contf continuous_on_subset path_continuous_image by blast + show "path_image (f \ g) \ S" + by (metis \path_image g \ U\ fim image_mono path_image_compose subset_trans) + show "pathstart (f \ g) = p a" + by (simp add: feq pastg pathstart_compose) + qed auto + then show ?thesis + by (metis \path g\ \path_image g \ U\ comp_apply pafig pastg) + qed + have "\l. \g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ + path h \ path_image h \ C \ pathstart h = a \ + (\t \ {0..1}. p(h t) = f(g t)) \ pathfinish h = l" for y + proof - + have "pathfinish h = pathfinish h'" + if g: "path g" "path_image g \ U" "pathstart g = z" "pathfinish g = y" + and h: "path h" "path_image h \ C" "pathstart h = a" + and phg: "\t. t \ {0..1} \ p(h t) = f(g t)" + and g': "path g'" "path_image g' \ U" "pathstart g' = z" "pathfinish g' = y" + and h': "path h'" "path_image h' \ C" "pathstart h' = a" + and phg': "\t. t \ {0..1} \ p(h' t) = f(g' t)" + for g h g' h' + proof - + obtain q where "path q" and piq: "path_image q \ C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a" + and homS: "homotopic_paths S (f \ g +++ reversepath g') (p \ q)" + using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join) + have papq: "path (p \ q)" + using homS homotopic_paths_imp_path by blast + have pipq: "path_image (p \ q) \ S" + using homS homotopic_paths_imp_subset by blast + obtain q' where "path q'" "path_image q' \ C" + and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q" + and pq'_eq: "\t. t \ {0..1} \ p (q' t) = (f \ g +++ reversepath g') t" + using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \path q\ piq refl refl] + by auto + have "q' t = (h \ op *\<^sub>R 2) t" if "0 \ t" "t \ 1/2" for t + proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \ op *\<^sub>R 2" "{0..1/2}" "f \ g \ op *\<^sub>R 2" t]) + show "q' 0 = (h \ op *\<^sub>R 2) 0" + by (metis \pathstart q' = pathstart q\ comp_def g h pastq pathstart_def pth_4(2)) + show "continuous_on {0..1/2} (f \ g \ op *\<^sub>R 2)" + apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path g\] continuous_on_subset [OF contf]) + using g(2) path_image_def by fastforce+ + show "(f \ g \ op *\<^sub>R 2) ` {0..1/2} \ S" + using g(2) path_image_def fim by fastforce + show "(h \ op *\<^sub>R 2) ` {0..1/2} \ C" + using h path_image_def by fastforce + show "q' ` {0..1/2} \ C" + using \path_image q' \ C\ path_image_def by fastforce + show "\x. x \ {0..1/2} \ (f \ g \ op *\<^sub>R 2) x = p (q' x)" + by (auto simp: joinpaths_def pq'_eq) + show "\x. x \ {0..1/2} \ (f \ g \ op *\<^sub>R 2) x = p ((h \ op *\<^sub>R 2) x)" + by (simp add: phg) + show "continuous_on {0..1/2} q'" + by (simp add: continuous_on_path \path q'\) + show "continuous_on {0..1/2} (h \ op *\<^sub>R 2)" + apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path h\], force) + done + qed (use that in auto) + moreover have "q' t = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \ 1" for t + proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \ (\t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)" t]) + show "q' 1 = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) 1" + using h' \pathfinish q' = pathfinish q\ pafiq + by (simp add: pathstart_def pathfinish_def reversepath_def) + show "continuous_on {1/2<..1} (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1))" + apply (intro continuous_intros continuous_on_compose continuous_on_path \path g'\ continuous_on_subset [OF contf]) + using g' apply simp_all + by (auto simp: path_image_def reversepath_def) + show "(f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ S" + using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def) + show "q' ` {1/2<..1} \ C" + using \path_image q' \ C\ path_image_def by fastforce + show "(reversepath h' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ C" + using h' by (simp add: path_image_def reversepath_def subset_eq) + show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p (q' x)" + by (auto simp: joinpaths_def pq'_eq) + show "\x. x \ {1/2<..1} \ + (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \ (\t. 2 *\<^sub>R t - 1)) x)" + by (simp add: phg' reversepath_def) + show "continuous_on {1/2<..1} q'" + by (auto intro: continuous_on_path [OF \path q'\]) + show "continuous_on {1/2<..1} (reversepath h' \ (\t. 2 *\<^sub>R t - 1))" + apply (intro continuous_intros continuous_on_compose continuous_on_path \path h'\) + using h' apply auto + done + qed (use that in auto) + ultimately have "q' t = (h +++ reversepath h') t" if "0 \ t" "t \ 1" for t + using that by (simp add: joinpaths_def) + then have "path(h +++ reversepath h')" + by (auto intro: path_eq [OF \path q'\]) + then show ?thesis + by (auto simp: \path h\ \path h'\) + qed + then show ?thesis by metis + qed + then obtain l :: "'c \ 'a" + where l: "\y g h. \path g; path_image g \ U; pathstart g = z; pathfinish g = y; + path h; path_image h \ C; pathstart h = a; + \t. t \ {0..1} \ p(h t) = f(g t)\ \ pathfinish h = l y" + by metis + show ?thesis + proof + show pleq: "p (l y) = f y" if "y \ U" for y + using*[OF \y \ U\] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one) + show "l z = a" + using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) + show LC: "l ` U \ C" + by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) + have "\T. openin (subtopology euclidean U) T \ y \ T \ T \ {x \ U. l x \ X}" + if X: "openin (subtopology euclidean C) X" and "y \ U" "l y \ X" for X y + proof - + have "X \ C" + using X openin_euclidean_subtopology_iff by blast + have "f y \ S" + using fim \y \ U\ by blast + then obtain W \ + where WV: "f y \ W \ openin (subtopology euclidean S) W \ + (\\ = {x. x \ C \ p x \ W} \ + (\U \ \. openin (subtopology euclidean C) U) \ + pairwise disjnt \ \ + (\U \ \. \q. homeomorphism U W p q))" + using cov by (force simp: covering_space_def) + then have "l y \ \\" + using \X \ C\ pleq that by auto + then obtain W' where "l y \ W'" and "W' \ \" + by blast + with WV obtain p' where opeCW': "openin (subtopology euclidean C) W'" + and homUW': "homeomorphism W' W p p'" + by blast + then have contp': "continuous_on W p'" and p'im: "p' ` W \ W'" + using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ + obtain V where "y \ V" "y \ U" and fimW: "f ` V \ W" "V \ U" + and "path_connected V" and opeUV: "openin (subtopology euclidean U) V" + proof - + have "openin (subtopology euclidean U) {c \ U. f c \ W}" + using WV contf continuous_on_open_gen fim by auto + then show ?thesis + using U WV + apply (auto simp: locally_path_connected) + apply (drule_tac x="{x. x \ U \ f x \ W}" in spec) + apply (drule_tac x=y in spec) + apply (auto simp: \y \ U\ intro: that) + done + qed + have "W' \ C" "W \ S" + using opeCW' WV openin_imp_subset by auto + have p'im: "p' ` W \ W'" + using homUW' homeomorphism_image2 by fastforce + show ?thesis + proof (intro exI conjI) + have "openin (subtopology euclidean S) {x \ W. p' x \ W' \ X}" + proof (rule openin_trans) + show "openin (subtopology euclidean W) {x \ W. p' x \ W' \ X}" + apply (rule continuous_openin_preimage [OF contp' p'im]) + using X \W' \ C\ apply (auto simp: openin_open) + done + show "openin (subtopology euclidean S) W" + using WV by blast + qed + then show "openin (subtopology euclidean U) (V \ {x. x \ U \ f x \ {x. x \ W \ p' x \ W' \ X}})" + by (intro openin_Int opeUV continuous_openin_preimage [OF contf fim]) + have "p' (f y) \ X" + using \l y \ W'\ homeomorphism_apply1 [OF homUW'] pleq \y \ U\ \l y \ X\ by fastforce + then show "y \ V \ {x \ U. f x \ {x \ W. p' x \ W' \ X}}" + using \y \ U\ \y \ V\ WV p'im by auto + show "V \ {x \ U. f x \ {x \ W. p' x \ W' \ X}} \ {x \ U. l x \ X}" + proof clarsimp + fix y' + assume y': "y' \ V" "y' \ U" "f y' \ W" "p' (f y') \ W'" "p' (f y') \ X" + then obtain \ where "path \" "path_image \ \ V" "pathstart \ = y" "pathfinish \ = y'" + by (meson \path_connected V\ \y \ V\ path_connected_def) + obtain pp qq where "path pp" "path_image pp \ U" + "pathstart pp = z" "pathfinish pp = y" + "path qq" "path_image qq \ C" "pathstart qq = a" + and pqqeq: "\t. t \ {0..1} \ p(qq t) = f(pp t)" + using*[OF \y \ U\] by blast + have finW: "\x. \0 \ x; x \ 1\ \ f (\ x) \ W" + using \path_image \ \ V\ by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD]) + have "pathfinish (qq +++ (p' \ f \ \)) = l y'" + proof (rule l [of "pp +++ \" y' "qq +++ (p' \ f \ \)" ]) + show "path (pp +++ \)" + by (simp add: \path \\ \path pp\ \pathfinish pp = y\ \pathstart \ = y\) + show "path_image (pp +++ \) \ U" + using \V \ U\ \path_image \ \ V\ \path_image pp \ U\ not_in_path_image_join by blast + show "pathstart (pp +++ \) = z" + by (simp add: \pathstart pp = z\) + show "pathfinish (pp +++ \) = y'" + by (simp add: \pathfinish \ = y'\) + have paqq: "pathfinish qq = pathstart (p' \ f \ \)" + apply (simp add: \pathstart \ = y\ pathstart_compose) + apply (metis (mono_tags, lifting) \l y \ W'\ \path pp\ \path qq\ \path_image pp \ U\ \path_image qq \ C\ + \pathfinish pp = y\ \pathstart pp = z\ \pathstart qq = a\ + homeomorphism_apply1 [OF homUW'] l pleq pqqeq \y \ U\) + done + have "continuous_on (path_image \) (p' \ f)" + proof (rule continuous_on_compose) + show "continuous_on (path_image \) f" + using \path_image \ \ V\ \V \ U\ contf continuous_on_subset by blast + show "continuous_on (f ` path_image \) p'" + apply (rule continuous_on_subset [OF contp']) + apply (auto simp: path_image_def pathfinish_def pathstart_def finW) + done + qed + then show "path (qq +++ (p' \ f \ \))" + using \path \\ \path qq\ paqq path_continuous_image path_join_imp by blast + show "path_image (qq +++ (p' \ f \ \)) \ C" + apply (rule subset_path_image_join) + apply (simp add: \path_image qq \ C\) + by (metis \W' \ C\ \path_image \ \ V\ dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose) + show "pathstart (qq +++ (p' \ f \ \)) = a" + by (simp add: \pathstart qq = a\) + show "p ((qq +++ (p' \ f \ \)) \) = f ((pp +++ \) \)" if \: "\ \ {0..1}" for \ + proof (simp add: joinpaths_def, safe) + show "p (qq (2*\)) = f (pp (2*\))" if "\*2 \ 1" + using \\ \ {0..1}\ pqqeq that by auto + show "p (p' (f (\ (2*\ - 1)))) = f (\ (2*\ - 1))" if "\ \*2 \ 1" + apply (rule homeomorphism_apply2 [OF homUW' finW]) + using that \ by auto + qed + qed + with \pathfinish \ = y'\ \p' (f y') \ X\ show "l y' \ X" + unfolding pathfinish_join by (simp add: pathfinish_def) + qed + qed + qed + then show "continuous_on U l" + using openin_subopen continuous_on_open_gen [OF LC] + by (metis (no_types, lifting) mem_Collect_eq) + qed +qed + + +corollary covering_space_lift_stronger: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and f :: "'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" "a \ C" "z \ U" + and U: "path_connected U" "locally path_connected U" + and contf: "continuous_on U f" and fim: "f ` U \ S" + and feq: "f z = p a" + and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ + \ \b. homotopic_paths S (f \ r) (linepath b b)" + obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" +proof (rule covering_space_lift_general [OF cov U contf fim feq]) + fix r + assume "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" + then obtain b where b: "homotopic_paths S (f \ r) (linepath b b)" + using hom by blast + then have "f (pathstart r) = b" + by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath) + then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" + by (simp add: b \pathstart r = z\) + then have "homotopic_paths S (f \ r) (p \ linepath a a)" + by (simp add: o_def feq linepath_def) + then show "\q. path q \ + path_image q \ C \ + pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" + by (force simp: \a \ C\) +qed auto + +corollary covering_space_lift_strong: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and f :: "'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" "a \ C" "z \ U" + and scU: "simply_connected U" and lpcU: "locally path_connected U" + and contf: "continuous_on U f" and fim: "f ` U \ S" + and feq: "f z = p a" + obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" +proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq]) + show "path_connected U" + using scU simply_connected_eq_contractible_loop_some by blast + fix r + assume r: "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" + have "linepath (f z) (f z) = f \ linepath z z" + by (simp add: o_def linepath_def) + then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" + by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path) + then show "\b. homotopic_paths S (f \ r) (linepath b b)" + by blast +qed blast + +corollary covering_space_lift: + fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" + and f :: "'c::real_normed_vector \ 'b" + assumes cov: "covering_space C p S" + and U: "simply_connected U" "locally path_connected U" + and contf: "continuous_on U f" and fim: "f ` U \ S" + obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" +proof (cases "U = {}") + case True + with that show ?thesis by auto +next + case False + then obtain z where "z \ U" by blast + then obtain a where "a \ C" "f z = p a" + by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff) + then show ?thesis + by (metis that covering_space_lift_strong [OF cov _ \z \ U\ U contf fim]) +qed + end diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Path_Connected.thy Thu Jan 05 22:38:06 2017 +0100 @@ -134,12 +134,18 @@ subsection\Basic lemmas about paths\ +lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" + using continuous_on_subset path_def by blast + lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast +lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" + by (auto simp: arc_def) + lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast @@ -1047,7 +1053,7 @@ done qed -subsection \Reparametrizing a closed curve to start at some chosen point\ +subsection \shiftpath: Reparametrizing a closed curve to start at some chosen point\ definition shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" @@ -1146,6 +1152,22 @@ by (auto simp add: image_iff) qed +lemma simple_path_shiftpath: + assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" + shows "simple_path (shiftpath a g)" + unfolding simple_path_def +proof (intro conjI impI ballI) + show "path (shiftpath a g)" + by (simp add: assms path_shiftpath simple_path_imp_path) + have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" + using assms by (simp add: simple_path_def) + show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" + if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y + using that a unfolding shiftpath_def + apply (simp add: split: if_split_asm) + apply (drule *; auto)+ + done +qed subsection \Special case of straight-line paths\ @@ -1499,8 +1521,8 @@ using path_connected_component_set by auto lemma path_connected_imp_connected: - assumes "path_connected s" - shows "connected s" + assumes "path_connected S" + shows "connected S" unfolding connected_def not_ex apply rule apply rule @@ -1509,10 +1531,10 @@ apply (elim conjE) proof - fix e1 e2 - assume as: "open e1" "open e2" "s \ e1 \ e2" "e1 \ e2 \ s = {}" "e1 \ s \ {}" "e2 \ s \ {}" - then obtain x1 x2 where obt:"x1 \ e1 \ s" "x2 \ e2 \ s" + assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" + then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto - then obtain g where g: "path g" "path_image g \ s" "pathstart g = x1" "pathfinish g = x2" + then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected convex_real_interval) @@ -1535,23 +1557,23 @@ qed lemma open_path_component: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - shows "open (path_component_set s x)" + fixes S :: "'a::real_normed_vector set" + assumes "open S" + shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y - assume as: "y \ path_component_set s x" - then have "y \ s" + assume as: "y \ path_component_set S x" + then have "y \ S" apply - apply (rule path_component_mem(2)) unfolding mem_Collect_eq apply auto done - then obtain e where e: "e > 0" "ball y e \ s" + then obtain e where e: "e > 0" "ball y e \ S" using assms[unfolded open_contains_ball] by auto - show "\e > 0. ball y e \ path_component_set s x" + show "\e > 0. ball y e \ path_component_set S x" apply (rule_tac x=e in exI) apply (rule,rule \e>0\) apply rule @@ -1559,7 +1581,7 @@ proof - fix z assume "dist y z < e" - then show "path_component s x z" + then show "path_component S x z" apply (rule_tac path_component_trans[of _ _ y]) defer apply (rule path_component_of_subset[OF e(2)]) @@ -1571,17 +1593,17 @@ qed lemma open_non_path_component: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - shows "open (s - path_component_set s x)" + fixes S :: "'a::real_normed_vector set" + assumes "open S" + shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y - assume as: "y \ s - path_component_set s x" - then obtain e where e: "e > 0" "ball y e \ s" + assume as: "y \ S - path_component_set S x" + then obtain e where e: "e > 0" "ball y e \ S" using assms [unfolded open_contains_ball] by auto - show "\e>0. ball y e \ s - path_component_set s x" + show "\e>0. ball y e \ S - path_component_set S x" apply (rule_tac x=e in exI) apply rule apply (rule \e>0\) @@ -1590,8 +1612,8 @@ defer proof (rule ccontr) fix z - assume "z \ ball y e" "\ z \ path_component_set s x" - then have "y \ path_component_set s x" + assume "z \ ball y e" "\ z \ path_component_set S x" + then have "y \ path_component_set S x" unfolding not_not mem_Collect_eq using \e>0\ apply - apply (rule path_component_trans, assumption) @@ -1605,42 +1627,42 @@ qed lemma connected_open_path_connected: - fixes s :: "'a::real_normed_vector set" - assumes "open s" - and "connected s" - shows "path_connected s" + fixes S :: "'a::real_normed_vector set" + assumes "open S" + and "connected S" + shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y - assume "x \ s" and "y \ s" - show "y \ path_component_set s x" + assume "x \ S" and "y \ S" + show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" - moreover have "path_component_set s x \ s \ {}" - using \x \ s\ path_component_eq_empty path_component_subset[of s x] + moreover have "path_component_set S x \ S \ {}" + using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False - using \y \ s\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] + using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, - of "path_component_set s x" "s - path_component_set s x"] + of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: - assumes "continuous_on s f" - and "path_connected s" - shows "path_connected (f ` s)" + assumes "continuous_on S f" + and "path_connected S" + shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' - assume "x' \ f ` s" "y' \ f ` s" - then obtain x y where x: "x \ s" and y: "y \ s" and x': "x' = f x" and y': "y' = f y" + assume "x' \ f ` S" "y' \ f ` S" + then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto - from x y obtain g where "path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y" + from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast - then show "\g. path g \ path_image g \ f ` s \ pathstart g = x' \ pathfinish g = y'" + then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' apply (rule_tac x="f \ g" in exI) unfolding path_defs @@ -1649,12 +1671,27 @@ done qed -lemma path_connected_segment: +lemma path_connected_translationI: + fixes a :: "'a :: topological_group_add" + assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" + by (intro path_connected_continuous_image assms continuous_intros) + +lemma path_connected_translation: + fixes a :: "'a :: topological_group_add" + shows "path_connected ((\x. a + x) ` S) = path_connected S" +proof - + have "\x y. op + (x::'a) ` op + (0 - x) ` y = y" + by (simp add: image_image) + then show ?thesis + by (metis (no_types) path_connected_translationI) +qed + +lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) -lemma path_connected_open_segment: +lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) @@ -1663,10 +1700,10 @@ "s homeomorphic t \ path_connected s \ path_connected t" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) -lemma path_connected_empty: "path_connected {}" +lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto -lemma path_connected_singleton: "path_connected {a}" +lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def apply clarify apply (rule_tac x="\x. a" in exI) @@ -1743,7 +1780,7 @@ unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) -lemma path_connected_path_component: +lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z @@ -1973,54 +2010,46 @@ by (simp add: path_connected_punctured_universe path_connected_imp_connected) lemma path_connected_sphere: - assumes "2 \ DIM('a::euclidean_space)" - shows "path_connected {x::'a. norm (x - a) = r}" -proof (rule linorder_cases [of r 0]) - assume "r < 0" - then have "{x::'a. norm(x - a) = r} = {}" - by auto + fixes a :: "'a :: euclidean_space" + assumes "2 \ DIM('a)" + shows "path_connected(sphere a r)" +proof (cases r "0::real" rule: linorder_cases) + case less then show ?thesis - using path_connected_empty by simp + by (simp add: path_connected_empty) next - assume "r = 0" + case equal then show ?thesis - using path_connected_singleton by simp + by (simp add: path_connected_singleton) next - assume r: "0 < r" - have *: "{x::'a. norm(x - a) = r} = (\x. a + r *\<^sub>R x) ` {x. norm x = 1}" - apply (rule set_eqI) - apply rule - unfolding image_iff - apply (rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI) - unfolding mem_Collect_eq norm_scaleR - using r - apply (auto simp add: scaleR_right_diff_distrib) - done - have **: "{x::'a. norm x = 1} = (\x. (1/norm x) *\<^sub>R x) ` (- {0})" - apply (rule set_eqI) - apply rule - unfolding image_iff - apply (rule_tac x=x in bexI) - unfolding mem_Collect_eq - apply (auto split: if_split_asm) - done - have "continuous_on (- {0}) (\x::'a. 1 / norm x)" - by (auto intro!: continuous_intros) + case greater + then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" + by (force simp: image_iff split: if_split_asm) + have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" + by (intro continuous_intros) auto + then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" + by (intro path_connected_continuous_image path_connected_punctured_universe assms) + with eq have "path_connected (sphere (0::'a) r)" + by auto + then have "path_connected(op +a ` (sphere (0::'a) r))" + by (simp add: path_connected_translation) then show ?thesis - unfolding * ** - using path_connected_punctured_universe[OF assms] - by (auto intro!: path_connected_continuous_image continuous_intros) -qed - -corollary connected_sphere: "2 \ DIM('a::euclidean_space) \ connected {x::'a. norm (x - a) = r}" - using path_connected_sphere path_connected_imp_connected - by auto + by (metis add.right_neutral sphere_translation) +qed + +lemma connected_sphere: + fixes a :: "'a :: euclidean_space" + assumes "2 \ DIM('a)" + shows "connected(sphere a r)" + using path_connected_sphere [OF assms] + by (simp add: path_connected_imp_connected) + corollary path_connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" and 2: "2 \ DIM('a)" shows "path_connected (- s)" -proof (cases "s={}") +proof (cases "s = {}") case True then show ?thesis using convex_imp_path_connected by auto next @@ -2104,11 +2133,11 @@ then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))" by (force simp: closed_segment_def intro!: path_connected_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" - apply (rule path_component_of_subset [of "{x. norm(x - a) = B}"]) + apply (rule path_component_of_subset [of "sphere a B"]) using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) - using \x \ a\ using \y \ a\ B apply (auto simp: C_def D_def) + using \x \ a\ using \y \ a\ B apply (auto simp: dist_norm C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) @@ -2193,6 +2222,52 @@ thus ?case by (metis Diff_insert) qed +lemma sphere_1D_doubleton_zero: + assumes 1: "DIM('a) = 1" and "r > 0" + obtains x y::"'a::euclidean_space" + where "sphere 0 r = {x,y} \ dist x y = 2*r" +proof - + obtain b::'a where b: "Basis = {b}" + using 1 card_1_singletonE by blast + show ?thesis + proof (intro that conjI) + have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x + proof - + have xb: "(x \ b) *\<^sub>R b = x" + using euclidean_representation [of x, unfolded b] by force + then have "norm ((x \ b) *\<^sub>R b) = norm x" + by simp + with b have "\x \ b\ = norm x" + using norm_Basis by fastforce + with xb show ?thesis + apply (simp add: abs_if split: if_split_asm) + apply (metis add.inverse_inverse real_vector.scale_minus_left xb) + done + qed + with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" + by (force simp: sphere_def dist_norm) + have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" + by (simp add: dist_norm) + also have "... = norm ((2*r) *\<^sub>R b)" + by (metis mult_2 scaleR_add_left) + also have "... = 2*r" + using \r > 0\ b norm_Basis by fastforce + finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . + qed +qed + +lemma sphere_1D_doubleton: + fixes a :: "'a :: euclidean_space" + assumes "DIM('a) = 1" and "r > 0" + obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" +proof - + have "sphere a r = op +a ` sphere 0 r" + by (metis add.right_neutral sphere_translation) + then show ?thesis + using sphere_1D_doubleton_zero [OF assms] + by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) +qed + lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" @@ -2219,6 +2294,7 @@ by (simp add: CS connected_Un) qed + subsection\Relations between components and path components\ lemma open_connected_component: @@ -2590,20 +2666,6 @@ by (metis bounded_subset [of "connected_component_set (- s) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) -lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)" - by (simp add: Int_commute frontier_def interior_closure) - -lemma frontier_interior_subset: "frontier(interior S) \ frontier S" - by (simp add: Diff_mono frontier_interiors interior_mono interior_subset) - -lemma connected_Int_frontier: - "\connected s; s \ t \ {}; s - t \ {}\ \ (s \ frontier t \ {})" - apply (simp add: frontier_interiors connected_openin, safe) - apply (drule_tac x="s \ interior t" in spec, safe) - apply (drule_tac [2] x="s \ interior (-t)" in spec) - apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) - done - lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" @@ -2735,7 +2797,7 @@ shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast -lemma components_univ [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" +lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: @@ -2837,6 +2899,26 @@ shows "\convex t; s \ t\ \ - t \ outside s" using outside_convex outside_mono by blast +lemma outside_Un_outside_Un: + fixes S :: "'a::real_normed_vector set" + assumes "S \ outside(T \ U) = {}" + shows "outside(T \ U) \ outside(T \ S)" +proof + fix x + assume x: "x \ outside (T \ U)" + have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y + proof - + have "Y \ connected_component_set (- (T \ U)) x" + by (simp add: connected_component_maximal that) + also have "... \ outside(T \ U)" + by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) + finally have "Y \ outside(T \ U)" . + with assms show ?thesis by auto + qed + with x show "x \ outside (T \ S)" + by (simp add: outside_connected_component_lt connected_component_def) meson +qed + lemma outside_frontier_misses_closure: fixes s :: "'a::real_normed_vector set" assumes "bounded s" @@ -3560,6 +3642,56 @@ by (metis homotopic_compose_continuous_right homotopic_with_imp_continuous homotopic_with_imp_subset2) +text\Homotopic triviality implicitly incorporates path-connectedness.\ +lemma homotopic_triviality: + fixes S :: "'a::real_normed_vector set" + shows "(\f g. continuous_on S f \ f ` S \ T \ + continuous_on S g \ g ` S \ T + \ homotopic_with (\x. True) S T f g) \ + (S = {} \ path_connected T) \ + (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with (\x. True) S T f (\x. c)))" + (is "?lhs = ?rhs") +proof (cases "S = {} \ T = {}") + case True then show ?thesis by auto +next + case False show ?thesis + proof + assume LHS [rule_format]: ?lhs + have pab: "path_component T a b" if "a \ T" "b \ T" for a b + proof - + have "homotopic_with (\x. True) S T (\x. a) (\x. b)" + by (simp add: LHS continuous_on_const image_subset_iff that) + then show ?thesis + using False homotopic_constant_maps by blast + qed + moreover + have "\c. homotopic_with (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f + by (metis (full_types) False LHS equals0I homotopic_constant_maps homotopic_with_imp_continuous homotopic_with_imp_subset2 pab that) + ultimately show ?rhs + by (simp add: path_connected_component) + next + assume RHS: ?rhs + with False have T: "path_connected T" + by blast + show ?lhs + proof clarify + fix f g + assume "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" + obtain c d where c: "homotopic_with (\x. True) S T f (\x. c)" and d: "homotopic_with (\x. True) S T g (\x. d)" + using False \continuous_on S f\ \f ` S \ T\ RHS \continuous_on S g\ \g ` S \ T\ by blast + then have "c \ T" "d \ T" + using False homotopic_with_imp_subset2 by fastforce+ + with T have "path_component T c d" + using path_connected_component by blast + then have "homotopic_with (\x. True) S T (\x. c) (\x. d)" + by (simp add: homotopic_constant_maps) + with c d show "homotopic_with (\x. True) S T f g" + by (meson homotopic_with_symD homotopic_with_trans) + qed + qed +qed + + subsection\Homotopy of paths, maintaining the same endpoints.\ @@ -4519,7 +4651,7 @@ shows "starlike S \ contractible S" using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) -lemma contractible_UNIV: "contractible (UNIV :: 'a::real_normed_vector set)" +lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" by (simp add: starlike_imp_contractible) lemma starlike_imp_simply_connected: @@ -4547,8 +4679,8 @@ shows "is_interval S \ simply_connected S" using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto -lemma contractible_empty: "contractible {}" - by (simp add: continuous_on_empty contractible_def homotopic_with) +lemma contractible_empty [simp]: "contractible {}" + by (simp add: contractible_def homotopic_with) lemma contractible_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" @@ -4557,7 +4689,7 @@ proof (cases "S = {}") case True with assms show ?thesis - by (simp add: contractible_empty subsetCE) + by (simp add: subsetCE) next case False show ?thesis @@ -4569,9 +4701,9 @@ lemma convex_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "convex S \ contractible S" -using contractible_empty convex_imp_starlike starlike_imp_contractible by auto - -lemma contractible_sing: + using contractible_empty convex_imp_starlike starlike_imp_contractible by blast + +lemma contractible_sing [simp]: fixes a :: "'a::real_normed_vector" shows "contractible {a}" by (rule convex_imp_contractible [OF convex_singleton]) @@ -4705,6 +4837,28 @@ apply (auto intro: P) done +lemma locally_Times: + fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" + assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" + shows "locally R (S \ T)" + unfolding locally_def +proof (clarify) + fix W x y + assume W: "openin (subtopology euclidean (S \ T)) W" and xy: "(x, y) \ W" + then obtain U V where "openin (subtopology euclidean S) U" "x \ U" + "openin (subtopology euclidean T) V" "y \ V" "U \ V \ W" + using Times_in_interior_subtopology by metis + then obtain U1 U2 V1 V2 + where opeS: "openin (subtopology euclidean S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" + and opeT: "openin (subtopology euclidean T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" + by (meson PS QT locallyE) + with \U \ V \ W\ show "\u v. openin (subtopology euclidean (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" + apply (rule_tac x="U1 \ V1" in exI) + apply (rule_tac x="U2 \ V2" in exI) + apply (auto simp: openin_Times R) + done +qed + proposition homeomorphism_locally_imp: fixes S :: "'a::metric_space set" and t :: "'b::t2_space set" @@ -5148,6 +5302,106 @@ using compact_eq_bounded_closed locally_mono by blast qed +lemma locally_compact_openin_Un: + fixes S :: "'a::euclidean_space set" + assumes LCS: "locally compact S" and LCT:"locally compact T" + and opS: "openin (subtopology euclidean (S \ T)) S" + and opT: "openin (subtopology euclidean (S \ T)) T" + shows "locally compact (S \ T)" +proof - + have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" + by (meson \x \ S\ opS openin_contains_cball) + then have "cball x e2 \ (S \ T) = cball x e2 \ S" + by force + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) + by (metis closed_Int closed_cball inf_left_commute) + qed + moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" + by (meson \x \ T\ opT openin_contains_cball) + then have "cball x e2 \ (S \ T) = cball x e2 \ T" + by force + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) + by (metis closed_Int closed_cball inf_left_commute) + qed + ultimately show ?thesis + by (force simp: locally_compact_Int_cball) +qed + +lemma locally_compact_closedin_Un: + fixes S :: "'a::euclidean_space set" + assumes LCS: "locally compact S" and LCT:"locally compact T" + and clS: "closedin (subtopology euclidean (S \ T)) S" + and clT: "closedin (subtopology euclidean (S \ T)) T" + shows "locally compact (S \ T)" +proof - + have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + moreover + have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" + using clT x by (fastforce simp: openin_contains_cball closedin_def) + then have "closed (cball x e2 \ T)" + proof - + have "{} = T - (T - cball x e2)" + using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto + then show ?thesis + by (simp add: Diff_Diff_Int inf_commute) + qed + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + moreover + have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" + using clS x by (fastforce simp: openin_contains_cball closedin_def) + then have "closed (cball x e2 \ S)" + by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + ultimately show ?thesis + by (auto simp: locally_compact_Int_cball) +qed + +lemma locally_compact_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" + by (auto simp: compact_Times locally_Times) + subsection\Important special cases of local connectedness and path connectedness\ lemma locally_connected_1: @@ -6519,6 +6773,64 @@ by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) qed +lemma connected_uncountable: + fixes S :: "'a::metric_space set" + assumes "connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" +proof - + have "continuous_on S (dist a)" + by (intro continuous_intros) + then have "connected (dist a ` S)" + by (metis connected_continuous_image \connected S\) + then have "closed_segment 0 (dist a b) \ (dist a ` S)" + by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) + then have "uncountable (dist a ` S)" + by (metis \a \ b\ countable_subset dist_eq_0_iff uncountable_closed_segment) + then show ?thesis + by blast +qed + +lemma path_connected_uncountable: + fixes S :: "'a::metric_space set" + assumes "path_connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" + using path_connected_imp_connected assms connected_uncountable by metis + +lemma connected_finite_iff_sing: + fixes S :: "'a::metric_space set" + assumes "connected S" + shows "finite S \ S = {} \ (\a. S = {a})" (is "_ = ?rhs") +proof - + have "uncountable S" if "\ ?rhs" + using connected_uncountable assms that by blast + then show ?thesis + using uncountable_infinite by auto +qed + +lemma connected_card_eq_iff_nontrivial: + fixes S :: "'a::metric_space set" + shows "connected S \ uncountable S \ ~(\a. S \ {a})" + apply (auto simp: countable_finite finite_subset) + by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) + +lemma simple_path_image_uncountable: + fixes g :: "real \ 'a::metric_space" + assumes "simple_path g" + shows "uncountable (path_image g)" +proof - + have "g 0 \ path_image g" "g (1/2) \ path_image g" + by (simp_all add: path_defs) + moreover have "g 0 \ g (1/2)" + using assms by (fastforce simp add: simple_path_def) + ultimately show ?thesis + apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image) + by blast +qed + +lemma arc_image_uncountable: + fixes g :: "real \ 'a::metric_space" + assumes "arc g" + shows "uncountable (path_image g)" + by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) + subsection\ Some simple positive connection theorems\ diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Jan 05 22:38:06 2017 +0100 @@ -15,24 +15,6 @@ Norm_Arith begin -(* FIXME: move to HOL/Real_Vector_Spaces.thy *) - -lemma scaleR_2: - fixes x :: "'a::real_vector" - shows "scaleR 2 x = x + x" - unfolding one_add_one [symmetric] scaleR_left_distrib by simp - -lemma scaleR_half_double [simp]: - fixes a :: "'a::real_normed_vector" - shows "(1 / 2) *\<^sub>R (a + a) = a" -proof - - have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" - by (metis scaleR_2 scaleR_scaleR) - then show ?thesis - by simp -qed - - (* FIXME: move elsewhere *) lemma Times_eq_image_sum: @@ -894,6 +876,11 @@ lemma closed_subset: "S \ T \ closed S \ closedin (subtopology euclidean T) S" by (auto simp add: closedin_closed) +lemma closedin_closed_subset: + "\closedin (subtopology euclidean U) V; T \ U; S = V \ T\ + \ closedin (subtopology euclidean T) S" + by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) + lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (subtopology euclidean T) S" @@ -1099,7 +1086,7 @@ lemma cball_min_Int: "cball a (min r s) = cball a r \ cball a s" by (simp add: set_eq_iff) -lemma cball_diff_eq_sphere: "cball a r - ball a r = {x. dist x a = r}" +lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" by (auto simp: cball_def ball_def dist_commute) lemma image_add_ball [simp]: @@ -2744,6 +2731,28 @@ lemma frontier_UNIV [simp]: "frontier UNIV = {}" using frontier_complement frontier_empty by fastforce +lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)" + by (simp add: Int_commute frontier_def interior_closure) + +lemma frontier_interior_subset: "frontier(interior S) \ frontier S" + by (simp add: Diff_mono frontier_interiors interior_mono interior_subset) + +lemma connected_Int_frontier: + "\connected s; s \ t \ {}; s - t \ {}\ \ (s \ frontier t \ {})" + apply (simp add: frontier_interiors connected_openin, safe) + apply (drule_tac x="s \ interior t" in spec, safe) + apply (drule_tac [2] x="s \ interior (-t)" in spec) + apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) + done + +lemma closure_Un_frontier: "closure S = S \ frontier S" +proof - + have "S \ interior S = S" + using interior_subset by auto + then show ?thesis + using closure_subset by (auto simp: frontier_def) +qed + subsection \Filters and the ``eventually true'' quantifier\ diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Library/Polynomial.thy Thu Jan 05 22:38:06 2017 +0100 @@ -12,6 +12,21 @@ "~~/src/HOL/Library/Infinite_Set" begin +subsection \Misc\ + +lemma quotient_of_denom_pos': "snd (quotient_of x) > 0" + using quotient_of_denom_pos [OF surjective_pairing] . + +lemma of_int_divide_in_Ints: + "b dvd a \ of_int a div of_int b \ (\ :: 'a :: idom_divide set)" +proof (cases "of_int b = (0 :: 'a)") + case False + assume "b dvd a" + then obtain c where "a = b * c" .. + with \of_int b \ (0::'a)\ show ?thesis by simp +qed auto + + subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) @@ -143,6 +158,33 @@ "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0", simp, simp add: leading_coeff_neq_0) +lemma eq_zero_or_degree_less: + assumes "degree p \ n" and "coeff p n = 0" + shows "p = 0 \ degree p < n" +proof (cases n) + case 0 + with \degree p \ n\ and \coeff p n = 0\ + have "coeff p (degree p) = 0" by simp + then have "p = 0" by simp + then show ?thesis .. +next + case (Suc m) + have "\i>n. coeff p i = 0" + using \degree p \ n\ by (simp add: coeff_eq_0) + then have "\i\n. coeff p i = 0" + using \coeff p n = 0\ by (simp add: le_less) + then have "\i>m. coeff p i = 0" + using \n = Suc m\ by (simp add: less_eq_Suc_le) + then have "degree p \ m" + by (rule degree_le) + then have "degree p < n" + using \n = Suc m\ by (simp add: less_Suc_eq_le) + then show ?thesis .. +qed + +lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" + using eq_zero_or_degree_less by fastforce + subsection \List-style constructor for polynomials\ @@ -481,6 +523,7 @@ "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) + subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" @@ -572,8 +615,22 @@ lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) - - + + +subsection \Leading coefficient\ + +abbreviation lead_coeff:: "'a::zero poly \ 'a" + where "lead_coeff p \ coeff p (degree p)" + +lemma lead_coeff_pCons[simp]: + "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" + "p = 0 \ lead_coeff (pCons a p) = a" + by auto + +lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" + by (cases "c = 0") (simp_all add: degree_monom_eq) + + subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add @@ -694,6 +751,16 @@ "degree (- p) = degree p" unfolding degree_def by simp +lemma lead_coeff_add_le: + assumes "degree p < degree q" + shows "lead_coeff (p + q) = lead_coeff q" + using assms + by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) + +lemma lead_coeff_minus: + "lead_coeff (- p) = - lead_coeff p" + by (metis coeff_minus degree_minus) + lemma degree_diff_le_max: fixes p q :: "'a :: ab_group_add poly" shows "degree (p - q) \ max (degree p) (degree q)" @@ -894,7 +961,16 @@ shows "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" by (rule coeffs_eqI) (auto simp add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) - + +lemma smult_eq_iff: + assumes "(b :: 'a :: field) \ 0" + shows "smult a p = smult b q \ smult (a / b) p = q" +proof + assume "smult a p = smult b q" + also from assms have "smult (inverse b) \ = q" by simp + finally show "smult (a / b) p = q" by (simp add: field_simps) +qed (insert assms, auto) + instantiation poly :: (comm_semiring_0) comm_semiring_0 begin @@ -1037,6 +1113,10 @@ "degree (p ^ n) \ degree p * n" by (induct n) (auto intro: order_trans degree_mult_le) +lemma coeff_0_power: + "coeff (p ^ n) 0 = coeff p 0 ^ n" + by (induction n) (simp_all add: coeff_mult) + lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" by (induct p, simp, simp add: algebra_simps) @@ -1064,6 +1144,40 @@ by (rule le_trans[OF degree_mult_le], insert insert, auto) qed simp +lemma coeff_0_prod_list: + "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" + by (induction xs) (simp_all add: coeff_mult) + +lemma coeff_monom_mult: + "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" +proof - + have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" + by (simp add: coeff_mult) + also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" + by (intro sum.cong) simp_all + also have "\ = (if k < n then 0 else c * coeff p (k - n))" by (simp add: sum.delta') + finally show ?thesis . +qed + +lemma monom_1_dvd_iff': + "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" + have "\\<^sub>\k. coeff p (k + n) = 0" + by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, + subst cofinite_eq_sequentially [symmetric]) transfer + hence coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def + by (subst poly.Abs_poly_inverse) simp_all + have "p = monom 1 n * r" + by (intro poly_eqI, subst coeff_monom_mult) (insert zero, simp_all) + thus "monom 1 n dvd p" by simp +qed + subsection \Mapping polynomials\ @@ -1169,7 +1283,7 @@ by (intro poly_eqI) (simp_all add: coeff_map_poly) -subsection \Conversions from natural numbers\ +subsection \Conversions from @{typ nat} and @{typ int} numbers\ lemma of_nat_poly: "of_nat n = [:of_nat n :: 'a :: comm_semiring_1:]" proof (induction n) @@ -1185,11 +1299,29 @@ lemma degree_of_nat [simp]: "degree (of_nat n) = 0" by (simp add: of_nat_poly) -lemma degree_numeral [simp]: "degree (numeral n) = 0" - by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp +lemma lead_coeff_of_nat [simp]: + "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})" + by (simp add: of_nat_poly) + +lemma of_int_poly: "of_int k = [:of_int k :: 'a :: comm_ring_1:]" + by (simp only: of_int_of_nat of_nat_poly) simp + +lemma degree_of_int [simp]: "degree (of_int k) = 0" + by (simp add: of_int_poly) + +lemma lead_coeff_of_int [simp]: + "lead_coeff (of_int k) = (of_int k :: 'a :: {comm_ring_1,ring_char_0})" + by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp + +lemma degree_numeral [simp]: "degree (numeral n) = 0" + by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp + +lemma lead_coeff_numeral [simp]: + "lead_coeff (numeral n) = numeral n" + by (simp add: numeral_poly) subsection \Lemmas about divisibility\ @@ -1231,6 +1363,28 @@ shows "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" by (auto elim: smult_dvd smult_dvd_cancel) +lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" +proof - + have "smult c p = [:c:] * p" by simp + also have "\ dvd 1 \ c dvd 1 \ p dvd 1" + proof safe + assume A: "[:c:] * p dvd 1" + thus "p dvd 1" by (rule dvd_mult_right) + from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE) + have "c dvd c * (coeff p 0 * coeff q 0)" by simp + also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) + also note B [symmetric] + finally show "c dvd 1" by simp + next + assume "c dvd 1" "p dvd 1" + from \c dvd 1\ obtain d where "1 = c * d" by (erule dvdE) + hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac) + hence "[:c:] dvd 1" by (rule dvdI) + from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp + qed + finally show ?thesis . +qed + subsection \Polynomials form an integral domain\ @@ -1296,6 +1450,27 @@ "[:a::'a::{comm_semiring_1,semiring_no_zero_divisors}:] dvd [:b:] \ a dvd b" by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) +lemma lead_coeff_mult: + fixes p q :: "'a :: {comm_semiring_0, semiring_no_zero_divisors} poly" + shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" + by (cases "p=0 \ q=0", auto simp add:coeff_mult_degree_sum degree_mult_eq) + +lemma lead_coeff_smult: + "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p" +proof - + have "smult c p = [:c:] * p" by simp + also have "lead_coeff \ = c * lead_coeff p" + by (subst lead_coeff_mult) simp_all + finally show ?thesis . +qed + +lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" + by simp + +lemma lead_coeff_power: + "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n" + by (induction n) (simp_all add: lead_coeff_mult) + subsection \Polynomials form an ordered integral domain\ @@ -1401,77 +1576,10 @@ text \TODO: Simplification rules for comparisons\ -subsection \Leading coefficient\ - -definition lead_coeff:: "'a::zero poly \ 'a" where - "lead_coeff p= coeff p (degree p)" - -lemma lead_coeff_pCons[simp]: - "p\0 \lead_coeff (pCons a p) = lead_coeff p" - "p=0 \ lead_coeff (pCons a p) = a" -unfolding lead_coeff_def by auto - -lemma lead_coeff_0[simp]:"lead_coeff 0 =0" - unfolding lead_coeff_def by auto - -lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" - by (induction xs) (simp_all add: coeff_mult) - -lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" - by (induction n) (simp_all add: coeff_mult) - -lemma lead_coeff_mult: - fixes p q::"'a :: {comm_semiring_0,semiring_no_zero_divisors} poly" - shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" -by (unfold lead_coeff_def,cases "p=0 \ q=0",auto simp add:coeff_mult_degree_sum degree_mult_eq) - -lemma lead_coeff_add_le: - assumes "degree p < degree q" - shows "lead_coeff (p+q) = lead_coeff q" -using assms unfolding lead_coeff_def -by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) - -lemma lead_coeff_minus: - "lead_coeff (-p) = - lead_coeff p" -by (metis coeff_minus degree_minus lead_coeff_def) - -lemma lead_coeff_smult: - "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p" -proof - - have "smult c p = [:c:] * p" by simp - also have "lead_coeff \ = c * lead_coeff p" - by (subst lead_coeff_mult) simp_all - finally show ?thesis . -qed - -lemma lead_coeff_eq_zero_iff [simp]: "lead_coeff p = 0 \ p = 0" - by (simp add: lead_coeff_def) - -lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" - by (simp add: lead_coeff_def) - -lemma lead_coeff_of_nat [simp]: - "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})" - by (induction n) (simp_all add: lead_coeff_def of_nat_poly) - -lemma lead_coeff_numeral [simp]: - "lead_coeff (numeral n) = numeral n" - unfolding lead_coeff_def - by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp - -lemma lead_coeff_power: - "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n" - by (induction n) (simp_all add: lead_coeff_mult) - -lemma lead_coeff_nonzero: "p \ 0 \ lead_coeff p \ 0" - by (simp add: lead_coeff_def) - -lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" - by (cases "c = 0") (simp_all add: lead_coeff_def degree_monom_eq) - - subsection \Synthetic division and polynomial roots\ +subsubsection \Synthetic division\ + text \ Synthetic division is simply division by the linear polynomial @{term "x - c"}. \ @@ -1539,9 +1647,12 @@ using synthetic_div_correct [of p c] by (simp add: algebra_simps) + +subsubsection \Polynomial roots\ + lemma poly_eq_0_iff_dvd: fixes c :: "'a::{comm_ring_1}" - shows "poly p c = 0 \ [:-c, 1:] dvd p" + shows "poly p c = 0 \ [:- c, 1:] dvd p" proof assume "poly p c = 0" with synthetic_div_correct' [of c p] @@ -1555,7 +1666,7 @@ lemma dvd_iff_poly_eq_0: fixes c :: "'a::{comm_ring_1}" - shows "[:c, 1:] dvd p \ poly p (-c) = 0" + shows "[:c, 1:] dvd p \ poly p (- c) = 0" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: @@ -1610,1318 +1721,8 @@ shows "(\x. poly p x = 0) \ p = 0" by (auto simp add: poly_eq_poly_eq_iff [symmetric]) - -subsection \Long division of polynomials\ - -inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" - where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" - | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" - | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y - \ x = q * y + r \ eucl_rel_poly x y (q, r)" - -lemma eucl_rel_poly_iff: - "eucl_rel_poly x y (q, r) \ - x = q * y + r \ - (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" - by (auto elim: eucl_rel_poly.cases - intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) - -lemma eucl_rel_poly_0: - "eucl_rel_poly 0 y (0, 0)" - unfolding eucl_rel_poly_iff by simp - -lemma eucl_rel_poly_by_0: - "eucl_rel_poly x 0 (0, x)" - unfolding eucl_rel_poly_iff by simp - -lemma eq_zero_or_degree_less: - assumes "degree p \ n" and "coeff p n = 0" - shows "p = 0 \ degree p < n" -proof (cases n) - case 0 - with \degree p \ n\ and \coeff p n = 0\ - have "coeff p (degree p) = 0" by simp - then have "p = 0" by simp - then show ?thesis .. -next - case (Suc m) - have "\i>n. coeff p i = 0" - using \degree p \ n\ by (simp add: coeff_eq_0) - then have "\i\n. coeff p i = 0" - using \coeff p n = 0\ by (simp add: le_less) - then have "\i>m. coeff p i = 0" - using \n = Suc m\ by (simp add: less_eq_Suc_le) - then have "degree p \ m" - by (rule degree_le) - then have "degree p < n" - using \n = Suc m\ by (simp add: less_Suc_eq_le) - then show ?thesis .. -qed - -lemma eucl_rel_poly_pCons: - assumes rel: "eucl_rel_poly x y (q, r)" - assumes y: "y \ 0" - assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" - shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" - (is "eucl_rel_poly ?x y (?q, ?r)") -proof - - have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" - using assms unfolding eucl_rel_poly_iff by simp_all - - have 1: "?x = ?q * y + ?r" - using b x by simp - - have 2: "?r = 0 \ degree ?r < degree y" - proof (rule eq_zero_or_degree_less) - show "degree ?r \ degree y" - proof (rule degree_diff_le) - show "degree (pCons a r) \ degree y" - using r by auto - show "degree (smult b y) \ degree y" - by (rule degree_smult_le) - qed - next - show "coeff ?r (degree y) = 0" - using \y \ 0\ unfolding b by simp - qed - - from 1 2 show ?thesis - unfolding eucl_rel_poly_iff - using \y \ 0\ by simp -qed - -lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" -apply (cases "y = 0") -apply (fast intro!: eucl_rel_poly_by_0) -apply (induct x) -apply (fast intro!: eucl_rel_poly_0) -apply (fast intro!: eucl_rel_poly_pCons) -done - -lemma eucl_rel_poly_unique: - assumes 1: "eucl_rel_poly x y (q1, r1)" - assumes 2: "eucl_rel_poly x y (q2, r2)" - shows "q1 = q2 \ r1 = r2" -proof (cases "y = 0") - assume "y = 0" with assms show ?thesis - by (simp add: eucl_rel_poly_iff) -next - assume [simp]: "y \ 0" - from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" - unfolding eucl_rel_poly_iff by simp_all - from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" - unfolding eucl_rel_poly_iff by simp_all - from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" - by (simp add: algebra_simps) - from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" - by (auto intro: degree_diff_less) - - show "q1 = q2 \ r1 = r2" - proof (rule ccontr) - assume "\ (q1 = q2 \ r1 = r2)" - with q3 have "q1 \ q2" and "r1 \ r2" by auto - with r3 have "degree (r2 - r1) < degree y" by simp - also have "degree y \ degree (q1 - q2) + degree y" by simp - also have "\ = degree ((q1 - q2) * y)" - using \q1 \ q2\ by (simp add: degree_mult_eq) - also have "\ = degree (r2 - r1)" - using q3 by simp - finally have "degree (r2 - r1) < degree (r2 - r1)" . - then show "False" by simp - qed -qed - -lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" -by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) - -lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" -by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) - -lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] - -lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] - - - -subsection \Pseudo-Division and Division of Polynomials\ - -text\This part is by René Thiemann and Akihisa Yamada.\ - -fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly - \ nat \ nat \ 'a poly \ 'a poly" where - "pseudo_divmod_main lc q r d dr (Suc n) = (let - rr = smult lc r; - qq = coeff r dr; - rrr = rr - monom qq n * d; - qqq = smult lc q + monom qq n - in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" -| "pseudo_divmod_main lc q r d dr 0 = (q,r)" - -definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where - "pseudo_divmod p q \ if q = 0 then (0,p) else - pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" - -lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" - using eq_zero_or_degree_less by fastforce - -lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" - and *: "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" - "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" - shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" - using * -proof (induct n arbitrary: q r dr) - case (Suc n q r dr) - let ?rr = "smult lc r" - let ?qq = "coeff r dr" - define b where [simp]: "b = monom ?qq n" - let ?rrr = "?rr - b * d" - let ?qqq = "smult lc q + b" - note res = Suc(3) - from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] - have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" - by (simp del: pseudo_divmod_main.simps) - have dr: "dr = n + degree d" using Suc(4) by auto - have "coeff (b * d) dr = coeff b n * coeff d (degree d)" - proof (cases "?qq = 0") - case False - hence n: "n = degree b" by (simp add: degree_monom_eq) - show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) - qed auto - also have "\ = lc * coeff b n" unfolding d by simp - finally have "coeff (b * d) dr = lc * coeff b n" . - moreover have "coeff ?rr dr = lc * coeff r dr" by simp - ultimately have c0: "coeff ?rrr dr = 0" by auto - have dr: "dr = n + degree d" using Suc(4) by auto - have deg_rr: "degree ?rr \ dr" using Suc(2) - using degree_smult_le dual_order.trans by blast - have deg_bd: "degree (b * d) \ dr" - unfolding dr - by(rule order.trans[OF degree_mult_le], auto simp: degree_monom_le) - have "degree ?rrr \ dr" - using degree_diff_le[OF deg_rr deg_bd] by auto - with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) - have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" - proof (cases dr) - case 0 - with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto - with deg_rrr have "degree ?rrr = 0" by simp - hence "\ a. ?rrr = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) - from this obtain a where rrr: "?rrr = [:a:]" by auto - show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp - qed (insert Suc(4), auto) - note IH = Suc(1)[OF deg_rrr res this] - show ?case - proof (intro conjI) - show "r' = 0 \ degree r' < degree d" using IH by blast - show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" - unfolding IH[THEN conjunct2,symmetric] - by (simp add: field_simps smult_add_right) - qed -qed auto - -lemma pseudo_divmod: - assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" - shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) - and "r = 0 \ degree r < degree g" (is ?B) -proof - - from *[unfolded pseudo_divmod_def Let_def] - have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) - note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] - have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ - degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" using g - by (cases "f = 0"; cases "coeffs g", auto simp: degree_eq_length_coeffs) - note main = main[OF this] - from main show "r = 0 \ degree r < degree g" by auto - show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" - by (subst main[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, - insert g, cases "f = 0"; cases "coeffs g", auto) -qed - -definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" - -lemma snd_pseudo_divmod_main: - "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" -by (induct n arbitrary: q q' lc r d dr; simp add: Let_def) - -definition pseudo_mod - :: "'a :: {comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where - "pseudo_mod f g = snd (pseudo_divmod f g)" - -lemma pseudo_mod: - fixes f g - defines "r \ pseudo_mod f g" - assumes g: "g \ 0" - shows "\ a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" -proof - - let ?cg = "coeff g (degree g)" - let ?cge = "?cg ^ (Suc (degree f) - degree g)" - define a where "a = ?cge" - obtain q where pdm: "pseudo_divmod f g = (q,r)" using r_def[unfolded pseudo_mod_def] - by (cases "pseudo_divmod f g", auto) - from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" - unfolding a_def by auto - show "r = 0 \ degree r < degree g" by fact - from g have "a \ 0" unfolding a_def by auto - thus "\ a q. a \ 0 \ smult a f = g * q + r" using id by auto -qed - -instantiation poly :: (idom_divide) idom_divide -begin - -fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly - \ nat \ nat \ 'a poly" where - "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in - if False \ a * lc = cr then (* False \ is only because of problem in function-package *) - divide_poly_main - lc - (q + mon) - (r - mon * d) - d (dr - 1) n else 0)" -| "divide_poly_main lc q r d dr 0 = q" - -definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where - "divide_poly f g = (if g = 0 then 0 else - divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" - -lemma divide_poly_main: - assumes d: "d \ 0" "lc = coeff d (degree d)" - and *: "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" - "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" - shows "q' = q + r" - using * -proof (induct n arbitrary: q r dr) - case (Suc n q r dr) - let ?rr = "d * r" - let ?a = "coeff ?rr dr" - let ?qq = "?a div lc" - define b where [simp]: "b = monom ?qq n" - let ?rrr = "d * (r - b)" - let ?qqq = "q + b" - note res = Suc(3) - have dr: "dr = n + degree d" using Suc(4) by auto - have lc: "lc \ 0" using d by auto - have "coeff (b * d) dr = coeff b n * coeff d (degree d)" - proof (cases "?qq = 0") - case False - hence n: "n = degree b" by (simp add: degree_monom_eq) - show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) - qed simp - also have "\ = lc * coeff b n" unfolding d by simp - finally have c2: "coeff (b * d) dr = lc * coeff b n" . - have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) - have c1: "coeff (d * r) dr = lc * coeff r n" - proof (cases "degree r = n") - case True - thus ?thesis using Suc(2) unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) - next - case False - have "degree r \ n" using dr Suc(2) by auto - (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) - with False have r_n: "degree r < n" by auto - hence right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) - have "coeff (d * r) dr = coeff (d * r) (degree d + n)" unfolding dr by (simp add: ac_simps) - also have "\ = 0" using r_n - by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 - coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) - finally show ?thesis unfolding right . - qed - have c0: "coeff ?rrr dr = 0" - and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 - unfolding b_def coeff_monom coeff_smult c1 using lc by auto - from res[unfolded divide_poly_main.simps[of lc q] Let_def] id - have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" - by (simp del: divide_poly_main.simps add: field_simps) - note IH = Suc(1)[OF _ res] - have dr: "dr = n + degree d" using Suc(4) by auto - have deg_rr: "degree ?rr \ dr" using Suc(2) by auto - have deg_bd: "degree (b * d) \ dr" - unfolding dr b_def by (rule order.trans[OF degree_mult_le], auto simp: degree_monom_le) - have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) - with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) - have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" - proof (cases dr) - case 0 - with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto - with deg_rrr have "degree ?rrr = 0" by simp - from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis - show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp - qed (insert Suc(4), auto) - note IH = IH[OF deg_rrr this] - show ?case using IH by simp -next - case (0 q r dr) - show ?case - proof (cases "r = 0") - case True - thus ?thesis using 0 by auto - next - case False - have "degree (d * r) = degree d + degree r" using d False - by (subst degree_mult_eq, auto) - thus ?thesis using 0 d by auto - qed -qed - -lemma fst_pseudo_divmod_main_as_divide_poly_main: - assumes d: "d \ 0" - defines lc: "lc \ coeff d (degree d)" - shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" -proof(induct n arbitrary: q r dr) - case 0 then show ?case by simp -next - case (Suc n) - note lc0 = leading_coeff_neq_0[OF d, folded lc] - then have "pseudo_divmod_main lc q r d dr (Suc n) = - pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) - (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" - by (simp add: Let_def ac_simps) - also have "fst ... = divide_poly_main lc - (smult (lc^n) (smult lc q + monom (coeff r dr) n)) - (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) - d (dr - 1) n" - unfolding Suc[unfolded divide_poly_main.simps Let_def].. - also have "... = divide_poly_main lc (smult (lc ^ Suc n) q) - (smult (lc ^ Suc n) r) d dr (Suc n)" - unfolding smult_monom smult_distribs mult_smult_left[symmetric] - using lc0 by (simp add: Let_def ac_simps) - finally show ?case. -qed - - -lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" -proof (induct n arbitrary: r d dr) - case (Suc n r d dr) - show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def - by (simp add: Suc del: divide_poly_main.simps) -qed simp - -lemma divide_poly: - assumes g: "g \ 0" - shows "(f * g) div g = (f :: 'a poly)" -proof - - have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) - = (f * g) div g" unfolding divide_poly_def Let_def by (simp add: ac_simps) - note main = divide_poly_main[OF g refl le_refl this] - { - fix f :: "'a poly" - assume "f \ 0" - hence "length (coeffs f) = Suc (degree f)" unfolding degree_eq_length_coeffs by auto - } note len = this - have "(f * g) div g = 0 + f" - proof (rule main, goal_cases) - case 1 - show ?case - proof (cases "f = 0") - case True - with g show ?thesis by (auto simp: degree_eq_length_coeffs) - next - case False - with g have fg: "g * f \ 0" by auto - show ?thesis unfolding len[OF fg] len[OF g] by auto - qed - qed - thus ?thesis by simp -qed - -lemma divide_poly_0: "f div 0 = (0 :: 'a poly)" - by (simp add: divide_poly_def Let_def divide_poly_main_0) - -instance - by standard (auto simp: divide_poly divide_poly_0) - -end - -instance poly :: (idom_divide) algebraic_semidom .. - -lemma div_const_poly_conv_map_poly: - assumes "[:c:] dvd p" - shows "p div [:c:] = map_poly (\x. x div c) p" -proof (cases "c = 0") - case False - from assms obtain q where p: "p = [:c:] * q" by (erule dvdE) - moreover { - have "smult c q = [:c:] * q" by simp - also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto) - finally have "smult c q div [:c:] = q" . - } - ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) -qed (auto intro!: poly_eqI simp: coeff_map_poly) - -lemma is_unit_monom_0: - fixes a :: "'a::field" - assumes "a \ 0" - shows "is_unit (monom a 0)" -proof - from assms show "1 = monom a 0 * monom (inverse a) 0" - by (simp add: mult_monom) -qed - -lemma is_unit_triv: - fixes a :: "'a::field" - assumes "a \ 0" - shows "is_unit [:a:]" - using assms by (simp add: is_unit_monom_0 monom_0 [symmetric]) - -lemma is_unit_iff_degree: - assumes "p \ (0 :: _ :: field poly)" - shows "is_unit p \ degree p = 0" (is "?P \ ?Q") -proof - assume ?Q - then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) - with assms show ?P by (simp add: is_unit_triv) -next - assume ?P - then obtain q where "q \ 0" "p * q = 1" .. - then have "degree (p * q) = degree 1" - by simp - with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" - by (simp add: degree_mult_eq) - then show ?Q by simp -qed - -lemma is_unit_pCons_iff: - "is_unit (pCons (a::_::field) p) \ p = 0 \ a \ 0" - by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree) - -lemma is_unit_monom_trival: - fixes p :: "'a::field poly" - assumes "is_unit p" - shows "monom (coeff p (degree p)) 0 = p" - using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) - -lemma is_unit_const_poly_iff: - "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \ c dvd 1" - by (auto simp: one_poly_def) - -lemma is_unit_polyE: - fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" - assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" -proof - - from assms obtain q where "1 = p * q" - by (rule dvdE) - then have "p \ 0" and "q \ 0" - by auto - from \1 = p * q\ have "degree 1 = degree (p * q)" - by simp - also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" - by (simp add: degree_mult_eq) - finally have "degree p = 0" by simp - with degree_eq_zeroE obtain c where c: "p = [:c:]" . - moreover with \p dvd 1\ have "c dvd 1" - by (simp add: is_unit_const_poly_iff) - ultimately show thesis - by (rule that) -qed - -lemma is_unit_polyE': - assumes "is_unit (p::_::field poly)" - obtains a where "p = monom a 0" and "a \ 0" -proof - - obtain a q where "p = pCons a q" by (cases p) - with assms have "p = [:a:]" and "a \ 0" - by (simp_all add: is_unit_pCons_iff) - with that show thesis by (simp add: monom_0) -qed - -lemma is_unit_poly_iff: - fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" - shows "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" - by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) - -instantiation poly :: ("{normalization_semidom, idom_divide}") normalization_semidom -begin - -definition unit_factor_poly :: "'a poly \ 'a poly" - where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0" - -definition normalize_poly :: "'a poly \ 'a poly" - where "normalize_poly p = map_poly (\x. x div unit_factor (lead_coeff p)) p" - -instance proof - fix p :: "'a poly" - show "unit_factor p * normalize p = p" - by (cases "p = 0") - (simp_all add: unit_factor_poly_def normalize_poly_def monom_0 - smult_conv_map_poly map_poly_map_poly o_def) -next - fix p :: "'a poly" - assume "is_unit p" - then obtain c where p: "p = [:c:]" "is_unit c" - by (auto simp: is_unit_poly_iff) - thus "normalize p = 1" - by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def) -next - fix p :: "'a poly" assume "p \ 0" - thus "is_unit (unit_factor p)" - by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff) -qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) - -end - -lemma normalize_poly_eq_div: - "normalize p = p div [:unit_factor (lead_coeff p):]" -proof (cases "p = 0") - case False - thus ?thesis - by (subst div_const_poly_conv_map_poly) - (auto simp: normalize_poly_def const_poly_dvd_iff lead_coeff_def ) -qed (auto simp: normalize_poly_def) - -lemma unit_factor_pCons: - "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)" - by (simp add: unit_factor_poly_def) - -lemma normalize_monom [simp]: - "normalize (monom a n) = monom (normalize a) n" - by (simp add: map_poly_monom normalize_poly_def) - -lemma unit_factor_monom [simp]: - "unit_factor (monom a n) = monom (unit_factor a) 0" - by (simp add: unit_factor_poly_def ) - -lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" - by (simp add: normalize_poly_def map_poly_pCons) - -lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)" -proof - - have "smult c p = [:c:] * p" by simp - also have "normalize \ = smult (normalize c) (normalize p)" - by (subst normalize_mult) (simp add: normalize_const_poly) - finally show ?thesis . -qed - - -subsubsection \Division in Field Polynomials\ - -text\ - This part connects the above result to the division of field polynomials. - Mainly imported from Isabelle 2016. -\ - -lemma pseudo_divmod_field: - assumes g: "(g::'a::field poly) \ 0" and *: "pseudo_divmod f g = (q,r)" - defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" - shows "f = g * smult (1/c) q + smult (1/c) r" -proof - - from leading_coeff_neq_0[OF g] have c0: "c \ 0" unfolding c_def by auto - from pseudo_divmod(1)[OF g *, folded c_def] - have "smult c f = g * q + r" by auto - also have "smult (1/c) ... = g * smult (1/c) q + smult (1/c) r" by (simp add: smult_add_right) - finally show ?thesis using c0 by auto -qed - -lemma divide_poly_main_field: - assumes d: "(d::'a::field poly) \ 0" - defines lc: "lc \ coeff d (degree d)" - shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1/lc)^n) q) (smult ((1/lc)^n) r) d dr n)" - unfolding lc - by(subst fst_pseudo_divmod_main_as_divide_poly_main, auto simp: d power_one_over) - -lemma divide_poly_field: - fixes f g :: "'a :: field poly" - defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" - shows "(f::'a::field poly) div g = fst (pseudo_divmod f' g)" -proof (cases "g = 0") - case True show ?thesis - unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) -next - case False - from leading_coeff_neq_0[OF False] have "degree f' = degree f" unfolding f'_def by auto - then show ?thesis - using length_coeffs_degree[of f'] length_coeffs_degree[of f] - unfolding divide_poly_def pseudo_divmod_def Let_def - divide_poly_main_field[OF False] - length_coeffs_degree[OF False] - f'_def - by force -qed - -instantiation poly :: (field) ring_div -begin - -definition modulo_poly where - mod_poly_def: "f mod g \ - if g = 0 then f - else pseudo_mod (smult ((1/coeff g (degree g)) ^ (Suc (degree f) - degree g)) f) g" - -lemma eucl_rel_poly: "eucl_rel_poly (x::'a::field poly) y (x div y, x mod y)" - unfolding eucl_rel_poly_iff -proof (intro conjI) - show "x = x div y * y + x mod y" - proof(cases "y = 0") - case True show ?thesis by(simp add: True divide_poly_def divide_poly_0 mod_poly_def) - next - case False - then have "pseudo_divmod (smult ((1 / coeff y (degree y)) ^ (Suc (degree x) - degree y)) x) y = - (x div y, x mod y)" - unfolding divide_poly_field mod_poly_def pseudo_mod_def by simp - from pseudo_divmod[OF False this] - show ?thesis using False - by (simp add: power_mult_distrib[symmetric] mult.commute) - qed - show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" - proof (cases "y = 0") - case True then show ?thesis by auto - next - case False - with pseudo_mod[OF this] show ?thesis unfolding mod_poly_def by simp - qed -qed - -lemma div_poly_eq: - "eucl_rel_poly (x::'a::field poly) y (q, r) \ x div y = q" - by(rule eucl_rel_poly_unique_div[OF eucl_rel_poly]) - -lemma mod_poly_eq: - "eucl_rel_poly (x::'a::field poly) y (q, r) \ x mod y = r" - by (rule eucl_rel_poly_unique_mod[OF eucl_rel_poly]) - -instance -proof - fix x y :: "'a poly" - from eucl_rel_poly[of x y,unfolded eucl_rel_poly_iff] - show "x div y * y + x mod y = x" by auto -next - fix x y z :: "'a poly" - assume "y \ 0" - hence "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)" - using eucl_rel_poly [of x y] - by (simp add: eucl_rel_poly_iff distrib_right) - thus "(x + z * y) div y = z + x div y" - by (rule div_poly_eq) -next - fix x y z :: "'a poly" - assume "x \ 0" - show "(x * y) div (x * z) = y div z" - proof (cases "y \ 0 \ z \ 0") - have "\x::'a poly. eucl_rel_poly x 0 (0, x)" - by (rule eucl_rel_poly_by_0) - then have [simp]: "\x::'a poly. x div 0 = 0" - by (rule div_poly_eq) - have "\x::'a poly. eucl_rel_poly 0 x (0, 0)" - by (rule eucl_rel_poly_0) - then have [simp]: "\x::'a poly. 0 div x = 0" - by (rule div_poly_eq) - case False then show ?thesis by auto - next - case True then have "y \ 0" and "z \ 0" by auto - with \x \ 0\ - have "\q r. eucl_rel_poly y z (q, r) \ eucl_rel_poly (x * y) (x * z) (q, x * r)" - by (auto simp add: eucl_rel_poly_iff algebra_simps) - (rule classical, simp add: degree_mult_eq) - moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" . - ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" . - then show ?thesis by (simp add: div_poly_eq) - qed -qed - -end - -lemma degree_mod_less: - "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" - using eucl_rel_poly [of x y] - unfolding eucl_rel_poly_iff by simp - -lemma div_poly_less: "degree (x::'a::field poly) < degree y \ x div y = 0" -proof - - assume "degree x < degree y" - hence "eucl_rel_poly x y (0, x)" - by (simp add: eucl_rel_poly_iff) - thus "x div y = 0" by (rule div_poly_eq) -qed - -lemma mod_poly_less: "degree x < degree y \ x mod y = x" -proof - - assume "degree x < degree y" - hence "eucl_rel_poly x y (0, x)" - by (simp add: eucl_rel_poly_iff) - thus "x mod y = x" by (rule mod_poly_eq) -qed - -lemma eucl_rel_poly_smult_left: - "eucl_rel_poly x y (q, r) - \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" - unfolding eucl_rel_poly_iff by (simp add: smult_add_right) - -lemma div_smult_left: "(smult (a::'a::field) x) div y = smult a (x div y)" - by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) - -lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" - by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) - -lemma poly_div_minus_left [simp]: - fixes x y :: "'a::field poly" - shows "(- x) div y = - (x div y)" - using div_smult_left [of "- 1::'a"] by simp - -lemma poly_mod_minus_left [simp]: - fixes x y :: "'a::field poly" - shows "(- x) mod y = - (x mod y)" - using mod_smult_left [of "- 1::'a"] by simp - -lemma eucl_rel_poly_add_left: - assumes "eucl_rel_poly x y (q, r)" - assumes "eucl_rel_poly x' y (q', r')" - shows "eucl_rel_poly (x + x') y (q + q', r + r')" - using assms unfolding eucl_rel_poly_iff - by (auto simp add: algebra_simps degree_add_less) - -lemma poly_div_add_left: - fixes x y z :: "'a::field poly" - shows "(x + y) div z = x div z + y div z" - using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] - by (rule div_poly_eq) - -lemma poly_mod_add_left: - fixes x y z :: "'a::field poly" - shows "(x + y) mod z = x mod z + y mod z" - using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] - by (rule mod_poly_eq) - -lemma poly_div_diff_left: - fixes x y z :: "'a::field poly" - shows "(x - y) div z = x div z - y div z" - by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) - -lemma poly_mod_diff_left: - fixes x y z :: "'a::field poly" - shows "(x - y) mod z = x mod z - y mod z" - by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) - -lemma eucl_rel_poly_smult_right: - "a \ 0 \ eucl_rel_poly x y (q, r) - \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" - unfolding eucl_rel_poly_iff by simp - -lemma div_smult_right: - "(a::'a::field) \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" - by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) - -lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" - by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) - -lemma poly_div_minus_right [simp]: - fixes x y :: "'a::field poly" - shows "x div (- y) = - (x div y)" - using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) - -lemma poly_mod_minus_right [simp]: - fixes x y :: "'a::field poly" - shows "x mod (- y) = x mod y" - using mod_smult_right [of "- 1::'a"] by simp - -lemma eucl_rel_poly_mult: - "eucl_rel_poly x y (q, r) \ eucl_rel_poly q z (q', r') - \ eucl_rel_poly x (y * z) (q', y * r' + r)" -apply (cases "z = 0", simp add: eucl_rel_poly_iff) -apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff) -apply (cases "r = 0") -apply (cases "r' = 0") -apply (simp add: eucl_rel_poly_iff) -apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq) -apply (cases "r' = 0") -apply (simp add: eucl_rel_poly_iff degree_mult_eq) -apply (simp add: eucl_rel_poly_iff field_simps) -apply (simp add: degree_mult_eq degree_add_less) -done - -lemma poly_div_mult_right: - fixes x y z :: "'a::field poly" - shows "x div (y * z) = (x div y) div z" - by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) - -lemma poly_mod_mult_right: - fixes x y z :: "'a::field poly" - shows "x mod (y * z) = y * (x div y mod z) + x mod y" - by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) - -lemma mod_pCons: - fixes a and x - assumes y: "y \ 0" - defines b: "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" - shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" -unfolding b -apply (rule mod_poly_eq) -apply (rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) -done - -definition pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" -where - "pdivmod p q = (p div q, p mod q)" - -lemma pdivmod_0: - "pdivmod 0 q = (0, 0)" - by (simp add: pdivmod_def) - -lemma pdivmod_pCons: - "pdivmod (pCons a p) q = - (if q = 0 then (0, pCons a p) else - (let (s, r) = pdivmod p q; - b = coeff (pCons a r) (degree q) / coeff q (degree q) - in (pCons b s, pCons a r - smult b q)))" - apply (simp add: pdivmod_def Let_def, safe) - apply (rule div_poly_eq) - apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl]) - apply (rule mod_poly_eq) - apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl]) - done - -lemma pdivmod_fold_coeffs: - "pdivmod p q = (if q = 0 then (0, p) - else fold_coeffs (\a (s, r). - let b = coeff (pCons a r) (degree q) / coeff q (degree q) - in (pCons b s, pCons a r - smult b q) - ) p (0, 0))" - apply (cases "q = 0") - apply (simp add: pdivmod_def) - apply (rule sym) - apply (induct p) - apply (simp_all add: pdivmod_0 pdivmod_pCons) - apply (case_tac "a = 0 \ p = 0") - apply (auto simp add: pdivmod_def) - done - -subsection \List-based versions for fast implementation\ -(* Subsection by: - Sebastiaan Joosten - René Thiemann - Akihisa Yamada - *) -fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where - "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" -| "minus_poly_rev_list xs [] = xs" -| "minus_poly_rev_list [] (y # ys) = []" - -fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list - \ nat \ 'a list \ 'a list" where - "pseudo_divmod_main_list lc q r d (Suc n) = (let - rr = map (op * lc) r; - a = hd r; - qqq = cCons a (map (op * lc) q); - rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) - in pseudo_divmod_main_list lc qqq rrr d n)" -| "pseudo_divmod_main_list lc q r d 0 = (q,r)" - -fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list - \ nat \ 'a list" where - "pseudo_mod_main_list lc r d (Suc n) = (let - rr = map (op * lc) r; - a = hd r; - rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) - in pseudo_mod_main_list lc rrr d n)" -| "pseudo_mod_main_list lc r d 0 = r" - - -fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list - \ nat \ 'a list \ 'a list" where - "divmod_poly_one_main_list q r d (Suc n) = (let - a = hd r; - qqq = cCons a q; - rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) - in divmod_poly_one_main_list qqq rr d n)" -| "divmod_poly_one_main_list q r d 0 = (q,r)" - -fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list - \ nat \ 'a list" where - "mod_poly_one_main_list r d (Suc n) = (let - a = hd r; - rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) - in mod_poly_one_main_list rr d n)" -| "mod_poly_one_main_list r d 0 = r" - -definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where - "pseudo_divmod_list p q = - (if q = [] then ([],p) else - (let rq = rev q; - (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in - (qu,rev re)))" - -definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where - "pseudo_mod_list p q = - (if q = [] then p else - (let rq = rev q; - re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in - (rev re)))" - -lemma minus_zero_does_nothing: -"(minus_poly_rev_list x (map (op * 0) y)) = (x :: 'a :: ring list)" - by(induct x y rule: minus_poly_rev_list.induct, auto) - -lemma length_minus_poly_rev_list[simp]: - "length (minus_poly_rev_list xs ys) = length xs" - by(induct xs ys rule: minus_poly_rev_list.induct, auto) - -lemma if_0_minus_poly_rev_list: - "(if a = 0 then x else minus_poly_rev_list x (map (op * a) y)) - = minus_poly_rev_list x (map (op * (a :: 'a :: ring)) y)" - by(cases "a=0",simp_all add:minus_zero_does_nothing) - -lemma Poly_append: - "Poly ((a::'a::comm_semiring_1 list) @ b) = Poly a + monom 1 (length a) * Poly b" - by (induct a,auto simp: monom_0 monom_Suc) - -lemma minus_poly_rev_list: "length p \ length (q :: 'a :: comm_ring_1 list) \ - Poly (rev (minus_poly_rev_list (rev p) (rev q))) - = Poly p - monom 1 (length p - length q) * Poly q" -proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) - case (1 x xs y ys) - have "length (rev q) \ length (rev p)" using 1 by simp - from this[folded 1(2,3)] have ys_xs:"length ys \ length xs" by simp - hence a:"Poly (rev (minus_poly_rev_list xs ys)) = - Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" - by(subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev],auto) - have "Poly p - monom 1 (length p - length q) * Poly q - = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" - by simp - also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" - unfolding 1(2,3) by simp - also have "\ = Poly (rev xs) + monom x (length xs) - - (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" using ys_xs - by (simp add:Poly_append distrib_left mult_monom smult_monom) - also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" - unfolding a diff_monom[symmetric] by(simp) - finally show ?case - unfolding 1(2,3)[symmetric] by (simp add: smult_monom Poly_append) -qed auto - -lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" - using smult_monom [of a _ n] by (metis mult_smult_left) - -lemma head_minus_poly_rev_list: - "length d \ length r \ d\[] \ - hd (minus_poly_rev_list (map (op * (last d :: 'a :: comm_ring)) r) (map (op * (hd r)) (rev d))) = 0" -proof(induct r) - case (Cons a rs) - thus ?case by(cases "rev d", simp_all add:ac_simps) -qed simp - -lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)" -proof (induct p) - case(Cons x xs) thus ?case by (cases "Poly xs = 0",auto) -qed simp - -lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" - by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) - -lemma pseudo_divmod_main_list_invar : - assumes leading_nonzero:"last d \ 0" - and lc:"last d = lc" - and dNonempty:"d \ []" - and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q',rev r')" - and "n = (1 + length r - length d)" - shows - "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = - (Poly q', Poly r')" -using assms(4-) -proof(induct "n" arbitrary: r q) -case (Suc n r q) - have ifCond: "\ Suc (length r) \ length d" using Suc.prems by simp - have rNonempty:"r \ []" - using ifCond dNonempty using Suc_leI length_greater_0_conv list.size(3) by fastforce - let ?a = "(hd (rev r))" - let ?rr = "map (op * lc) (rev r)" - let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map (op * ?a) (rev d))))" - let ?qq = "cCons ?a (map (op * lc) q)" - have n: "n = (1 + length r - length d - 1)" - using ifCond Suc(3) by simp - have rr_val:"(length ?rrr) = (length r - 1)" using ifCond by auto - hence rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" - using rNonempty ifCond unfolding One_nat_def by auto - from ifCond have id: "Suc (length r) - length d = Suc (length r - length d)" by auto - have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" - using Suc.prems ifCond by (simp add:Let_def if_0_minus_poly_rev_list id) - hence v:"pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" - using n by auto - have sucrr:"Suc (length r) - length d = Suc (length r - length d)" - using Suc_diff_le ifCond not_less_eq_eq by blast - have n_ok : "n = 1 + (length ?rrr) - length d" using Suc(3) rNonempty by simp - have cong: "\ x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ - pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp - have hd_rev:"coeff (Poly r) (length r - Suc 0) = hd (rev r)" - using last_coeff_is_hd[OF rNonempty] by simp - show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def - proof (rule cong[OF _ _ refl], goal_cases) - case 1 - show ?case unfolding monom_Suc hd_rev[symmetric] - by (simp add: smult_monom Poly_map) - next - case 2 - show ?case - proof (subst Poly_on_rev_starting_with_0, goal_cases) - show "hd (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))) = 0" - by (fold lc, subst head_minus_poly_rev_list, insert ifCond dNonempty,auto) - from ifCond have "length d \ length r" by simp - then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = - Poly (rev (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))))" - by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] - minus_poly_rev_list) - qed - qed simp -qed simp - -lemma pseudo_divmod_impl[code]: "pseudo_divmod (f::'a::comm_ring_1 poly) g = - map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" -proof (cases "g=0") -case False - hence coeffs_g_nonempty:"(coeffs g) \ []" by simp - hence lastcoeffs:"last (coeffs g) = coeff g (degree g)" - by (simp add: hd_rev last_coeffs_eq_coeff_degree not_0_coeffs_not_Nil) - obtain q r where qr: "pseudo_divmod_main_list - (last (coeffs g)) (rev []) - (rev (coeffs f)) (rev (coeffs g)) - (1 + length (coeffs f) - - length (coeffs g)) = (q,rev (rev r))" by force - then have qr': "pseudo_divmod_main_list - (hd (rev (coeffs g))) [] - (rev (coeffs f)) (rev (coeffs g)) - (1 + length (coeffs f) - - length (coeffs g)) = (q,r)" using hd_rev[OF coeffs_g_nonempty] by(auto) - from False have cg: "(coeffs g = []) = False" by auto - have last_non0:"last (coeffs g) \ 0" using False by (simp add:last_coeffs_not_0) - show ?thesis - unfolding pseudo_divmod_def pseudo_divmod_list_def Let_def qr' map_prod_def split cg if_False - pseudo_divmod_main_list_invar[OF last_non0 _ _ qr,unfolded lastcoeffs,simplified,symmetric,OF False] - poly_of_list_def - using False by (auto simp: degree_eq_length_coeffs) -next - case True - show ?thesis unfolding True unfolding pseudo_divmod_def pseudo_divmod_list_def - by auto -qed - -lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q - xs ys n) = pseudo_mod_main_list l xs ys n" - by (induct n arbitrary: l q xs ys, auto simp: Let_def) - -lemma pseudo_mod_impl[code]: "pseudo_mod f g = - poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" -proof - - have snd_case: "\ f g p. snd ((\ (x,y). (f x, g y)) p) = g (snd p)" - by auto - show ?thesis - unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def - pseudo_mod_list_def Let_def - by (simp add: snd_case pseudo_mod_main_list) -qed - - -(* *************** *) -subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ - -lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ pdivmod p q = (r, s)" - by (metis pdivmod_def eucl_rel_poly eucl_rel_poly_unique) - -lemma pdivmod_via_pseudo_divmod: "pdivmod f g = (if g = 0 then (0,f) - else let - ilc = inverse (coeff g (degree g)); - h = smult ilc g; - (q,r) = pseudo_divmod f h - in (smult ilc q, r))" (is "?l = ?r") -proof (cases "g = 0") - case False - define lc where "lc = inverse (coeff g (degree g))" - define h where "h = smult lc g" - from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" unfolding h_def lc_def by auto - hence h0: "h \ 0" by auto - obtain q r where p: "pseudo_divmod f h = (q,r)" by force - from False have id: "?r = (smult lc q, r)" - unfolding Let_def h_def[symmetric] lc_def[symmetric] p by auto - from pseudo_divmod[OF h0 p, unfolded h1] - have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto - have "eucl_rel_poly f h (q, r)" unfolding eucl_rel_poly_iff using f r h0 by auto - hence "pdivmod f h = (q,r)" by (simp add: pdivmod_pdivmodrel) - hence "pdivmod f g = (smult lc q, r)" - unfolding pdivmod_def h_def div_smult_right[OF lc] mod_smult_right[OF lc] - using lc by auto - with id show ?thesis by auto -qed (auto simp: pdivmod_def) - -lemma pdivmod_via_pseudo_divmod_list: "pdivmod f g = (let - cg = coeffs g - in if cg = [] then (0,f) - else let - cf = coeffs f; - ilc = inverse (last cg); - ch = map (op * ilc) cg; - (q,r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) - in (poly_of_list (map (op * ilc) q), poly_of_list (rev r)))" -proof - - note d = pdivmod_via_pseudo_divmod - pseudo_divmod_impl pseudo_divmod_list_def - show ?thesis - proof (cases "g = 0") - case True thus ?thesis unfolding d by auto - next - case False - define ilc where "ilc = inverse (coeff g (degree g))" - from False have ilc: "ilc \ 0" unfolding ilc_def by auto - with False have id: "(g = 0) = False" "(coeffs g = []) = False" - "last (coeffs g) = coeff g (degree g)" - "(coeffs (smult ilc g) = []) = False" - by (auto simp: last_coeffs_eq_coeff_degree) - have id2: "hd (rev (coeffs (smult ilc g))) = 1" - by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) - have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" - "rev (coeffs (smult ilc g)) = rev (map (op * ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto - obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map (op * ilc) (coeffs g))) - (1 + length (coeffs f) - length (coeffs g)) = (q,r)" by force - show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 - unfolding id3 pair map_prod_def split by (auto simp: Poly_map) - qed -qed - -lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" -proof (intro ext, goal_cases) - case (1 q r d n) - { - fix xs :: "'a list" - have "map (op * 1) xs = xs" by (induct xs, auto) - } note [simp] = this - show ?case by (induct n arbitrary: q r d, auto simp: Let_def) -qed - -fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list - \ nat \ 'a list" where - "divide_poly_main_list lc q r d (Suc n) = (let - cr = hd r - in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let - a = cr div lc; - qq = cCons a q; - rr = minus_poly_rev_list r (map (op * a) d) - in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" -| "divide_poly_main_list lc q r d 0 = q" - - -lemma divide_poly_main_list_simp[simp]: "divide_poly_main_list lc q r d (Suc n) = (let - cr = hd r; - a = cr div lc; - qq = cCons a q; - rr = minus_poly_rev_list r (map (op * a) d) - in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" - by (simp add: Let_def minus_zero_does_nothing) - -declare divide_poly_main_list.simps(1)[simp del] - -definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where - "divide_poly_list f g = - (let cg = coeffs g - in if cg = [] then g - else let cf = coeffs f; cgr = rev cg - in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" - -lemmas pdivmod_via_divmod_list[code] = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] - -lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" - by (induct n arbitrary: q r d, auto simp: Let_def) - -lemma mod_poly_code[code]: "f mod g = - (let cg = coeffs g - in if cg = [] then f - else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg; - r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) - in poly_of_list (rev r))" (is "?l = ?r") -proof - - have "?l = snd (pdivmod f g)" unfolding pdivmod_def by simp - also have "\ = ?r" unfolding pdivmod_via_divmod_list Let_def - mod_poly_one_main_list[symmetric, of _ _ _ Nil] by (auto split: prod.splits) - finally show ?thesis . -qed - -definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where - "div_field_poly_impl f g = ( - let cg = coeffs g - in if cg = [] then 0 - else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg; - q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) - in poly_of_list ((map (op * ilc) q)))" - -text \We do not declare the following lemma as code equation, since then polynomial division - on non-fields will no longer be executable. However, a code-unfold is possible, since - \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ -lemma div_field_poly_impl[code_unfold]: "op div = div_field_poly_impl" -proof (intro ext) - fix f g :: "'a poly" - have "f div g = fst (pdivmod f g)" unfolding pdivmod_def by simp - also have "\ = div_field_poly_impl f g" unfolding - div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) - finally show "f div g = div_field_poly_impl f g" . -qed - - -lemma divide_poly_main_list: - assumes lc0: "lc \ 0" - and lc:"last d = lc" - and d:"d \ []" - and "n = (1 + length r - length d)" - shows - "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = - divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" -using assms(4-) -proof(induct "n" arbitrary: r q) -case (Suc n r q) - have ifCond: "\ Suc (length r) \ length d" using Suc.prems by simp - have r: "r \ []" - using ifCond d using Suc_leI length_greater_0_conv list.size(3) by fastforce - then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases, auto) - from d lc obtain dd where d: "d = dd @ [lc]" - by (cases d rule: rev_cases, auto) - from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) - from ifCond have len: "length dd \ length rr" by (simp add: r d) - show ?case - proof (cases "lcr div lc * lc = lcr") - case False - thus ?thesis unfolding Suc(2)[symmetric] using r d - by (auto simp add: Let_def nth_default_append) - next - case True - hence id: - "?thesis = (Poly (divide_poly_main_list lc (cCons (lcr div lc) q) - (rev (rev (minus_poly_rev_list (rev rr) (rev (map (op * (lcr div lc)) dd))))) (rev d) n) = - divide_poly_main lc - (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) - (Poly r - monom (lcr div lc) n * Poly d) - (Poly d) (length rr - 1) n)" - using r d - by (cases r rule: rev_cases; cases "d" rule: rev_cases; - auto simp add: Let_def rev_map nth_default_append) - have cong: "\ x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ - divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp - show ?thesis unfolding id - proof (subst Suc(1), simp add: n, - subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) - case 2 - have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" - by (simp add: mult_monom len True) - thus ?case unfolding r d Poly_append n ring_distribs - by (auto simp: Poly_map smult_monom smult_monom_mult) - qed (auto simp: len monom_Suc smult_monom) - qed -qed simp - - -lemma divide_poly_list[code]: "f div g = divide_poly_list f g" -proof - - note d = divide_poly_def divide_poly_list_def - show ?thesis - proof (cases "g = 0") - case True - show ?thesis unfolding d True by auto - next - case False - then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases, auto) - with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto - from cg False have lcg: "coeff g (degree g) = lcg" - using last_coeffs_eq_coeff_degree last_snoc by force - with False have lcg0: "lcg \ 0" by auto - from cg have ltp: "Poly (cg @ [lcg]) = g" - using Poly_coeffs [of g] by auto - show ?thesis unfolding d cg Let_def id if_False poly_of_list_def - by (subst divide_poly_main_list, insert False cg lcg0, auto simp: lcg ltp, - simp add: degree_eq_length_coeffs) - qed -qed - -subsection \Order of polynomial roots\ + +subsubsection \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where @@ -2986,6 +1787,124 @@ lemma order_0I: "poly p a \ 0 \ order a p = 0" by (subst (asm) order_root) auto +lemma order_unique_lemma: + fixes p :: "'a::idom poly" + assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" + shows "n = order a p" +unfolding Polynomial.order_def +apply (rule Least_equality [symmetric]) +apply (fact assms) +apply (rule classical) +apply (erule notE) +unfolding not_less_eq_eq +using assms(1) apply (rule power_le_dvd) +apply assumption + done + +lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" +proof - + define i where "i = order a p" + define j where "j = order a q" + define t where "t = [:-a, 1:]" + have t_dvd_iff: "\u. t dvd u \ poly u a = 0" + unfolding t_def by (simp add: dvd_iff_poly_eq_0) + assume "p * q \ 0" + then show "order a (p * q) = i + j" + apply clarsimp + apply (drule order [where a=a and p=p, folded i_def t_def]) + apply (drule order [where a=a and p=q, folded j_def t_def]) + apply clarify + apply (erule dvdE)+ + apply (rule order_unique_lemma [symmetric], fold t_def) + apply (simp_all add: power_add t_dvd_iff) + done +qed + +lemma order_smult: + assumes "c \ 0" + shows "order x (smult c p) = order x p" +proof (cases "p = 0") + case False + have "smult c p = [:c:] * p" by simp + also from assms False have "order x \ = order x [:c:] + order x p" + by (subst order_mult) simp_all + also from assms have "order x [:c:] = 0" by (intro order_0I) auto + finally show ?thesis by simp +qed simp + +(* Next two lemmas contributed by Wenda Li *) +lemma order_1_eq_0 [simp]:"order x 1 = 0" + by (metis order_root poly_1 zero_neq_one) + +lemma order_power_n_n: "order a ([:-a,1:]^n)=n" +proof (induct n) (*might be proved more concisely using nat_less_induct*) + case 0 + thus ?case by (metis order_root poly_1 power_0 zero_neq_one) +next + case (Suc n) + have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" + by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral + one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) + moreover have "order a [:-a,1:]=1" unfolding order_def + proof (rule Least_equality,rule ccontr) + assume "\ \ [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" + hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp + hence "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:] )" + by (rule dvd_imp_degree_le,auto) + thus False by auto + next + fix y assume asm:"\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" + show "1 \ y" + proof (rule ccontr) + assume "\ 1 \ y" + hence "y=0" by auto + hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto + thus False using asm by auto + qed + qed + ultimately show ?case using Suc by auto +qed + +lemma order_0_monom [simp]: + assumes "c \ 0" + shows "order 0 (monom c n) = n" + using assms order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) + +lemma dvd_imp_order_le: + "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" + by (auto simp: order_mult elim: dvdE) + +text\Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\ + +lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" +apply (cases "p = 0", auto) +apply (drule order_2 [where a=a and p=p]) +apply (metis not_less_eq_eq power_le_dvd) +apply (erule power_le_dvd [OF order_1]) +done + +lemma order_decomp: + assumes "p \ 0" + shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" +proof - + from assms have A: "[:- a, 1:] ^ order a p dvd p" + and B: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) + from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. + with B have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" + by simp + then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" + by simp + then have D: "\ [:- a, 1:] dvd q" + using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] + by auto + from C D show ?thesis by blast +qed + +lemma monom_1_dvd_iff: + assumes "p \ 0" + shows "monom 1 n dvd p \ n \ order 0 p" + using assms order_divides[of 0 n p] by (simp add: monom_altdef) + subsection \Additional induction rules on polynomials\ @@ -3055,7 +1974,7 @@ finally show ?thesis . qed - + subsection \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) @@ -3218,7 +2137,7 @@ shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" proof (induct p) case 0 - thus ?case unfolding lead_coeff_def by auto + thus ?case by auto next case (pCons a p) have "degree ( q * pcompose p q) = 0 \ ?case" @@ -3230,9 +2149,13 @@ qed moreover have "degree ( q * pcompose p q) > 0 \ ?case" proof - - assume "degree ( q * pcompose p q) > 0" - hence "lead_coeff (pcompose (pCons a p) q) =lead_coeff ( q * pcompose p q)" - by (auto simp add:pcompose_pCons lead_coeff_add_le) + assume "degree (q * pcompose p q) > 0" + then have "degree [:a:] < degree (q * pcompose p q)" + by simp + then have "lead_coeff ([:a:] + q * p \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" + by (rule lead_coeff_add_le) + then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" + by (simp add: pcompose_pCons) also have "... = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "... = lead_coeff p * lead_coeff q ^ (degree p + 1)" @@ -3254,7 +2177,6 @@ lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" by (auto simp add: nth_default_def add_ac) - lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" proof - from MOST_coeff_eq_0[of p] obtain m where "\k>m. coeff p k = 0" by (auto simp: MOST_nat) @@ -3359,7 +2281,7 @@ by (induction rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" - by (simp add: coeff_reflect_poly lead_coeff_def) + by (simp add: coeff_reflect_poly) lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" by (simp add: poly_0_coeff_0) @@ -3442,7 +2364,7 @@ reflect_poly_power reflect_poly_prod reflect_poly_prod_list -subsection \Derivatives of univariate polynomials\ +subsection \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ 'a poly" where @@ -3735,6 +2657,136 @@ qed qed +lemma lemma_order_pderiv1: + "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" +apply (simp only: pderiv_mult pderiv_power_Suc) +apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons) +done + +lemma lemma_order_pderiv: + fixes p :: "'a :: field_char_0 poly" + assumes n: "0 < n" + and pd: "pderiv p \ 0" + and pe: "p = [:- a, 1:] ^ n * q" + and nd: "~ [:- a, 1:] dvd q" + shows "n = Suc (order a (pderiv p))" +using n +proof - + have "pderiv ([:- a, 1:] ^ n * q) \ 0" + using assms by auto + obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" + using assms by (cases n) auto + have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" + by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) + have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" + proof (rule order_unique_lemma) + show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" + apply (subst lemma_order_pderiv1) + apply (rule dvd_add) + apply (metis dvdI dvd_mult2 power_Suc2) + apply (metis dvd_smult dvd_triv_right) + done + next + show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" + apply (subst lemma_order_pderiv1) + by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) + qed + then show ?thesis + by (metis \n = Suc n'\ pe) +qed + +lemma order_pderiv: + "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ \ + (order a p = Suc (order a (pderiv p)))" +apply (case_tac "p = 0", simp) +apply (drule_tac a = a and p = p in order_decomp) +using neq0_conv +apply (blast intro: lemma_order_pderiv) +done + +lemma poly_squarefree_decomp_order: + assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" + and p: "p = q * d" + and p': "pderiv p = e * d" + and d: "d = r * p + s * pderiv p" + shows "order a q = (if order a p = 0 then 0 else 1)" +proof (rule classical) + assume 1: "order a q \ (if order a p = 0 then 0 else 1)" + from \pderiv p \ 0\ have "p \ 0" by auto + with p have "order a p = order a q + order a d" + by (simp add: order_mult) + with 1 have "order a p \ 0" by (auto split: if_splits) + have "order a (pderiv p) = order a e + order a d" + using \pderiv p \ 0\ \pderiv p = e * d\ by (simp add: order_mult) + have "order a p = Suc (order a (pderiv p))" + using \pderiv p \ 0\ \order a p \ 0\ by (rule order_pderiv) + have "d \ 0" using \p \ 0\ \p = q * d\ by simp + have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" + apply (simp add: d) + apply (rule dvd_add) + apply (rule dvd_mult) + apply (simp add: order_divides \p \ 0\ + \order a p = Suc (order a (pderiv p))\) + apply (rule dvd_mult) + apply (simp add: order_divides) + done + then have "order a (pderiv p) \ order a d" + using \d \ 0\ by (simp add: order_divides) + show ?thesis + using \order a p = order a q + order a d\ + using \order a (pderiv p) = order a e + order a d\ + using \order a p = Suc (order a (pderiv p))\ + using \order a (pderiv p) \ order a d\ + by auto +qed + +lemma poly_squarefree_decomp_order2: + "\pderiv p \ (0 :: 'a :: field_char_0 poly); + p = q * d; + pderiv p = e * d; + d = r * p + s * pderiv p + \ \ \a. order a q = (if order a p = 0 then 0 else 1)" +by (blast intro: poly_squarefree_decomp_order) + +lemma order_pderiv2: + "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ + \ (order a (pderiv p) = n) = (order a p = Suc n)" +by (auto dest: order_pderiv) + +definition rsquarefree :: "'a::idom poly \ bool" + where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" + +lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h :: 'a :: {semidom,semiring_char_0}:]" + by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) + +lemma rsquarefree_roots: + fixes p :: "'a :: field_char_0 poly" + shows "rsquarefree p = (\a. \(poly p a = 0 \ poly (pderiv p) a = 0))" +apply (simp add: rsquarefree_def) +apply (case_tac "p = 0", simp, simp) +apply (case_tac "pderiv p = 0") +apply simp +apply (drule pderiv_iszero, clarsimp) +apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) +apply (force simp add: order_root order_pderiv2) + done + +lemma poly_squarefree_decomp: + assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" + and "p = q * d" + and "pderiv p = e * d" + and "d = r * p + s * pderiv p" + shows "rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" +proof - + from \pderiv p \ 0\ have "p \ 0" by auto + with \p = q * d\ have "q \ 0" by simp + have "\a. order a q = (if order a p = 0 then 0 else 1)" + using assms by (rule poly_squarefree_decomp_order2) + with \p \ 0\ \q \ 0\ show ?thesis + by (simp add: rsquarefree_def order_root) +qed + subsection \Algebraic numbers\ @@ -3746,7 +2798,6 @@ The equivalence is obvious since any rational polynomial can be multiplied with the LCM of its coefficients, yielding an integer polynomial with the same roots. \ -subsection \Algebraic numbers\ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" @@ -3761,25 +2812,6 @@ obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast -lemma quotient_of_denom_pos': "snd (quotient_of x) > 0" - using quotient_of_denom_pos[OF surjective_pairing] . - -lemma of_int_div_in_Ints: - "b dvd a \ of_int a div of_int b \ (\ :: 'a :: ring_div set)" -proof (cases "of_int b = (0 :: 'a)") - assume "b dvd a" "of_int b \ (0::'a)" - then obtain c where "a = b * c" by (elim dvdE) - with \of_int b \ (0::'a)\ show ?thesis by simp -qed auto - -lemma of_int_divide_in_Ints: - "b dvd a \ of_int a / of_int b \ (\ :: 'a :: field set)" -proof (cases "of_int b = (0 :: 'a)") - assume "b dvd a" "of_int b \ (0::'a)" - then obtain c where "a = b * c" by (elim dvdE) - with \of_int b \ (0::'a)\ show ?thesis by simp -qed auto - lemma algebraic_altdef: fixes p :: "'a :: field_char_0 poly" shows "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" @@ -3834,285 +2866,1426 @@ qed -text\Lemmas for Derivatives\ - -lemma order_unique_lemma: - fixes p :: "'a::idom poly" - assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" - shows "n = order a p" -unfolding Polynomial.order_def -apply (rule Least_equality [symmetric]) -apply (fact assms) -apply (rule classical) -apply (erule notE) -unfolding not_less_eq_eq -using assms(1) apply (rule power_le_dvd) -apply assumption -done - -lemma lemma_order_pderiv1: - "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + - smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" -apply (simp only: pderiv_mult pderiv_power_Suc) -apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons) -done - -lemma lemma_order_pderiv: - fixes p :: "'a :: field_char_0 poly" - assumes n: "0 < n" - and pd: "pderiv p \ 0" - and pe: "p = [:- a, 1:] ^ n * q" - and nd: "~ [:- a, 1:] dvd q" - shows "n = Suc (order a (pderiv p))" -using n +subsection \Content and primitive part of a polynomial\ + +definition content :: "('a :: semiring_Gcd poly) \ 'a" where + "content p = Gcd (set (coeffs p))" + +lemma content_0 [simp]: "content 0 = 0" + by (simp add: content_def) + +lemma content_1 [simp]: "content 1 = 1" + by (simp add: content_def) + +lemma content_const [simp]: "content [:c:] = normalize c" + by (simp add: content_def cCons_def) + +lemma const_poly_dvd_iff_dvd_content: + fixes c :: "'a :: semiring_Gcd" + shows "[:c:] dvd p \ c dvd content p" +proof (cases "p = 0") + case [simp]: False + have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) + also have "\ \ (\a\set (coeffs p). c dvd a)" + proof safe + fix n :: nat assume "\a\set (coeffs p). c dvd a" + thus "c dvd coeff p n" + by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) + qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) + also have "\ \ c dvd content p" + by (simp add: content_def dvd_Gcd_iff mult.commute [of "unit_factor x" for x] + dvd_mult_unit_iff) + finally show ?thesis . +qed simp_all + +lemma content_dvd [simp]: "[:content p:] dvd p" + by (subst const_poly_dvd_iff_dvd_content) simp_all + +lemma content_dvd_coeff [simp]: "content p dvd coeff p n" + by (cases "n \ degree p") + (auto simp: content_def coeffs_def not_le coeff_eq_0 simp del: upt_Suc intro: Gcd_dvd) + +lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" + by (simp add: content_def Gcd_dvd) + +lemma normalize_content [simp]: "normalize (content p) = content p" + by (simp add: content_def) + +lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" +proof + assume "is_unit (content p)" + hence "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) + thus "content p = 1" by simp +qed auto + +lemma content_smult [simp]: "content (smult c p) = normalize c * content p" + by (simp add: content_def coeffs_smult Gcd_mult) + +lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" + by (auto simp: content_def simp: poly_eq_iff coeffs_def) + +definition primitive_part :: "'a :: {semiring_Gcd,idom_divide} poly \ 'a poly" where + "primitive_part p = (if p = 0 then 0 else map_poly (\x. x div content p) p)" + +lemma primitive_part_0 [simp]: "primitive_part 0 = 0" + by (simp add: primitive_part_def) + +lemma content_times_primitive_part [simp]: + fixes p :: "'a :: {idom_divide, semiring_Gcd} poly" + shows "smult (content p) (primitive_part p) = p" +proof (cases "p = 0") + case False + thus ?thesis + unfolding primitive_part_def + by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs + intro: map_poly_idI) +qed simp_all + +lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" +proof (cases "p = 0") + case False + hence "primitive_part p = map_poly (\x. x div content p) p" + by (simp add: primitive_part_def) + also from False have "\ = 0 \ p = 0" + by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) + finally show ?thesis using False by simp +qed simp + +lemma content_primitive_part [simp]: + assumes "p \ 0" + shows "content (primitive_part p) = 1" proof - - have "pderiv ([:- a, 1:] ^ n * q) \ 0" - using assms by auto - obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" - using assms by (cases n) auto - have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" - by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) - have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" - proof (rule order_unique_lemma) - show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" - apply (subst lemma_order_pderiv1) - apply (rule dvd_add) - apply (metis dvdI dvd_mult2 power_Suc2) - apply (metis dvd_smult dvd_triv_right) - done + have "p = smult (content p) (primitive_part p)" by simp + also have "content \ = content p * content (primitive_part p)" + by (simp del: content_times_primitive_part) + finally show ?thesis using assms by simp +qed + +lemma content_decompose: + fixes p :: "'a :: semiring_Gcd poly" + obtains p' where "p = smult (content p) p'" "content p' = 1" +proof (cases "p = 0") + case True + thus ?thesis by (intro that[of 1]) simp_all +next + case False + from content_dvd[of p] obtain r where r: "p = [:content p:] * r" by (erule dvdE) + have "content p * 1 = content p * content r" by (subst r) simp + with False have "content r = 1" by (subst (asm) mult_left_cancel) simp_all + with r show ?thesis by (intro that[of r]) simp_all +qed + +lemma content_dvd_contentI [intro]: + "p dvd q \ content p dvd content q" + using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast + +lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" + by (simp add: primitive_part_def map_poly_pCons) + +lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" + by (auto simp: primitive_part_def) + +lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" +proof (cases "p = 0") + case False + have "p = smult (content p) (primitive_part p)" by simp + also from False have "degree \ = degree (primitive_part p)" + by (subst degree_smult_eq) simp_all + finally show ?thesis .. +qed simp_all + + +subsection \Division of polynomials\ + +subsubsection \Division in general\ + +instantiation poly :: (idom_divide) idom_divide +begin + +fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly + \ nat \ nat \ 'a poly" where + "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in + if False \ a * lc = cr then (* False \ is only because of problem in function-package *) + divide_poly_main + lc + (q + mon) + (r - mon * d) + d (dr - 1) n else 0)" +| "divide_poly_main lc q r d dr 0 = q" + +definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where + "divide_poly f g = (if g = 0 then 0 else + divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" + +lemma divide_poly_main: + assumes d: "d \ 0" "lc = coeff d (degree d)" + and *: "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" + "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" + shows "q' = q + r" + using * +proof (induct n arbitrary: q r dr) + case (Suc n q r dr) + let ?rr = "d * r" + let ?a = "coeff ?rr dr" + let ?qq = "?a div lc" + define b where [simp]: "b = monom ?qq n" + let ?rrr = "d * (r - b)" + let ?qqq = "q + b" + note res = Suc(3) + have dr: "dr = n + degree d" using Suc(4) by auto + have lc: "lc \ 0" using d by auto + have "coeff (b * d) dr = coeff b n * coeff d (degree d)" + proof (cases "?qq = 0") + case False + hence n: "n = degree b" by (simp add: degree_monom_eq) + show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) + qed simp + also have "\ = lc * coeff b n" unfolding d by simp + finally have c2: "coeff (b * d) dr = lc * coeff b n" . + have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) + have c1: "coeff (d * r) dr = lc * coeff r n" + proof (cases "degree r = n") + case True + thus ?thesis using Suc(2) unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) next - show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" - apply (subst lemma_order_pderiv1) - by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) + case False + have "degree r \ n" using dr Suc(2) by auto + (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) + with False have r_n: "degree r < n" by auto + hence right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) + have "coeff (d * r) dr = coeff (d * r) (degree d + n)" unfolding dr by (simp add: ac_simps) + also have "\ = 0" using r_n + by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 + coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) + finally show ?thesis unfolding right . + qed + have c0: "coeff ?rrr dr = 0" + and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 + unfolding b_def coeff_monom coeff_smult c1 using lc by auto + from res[unfolded divide_poly_main.simps[of lc q] Let_def] id + have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" + by (simp del: divide_poly_main.simps add: field_simps) + note IH = Suc(1)[OF _ res] + have dr: "dr = n + degree d" using Suc(4) by auto + have deg_rr: "degree ?rr \ dr" using Suc(2) by auto + have deg_bd: "degree (b * d) \ dr" + unfolding dr b_def by (rule order.trans[OF degree_mult_le], auto simp: degree_monom_le) + have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) + with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" + by (rule coeff_0_degree_minus_1) + have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" + proof (cases dr) + case 0 + with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto + with deg_rrr have "degree ?rrr = 0" by simp + from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis + show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp + qed (insert Suc(4), auto) + note IH = IH[OF deg_rrr this] + show ?case using IH by simp +next + case (0 q r dr) + show ?case + proof (cases "r = 0") + case True + thus ?thesis using 0 by auto + next + case False + have "degree (d * r) = degree d + degree r" using d False + by (subst degree_mult_eq, auto) + thus ?thesis using 0 d by auto + qed +qed + +lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" +proof (induct n arbitrary: r d dr) + case (Suc n r d dr) + show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def + by (simp add: Suc del: divide_poly_main.simps) +qed simp + +lemma divide_poly: + assumes g: "g \ 0" + shows "(f * g) div g = (f :: 'a poly)" +proof - + have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) + = (f * g) div g" unfolding divide_poly_def Let_def by (simp add: ac_simps) + note main = divide_poly_main[OF g refl le_refl this] + { + fix f :: "'a poly" + assume "f \ 0" + hence "length (coeffs f) = Suc (degree f)" unfolding degree_eq_length_coeffs by auto + } note len = this + have "(f * g) div g = 0 + f" + proof (rule main, goal_cases) + case 1 + show ?case + proof (cases "f = 0") + case True + with g show ?thesis by (auto simp: degree_eq_length_coeffs) + next + case False + with g have fg: "g * f \ 0" by auto + show ?thesis unfolding len[OF fg] len[OF g] by auto + qed qed - then show ?thesis - by (metis \n = Suc n'\ pe) + thus ?thesis by simp +qed + +lemma divide_poly_0: "f div 0 = (0 :: 'a poly)" + by (simp add: divide_poly_def Let_def divide_poly_main_0) + +instance + by standard (auto simp: divide_poly divide_poly_0) + +end + +instance poly :: (idom_divide) algebraic_semidom .. + +lemma div_const_poly_conv_map_poly: + assumes "[:c:] dvd p" + shows "p div [:c:] = map_poly (\x. x div c) p" +proof (cases "c = 0") + case False + from assms obtain q where p: "p = [:c:] * q" by (erule dvdE) + moreover { + have "smult c q = [:c:] * q" by simp + also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto) + finally have "smult c q div [:c:] = q" . + } + ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) +qed (auto intro!: poly_eqI simp: coeff_map_poly) + +lemma is_unit_monom_0: + fixes a :: "'a::field" + assumes "a \ 0" + shows "is_unit (monom a 0)" +proof + from assms show "1 = monom a 0 * monom (inverse a) 0" + by (simp add: mult_monom) qed -lemma order_decomp: - assumes "p \ 0" - shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" +lemma is_unit_triv: + fixes a :: "'a::field" + assumes "a \ 0" + shows "is_unit [:a:]" + using assms by (simp add: is_unit_monom_0 monom_0 [symmetric]) + +lemma is_unit_iff_degree: + assumes "p \ (0 :: _ :: field poly)" + shows "is_unit p \ degree p = 0" (is "?P \ ?Q") +proof + assume ?Q + then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) + with assms show ?P by (simp add: is_unit_triv) +next + assume ?P + then obtain q where "q \ 0" "p * q = 1" .. + then have "degree (p * q) = degree 1" + by simp + with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" + by (simp add: degree_mult_eq) + then show ?Q by simp +qed + +lemma is_unit_pCons_iff: + "is_unit (pCons (a::_::field) p) \ p = 0 \ a \ 0" + by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree) + +lemma is_unit_monom_trival: + fixes p :: "'a::field poly" + assumes "is_unit p" + shows "monom (coeff p (degree p)) 0 = p" + using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) + +lemma is_unit_const_poly_iff: + "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \ c dvd 1" + by (auto simp: one_poly_def) + +lemma is_unit_polyE: + fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" + assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" proof - - from assms have A: "[:- a, 1:] ^ order a p dvd p" - and B: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) - from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. - with B have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" + from assms obtain q where "1 = p * q" + by (rule dvdE) + then have "p \ 0" and "q \ 0" + by auto + from \1 = p * q\ have "degree 1 = degree (p * q)" by simp - then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" - by simp - then have D: "\ [:- a, 1:] dvd q" - using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] - by auto - from C D show ?thesis by blast + also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" + by (simp add: degree_mult_eq) + finally have "degree p = 0" by simp + with degree_eq_zeroE obtain c where c: "p = [:c:]" . + moreover with \p dvd 1\ have "c dvd 1" + by (simp add: is_unit_const_poly_iff) + ultimately show thesis + by (rule that) +qed + +lemma is_unit_polyE': + assumes "is_unit (p::_::field poly)" + obtains a where "p = monom a 0" and "a \ 0" +proof - + obtain a q where "p = pCons a q" by (cases p) + with assms have "p = [:a:]" and "a \ 0" + by (simp_all add: is_unit_pCons_iff) + with that show thesis by (simp add: monom_0) qed -lemma order_pderiv: - "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ \ - (order a p = Suc (order a (pderiv p)))" -apply (case_tac "p = 0", simp) -apply (drule_tac a = a and p = p in order_decomp) -using neq0_conv -apply (blast intro: lemma_order_pderiv) -done - -lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" +lemma is_unit_poly_iff: + fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" + shows "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" + by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) + + +subsubsection \Pseudo-Division\ + +text\This part is by René Thiemann and Akihisa Yamada.\ + +fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly + \ nat \ nat \ 'a poly \ 'a poly" where + "pseudo_divmod_main lc q r d dr (Suc n) = (let + rr = smult lc r; + qq = coeff r dr; + rrr = rr - monom qq n * d; + qqq = smult lc q + monom qq n + in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" +| "pseudo_divmod_main lc q r d dr 0 = (q,r)" + +definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where + "pseudo_divmod p q \ if q = 0 then (0,p) else + pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" + +lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" + and *: "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" + "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" + shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" + using * +proof (induct n arbitrary: q r dr) + case (Suc n q r dr) + let ?rr = "smult lc r" + let ?qq = "coeff r dr" + define b where [simp]: "b = monom ?qq n" + let ?rrr = "?rr - b * d" + let ?qqq = "smult lc q + b" + note res = Suc(3) + from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] + have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" + by (simp del: pseudo_divmod_main.simps) + have dr: "dr = n + degree d" using Suc(4) by auto + have "coeff (b * d) dr = coeff b n * coeff d (degree d)" + proof (cases "?qq = 0") + case False + hence n: "n = degree b" by (simp add: degree_monom_eq) + show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) + qed auto + also have "\ = lc * coeff b n" unfolding d by simp + finally have "coeff (b * d) dr = lc * coeff b n" . + moreover have "coeff ?rr dr = lc * coeff r dr" by simp + ultimately have c0: "coeff ?rrr dr = 0" by auto + have dr: "dr = n + degree d" using Suc(4) by auto + have deg_rr: "degree ?rr \ dr" using Suc(2) + using degree_smult_le dual_order.trans by blast + have deg_bd: "degree (b * d) \ dr" + unfolding dr + by(rule order.trans[OF degree_mult_le], auto simp: degree_monom_le) + have "degree ?rrr \ dr" + using degree_diff_le[OF deg_rr deg_bd] by auto + with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) + have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" + proof (cases dr) + case 0 + with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto + with deg_rrr have "degree ?rrr = 0" by simp + hence "\ a. ?rrr = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) + from this obtain a where rrr: "?rrr = [:a:]" by auto + show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp + qed (insert Suc(4), auto) + note IH = Suc(1)[OF deg_rrr res this] + show ?case + proof (intro conjI) + show "r' = 0 \ degree r' < degree d" using IH by blast + show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" + unfolding IH[THEN conjunct2,symmetric] + by (simp add: field_simps smult_add_right) + qed +qed auto + +lemma pseudo_divmod: + assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" + shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) + and "r = 0 \ degree r < degree g" (is ?B) proof - - define i where "i = order a p" - define j where "j = order a q" - define t where "t = [:-a, 1:]" - have t_dvd_iff: "\u. t dvd u \ poly u a = 0" - unfolding t_def by (simp add: dvd_iff_poly_eq_0) - assume "p * q \ 0" - then show "order a (p * q) = i + j" - apply clarsimp - apply (drule order [where a=a and p=p, folded i_def t_def]) - apply (drule order [where a=a and p=q, folded j_def t_def]) - apply clarify - apply (erule dvdE)+ - apply (rule order_unique_lemma [symmetric], fold t_def) - apply (simp_all add: power_add t_dvd_iff) - done + from *[unfolded pseudo_divmod_def Let_def] + have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) + note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] + have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ + degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" using g + by (cases "f = 0"; cases "coeffs g", auto simp: degree_eq_length_coeffs) + note main = main[OF this] + from main show "r = 0 \ degree r < degree g" by auto + show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" + by (subst main[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, + insert g, cases "f = 0"; cases "coeffs g", auto) +qed + +definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" + +lemma snd_pseudo_divmod_main: + "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" +by (induct n arbitrary: q q' lc r d dr; simp add: Let_def) + +definition pseudo_mod + :: "'a :: {comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where + "pseudo_mod f g = snd (pseudo_divmod f g)" + +lemma pseudo_mod: + fixes f g + defines "r \ pseudo_mod f g" + assumes g: "g \ 0" + shows "\ a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" +proof - + let ?cg = "coeff g (degree g)" + let ?cge = "?cg ^ (Suc (degree f) - degree g)" + define a where "a = ?cge" + obtain q where pdm: "pseudo_divmod f g = (q,r)" using r_def[unfolded pseudo_mod_def] + by (cases "pseudo_divmod f g", auto) + from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" + unfolding a_def by auto + show "r = 0 \ degree r < degree g" by fact + from g have "a \ 0" unfolding a_def by auto + thus "\ a q. a \ 0 \ smult a f = g * q + r" using id by auto +qed + +lemma fst_pseudo_divmod_main_as_divide_poly_main: + assumes d: "d \ 0" + defines lc: "lc \ coeff d (degree d)" + shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" +proof(induct n arbitrary: q r dr) + case 0 then show ?case by simp +next + case (Suc n) + note lc0 = leading_coeff_neq_0[OF d, folded lc] + then have "pseudo_divmod_main lc q r d dr (Suc n) = + pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) + (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" + by (simp add: Let_def ac_simps) + also have "fst ... = divide_poly_main lc + (smult (lc^n) (smult lc q + monom (coeff r dr) n)) + (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) + d (dr - 1) n" + unfolding Suc[unfolded divide_poly_main.simps Let_def].. + also have "... = divide_poly_main lc (smult (lc ^ Suc n) q) + (smult (lc ^ Suc n) r) d dr (Suc n)" + unfolding smult_monom smult_distribs mult_smult_left[symmetric] + using lc0 by (simp add: Let_def ac_simps) + finally show ?case. qed -lemma order_smult: - assumes "c \ 0" - shows "order x (smult c p) = order x p" + +subsubsection \Division in polynomials over fields\ + +lemma pseudo_divmod_field: + assumes g: "(g::'a::field poly) \ 0" and *: "pseudo_divmod f g = (q,r)" + defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" + shows "f = g * smult (1/c) q + smult (1/c) r" +proof - + from leading_coeff_neq_0[OF g] have c0: "c \ 0" unfolding c_def by auto + from pseudo_divmod(1)[OF g *, folded c_def] + have "smult c f = g * q + r" by auto + also have "smult (1/c) ... = g * smult (1/c) q + smult (1/c) r" by (simp add: smult_add_right) + finally show ?thesis using c0 by auto +qed + +lemma divide_poly_main_field: + assumes d: "(d::'a::field poly) \ 0" + defines lc: "lc \ coeff d (degree d)" + shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1/lc)^n) q) (smult ((1/lc)^n) r) d dr n)" + unfolding lc + by(subst fst_pseudo_divmod_main_as_divide_poly_main, auto simp: d power_one_over) + +lemma divide_poly_field: + fixes f g :: "'a :: field poly" + defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" + shows "(f::'a::field poly) div g = fst (pseudo_divmod f' g)" +proof (cases "g = 0") + case True show ?thesis + unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) +next + case False + from leading_coeff_neq_0[OF False] have "degree f' = degree f" unfolding f'_def by auto + then show ?thesis + using length_coeffs_degree[of f'] length_coeffs_degree[of f] + unfolding divide_poly_def pseudo_divmod_def Let_def + divide_poly_main_field[OF False] + length_coeffs_degree[OF False] + f'_def + by force +qed + +instantiation poly :: ("{normalization_semidom, idom_divide}") normalization_semidom +begin + +definition unit_factor_poly :: "'a poly \ 'a poly" + where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0" + +definition normalize_poly :: "'a poly \ 'a poly" + where "normalize_poly p = map_poly (\x. x div unit_factor (lead_coeff p)) p" + +instance proof + fix p :: "'a poly" + show "unit_factor p * normalize p = p" + by (cases "p = 0") + (simp_all add: unit_factor_poly_def normalize_poly_def monom_0 + smult_conv_map_poly map_poly_map_poly o_def) +next + fix p :: "'a poly" + assume "is_unit p" + then obtain c where p: "p = [:c:]" "is_unit c" + by (auto simp: is_unit_poly_iff) + thus "normalize p = 1" + by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def) +next + fix p :: "'a poly" assume "p \ 0" + thus "is_unit (unit_factor p)" + by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff) +qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) + +end + +lemma normalize_poly_eq_div: + "normalize p = p div [:unit_factor (lead_coeff p):]" proof (cases "p = 0") case False + thus ?thesis + by (subst div_const_poly_conv_map_poly) + (auto simp: normalize_poly_def const_poly_dvd_iff) +qed (auto simp: normalize_poly_def) + +lemma unit_factor_pCons: + "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)" + by (simp add: unit_factor_poly_def) + +lemma normalize_monom [simp]: + "normalize (monom a n) = monom (normalize a) n" + by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_def degree_monom_eq) + +lemma unit_factor_monom [simp]: + "unit_factor (monom a n) = monom (unit_factor a) 0" + by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) + +lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" + by (simp add: normalize_poly_def map_poly_pCons) + +lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)" +proof - have "smult c p = [:c:] * p" by simp - also from assms False have "order x \ = order x [:c:] + order x p" - by (subst order_mult) simp_all - also from assms have "order x [:c:] = 0" by (intro order_0I) auto - finally show ?thesis by simp -qed simp - -(* Next two lemmas contributed by Wenda Li *) -lemma order_1_eq_0 [simp]:"order x 1 = 0" - by (metis order_root poly_1 zero_neq_one) - -lemma order_power_n_n: "order a ([:-a,1:]^n)=n" -proof (induct n) (*might be proved more concisely using nat_less_induct*) - case 0 - thus ?case by (metis order_root poly_1 power_0 zero_neq_one) -next - case (Suc n) - have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" - by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral - one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) - moreover have "order a [:-a,1:]=1" unfolding order_def - proof (rule Least_equality,rule ccontr) - assume "\ \ [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" - hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp - hence "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:] )" - by (rule dvd_imp_degree_le,auto) - thus False by auto - next - fix y assume asm:"\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" - show "1 \ y" - proof (rule ccontr) - assume "\ 1 \ y" - hence "y=0" by auto - hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto - thus False using asm by auto - qed - qed - ultimately show ?case using Suc by auto + also have "normalize \ = smult (normalize c) (normalize p)" + by (subst normalize_mult) (simp add: normalize_const_poly) + finally show ?thesis . qed -lemma order_0_monom [simp]: - assumes "c \ 0" - shows "order 0 (monom c n) = n" - using assms order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) - -lemma dvd_imp_order_le: - "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" - by (auto simp: order_mult elim: dvdE) - -text\Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\ - -lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" -apply (cases "p = 0", auto) -apply (drule order_2 [where a=a and p=p]) -apply (metis not_less_eq_eq power_le_dvd) -apply (erule power_le_dvd [OF order_1]) -done - -lemma monom_1_dvd_iff: - assumes "p \ 0" - shows "monom 1 n dvd p \ n \ Polynomial.order 0 p" - using assms order_divides[of 0 n p] by (simp add: monom_altdef) - -lemma poly_squarefree_decomp_order: - assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" - and p: "p = q * d" - and p': "pderiv p = e * d" - and d: "d = r * p + s * pderiv p" - shows "order a q = (if order a p = 0 then 0 else 1)" -proof (rule classical) - assume 1: "order a q \ (if order a p = 0 then 0 else 1)" - from \pderiv p \ 0\ have "p \ 0" by auto - with p have "order a p = order a q + order a d" - by (simp add: order_mult) - with 1 have "order a p \ 0" by (auto split: if_splits) - have "order a (pderiv p) = order a e + order a d" - using \pderiv p \ 0\ \pderiv p = e * d\ by (simp add: order_mult) - have "order a p = Suc (order a (pderiv p))" - using \pderiv p \ 0\ \order a p \ 0\ by (rule order_pderiv) - have "d \ 0" using \p \ 0\ \p = q * d\ by simp - have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" - apply (simp add: d) - apply (rule dvd_add) - apply (rule dvd_mult) - apply (simp add: order_divides \p \ 0\ - \order a p = Suc (order a (pderiv p))\) - apply (rule dvd_mult) - apply (simp add: order_divides) - done - then have "order a (pderiv p) \ order a d" - using \d \ 0\ by (simp add: order_divides) - show ?thesis - using \order a p = order a q + order a d\ - using \order a (pderiv p) = order a e + order a d\ - using \order a p = Suc (order a (pderiv p))\ - using \order a (pderiv p) \ order a d\ - by auto -qed - -lemma poly_squarefree_decomp_order2: - "\pderiv p \ (0 :: 'a :: field_char_0 poly); - p = q * d; - pderiv p = e * d; - d = r * p + s * pderiv p - \ \ \a. order a q = (if order a p = 0 then 0 else 1)" -by (blast intro: poly_squarefree_decomp_order) - -lemma order_pderiv2: - "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ - \ (order a (pderiv p) = n) = (order a p = Suc n)" -by (auto dest: order_pderiv) - -definition - rsquarefree :: "'a::idom poly => bool" where - "rsquarefree p = (p \ 0 & (\a. (order a p = 0) | (order a p = 1)))" - -lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h :: 'a :: {semidom,semiring_char_0}:]" - by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) - -lemma rsquarefree_roots: - fixes p :: "'a :: field_char_0 poly" - shows "rsquarefree p = (\a. \(poly p a = 0 \ poly (pderiv p) a = 0))" -apply (simp add: rsquarefree_def) -apply (case_tac "p = 0", simp, simp) -apply (case_tac "pderiv p = 0") -apply simp -apply (drule pderiv_iszero, clarsimp) -apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) -apply (force simp add: order_root order_pderiv2) -done - -lemma poly_squarefree_decomp: - assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" - and "p = q * d" - and "pderiv p = e * d" - and "d = r * p + s * pderiv p" - shows "rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" +lemma smult_content_normalize_primitive_part [simp]: + "smult (content p) (normalize (primitive_part p)) = normalize p" proof - - from \pderiv p \ 0\ have "p \ 0" by auto - with \p = q * d\ have "q \ 0" by simp - have "\a. order a q = (if order a p = 0 then 0 else 1)" - using assms by (rule poly_squarefree_decomp_order2) - with \p \ 0\ \q \ 0\ show ?thesis - by (simp add: rsquarefree_def order_root) -qed - -lemma coeff_monom_mult: - "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" -proof - - have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" - by (simp add: coeff_mult) - also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" - by (intro sum.cong) simp_all - also have "\ = (if k < n then 0 else c * coeff p (k - n))" by (simp add: sum.delta') + have "smult (content p) (normalize (primitive_part p)) = + normalize ([:content p:] * primitive_part p)" + by (subst normalize_mult) (simp_all add: normalize_const_poly) + also have "[:content p:] * primitive_part p = p" by simp finally show ?thesis . qed -lemma monom_1_dvd_iff': - "monom 1 n dvd p \ (\k 'a poly \ 'a poly \ 'a poly \ bool" + where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" + | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" + | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y + \ x = q * y + r \ eucl_rel_poly x y (q, r)" + +lemma eucl_rel_poly_iff: + "eucl_rel_poly x y (q, r) \ + x = q * y + r \ + (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" + by (auto elim: eucl_rel_poly.cases + intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) + +lemma eucl_rel_poly_0: + "eucl_rel_poly 0 y (0, 0)" + unfolding eucl_rel_poly_iff by simp + +lemma eucl_rel_poly_by_0: + "eucl_rel_poly x 0 (0, x)" + unfolding eucl_rel_poly_iff by simp + +lemma eucl_rel_poly_pCons: + assumes rel: "eucl_rel_poly x y (q, r)" + assumes y: "y \ 0" + assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" + shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" + (is "eucl_rel_poly ?x y (?q, ?r)") +proof - + have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" + using assms unfolding eucl_rel_poly_iff by simp_all + + have 1: "?x = ?q * y + ?r" + using b x by simp + + have 2: "?r = 0 \ degree ?r < degree y" + proof (rule eq_zero_or_degree_less) + show "degree ?r \ degree y" + proof (rule degree_diff_le) + show "degree (pCons a r) \ degree y" + using r by auto + show "degree (smult b y) \ degree y" + by (rule degree_smult_le) + qed + next + show "coeff ?r (degree y) = 0" + using \y \ 0\ unfolding b by simp + qed + + from 1 2 show ?thesis + unfolding eucl_rel_poly_iff + using \y \ 0\ by simp +qed + +lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" +apply (cases "y = 0") +apply (fast intro!: eucl_rel_poly_by_0) +apply (induct x) +apply (fast intro!: eucl_rel_poly_0) +apply (fast intro!: eucl_rel_poly_pCons) +done + +lemma eucl_rel_poly_unique: + assumes 1: "eucl_rel_poly x y (q1, r1)" + assumes 2: "eucl_rel_poly x y (q2, r2)" + shows "q1 = q2 \ r1 = r2" +proof (cases "y = 0") + assume "y = 0" with assms show ?thesis + by (simp add: eucl_rel_poly_iff) +next + assume [simp]: "y \ 0" + from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" + unfolding eucl_rel_poly_iff by simp_all + from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" + unfolding eucl_rel_poly_iff by simp_all + from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" + by (simp add: algebra_simps) + from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" + by (auto intro: degree_diff_less) + + show "q1 = q2 \ r1 = r2" + proof (rule ccontr) + assume "\ (q1 = q2 \ r1 = r2)" + with q3 have "q1 \ q2" and "r1 \ r2" by auto + with r3 have "degree (r2 - r1) < degree y" by simp + also have "degree y \ degree (q1 - q2) + degree y" by simp + also have "\ = degree ((q1 - q2) * y)" + using \q1 \ q2\ by (simp add: degree_mult_eq) + also have "\ = degree (r2 - r1)" + using q3 by simp + finally have "degree (r2 - r1) < degree (r2 - r1)" . + then show "False" by simp + qed +qed + +lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" +by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) + +lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" +by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) + +lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] + +lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] + +instantiation poly :: (field) ring_div +begin + +definition modulo_poly where + mod_poly_def: "f mod g \ + if g = 0 then f + else pseudo_mod (smult ((1/coeff g (degree g)) ^ (Suc (degree f) - degree g)) f) g" + +lemma eucl_rel_poly: "eucl_rel_poly (x::'a::field poly) y (x div y, x mod y)" + unfolding eucl_rel_poly_iff +proof (intro conjI) + show "x = x div y * y + x mod y" + proof(cases "y = 0") + case True show ?thesis by(simp add: True divide_poly_def divide_poly_0 mod_poly_def) + next + case False + then have "pseudo_divmod (smult ((1 / coeff y (degree y)) ^ (Suc (degree x) - degree y)) x) y = + (x div y, x mod y)" + unfolding divide_poly_field mod_poly_def pseudo_mod_def by simp + from pseudo_divmod[OF False this] + show ?thesis using False + by (simp add: power_mult_distrib[symmetric] mult.commute) + qed + show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" + proof (cases "y = 0") + case True then show ?thesis by auto + next + case False + with pseudo_mod[OF this] show ?thesis unfolding mod_poly_def by simp + qed +qed + +lemma div_poly_eq: + "eucl_rel_poly (x::'a::field poly) y (q, r) \ x div y = q" + by(rule eucl_rel_poly_unique_div[OF eucl_rel_poly]) + +lemma mod_poly_eq: + "eucl_rel_poly (x::'a::field poly) y (q, r) \ x mod y = r" + by (rule eucl_rel_poly_unique_mod[OF eucl_rel_poly]) + +instance proof - assume "monom 1 n dvd p" - then obtain r where r: "p = monom 1 n * r" by (elim dvdE) - thus "\k 0" + hence "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)" + using eucl_rel_poly [of x y] + by (simp add: eucl_rel_poly_iff distrib_right) + thus "(x + z * y) div y = z + x div y" + by (rule div_poly_eq) next - assume zero: "(\kk. coeff p (k + n))" - have "\\<^sub>\k. coeff p (k + n) = 0" - by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, - subst cofinite_eq_sequentially [symmetric]) transfer - hence coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def - by (subst poly.Abs_poly_inverse) simp_all - have "p = monom 1 n * r" - by (intro poly_eqI, subst coeff_monom_mult) (insert zero, simp_all) - thus "monom 1 n dvd p" by simp + fix x y z :: "'a poly" + assume "x \ 0" + show "(x * y) div (x * z) = y div z" + proof (cases "y \ 0 \ z \ 0") + have "\x::'a poly. eucl_rel_poly x 0 (0, x)" + by (rule eucl_rel_poly_by_0) + then have [simp]: "\x::'a poly. x div 0 = 0" + by (rule div_poly_eq) + have "\x::'a poly. eucl_rel_poly 0 x (0, 0)" + by (rule eucl_rel_poly_0) + then have [simp]: "\x::'a poly. 0 div x = 0" + by (rule div_poly_eq) + case False then show ?thesis by auto + next + case True then have "y \ 0" and "z \ 0" by auto + with \x \ 0\ + have "\q r. eucl_rel_poly y z (q, r) \ eucl_rel_poly (x * y) (x * z) (q, x * r)" + by (auto simp add: eucl_rel_poly_iff algebra_simps) + (rule classical, simp add: degree_mult_eq) + moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" . + ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" . + then show ?thesis by (simp add: div_poly_eq) + qed +qed + +end + +lemma degree_mod_less: + "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" + using eucl_rel_poly [of x y] + unfolding eucl_rel_poly_iff by simp + +lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" + using degree_mod_less[of b a] by auto + +lemma div_poly_less: "degree (x::'a::field poly) < degree y \ x div y = 0" +proof - + assume "degree x < degree y" + hence "eucl_rel_poly x y (0, x)" + by (simp add: eucl_rel_poly_iff) + thus "x div y = 0" by (rule div_poly_eq) +qed + +lemma mod_poly_less: "degree x < degree y \ x mod y = x" +proof - + assume "degree x < degree y" + hence "eucl_rel_poly x y (0, x)" + by (simp add: eucl_rel_poly_iff) + thus "x mod y = x" by (rule mod_poly_eq) +qed + +lemma eucl_rel_poly_smult_left: + "eucl_rel_poly x y (q, r) + \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" + unfolding eucl_rel_poly_iff by (simp add: smult_add_right) + +lemma div_smult_left: "(smult (a::'a::field) x) div y = smult a (x div y)" + by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) + +lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" + by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) + +lemma poly_div_minus_left [simp]: + fixes x y :: "'a::field poly" + shows "(- x) div y = - (x div y)" + using div_smult_left [of "- 1::'a"] by simp + +lemma poly_mod_minus_left [simp]: + fixes x y :: "'a::field poly" + shows "(- x) mod y = - (x mod y)" + using mod_smult_left [of "- 1::'a"] by simp + +lemma eucl_rel_poly_add_left: + assumes "eucl_rel_poly x y (q, r)" + assumes "eucl_rel_poly x' y (q', r')" + shows "eucl_rel_poly (x + x') y (q + q', r + r')" + using assms unfolding eucl_rel_poly_iff + by (auto simp add: algebra_simps degree_add_less) + +lemma poly_div_add_left: + fixes x y z :: "'a::field poly" + shows "(x + y) div z = x div z + y div z" + using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] + by (rule div_poly_eq) + +lemma poly_mod_add_left: + fixes x y z :: "'a::field poly" + shows "(x + y) mod z = x mod z + y mod z" + using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] + by (rule mod_poly_eq) + +lemma poly_div_diff_left: + fixes x y z :: "'a::field poly" + shows "(x - y) div z = x div z - y div z" + by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) + +lemma poly_mod_diff_left: + fixes x y z :: "'a::field poly" + shows "(x - y) mod z = x mod z - y mod z" + by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) + +lemma eucl_rel_poly_smult_right: + "a \ 0 \ eucl_rel_poly x y (q, r) + \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" + unfolding eucl_rel_poly_iff by simp + +lemma div_smult_right: + "(a::'a::field) \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" + by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) + +lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" + by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) + +lemma poly_div_minus_right [simp]: + fixes x y :: "'a::field poly" + shows "x div (- y) = - (x div y)" + using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) + +lemma poly_mod_minus_right [simp]: + fixes x y :: "'a::field poly" + shows "x mod (- y) = x mod y" + using mod_smult_right [of "- 1::'a"] by simp + +lemma eucl_rel_poly_mult: + "eucl_rel_poly x y (q, r) \ eucl_rel_poly q z (q', r') + \ eucl_rel_poly x (y * z) (q', y * r' + r)" +apply (cases "z = 0", simp add: eucl_rel_poly_iff) +apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff) +apply (cases "r = 0") +apply (cases "r' = 0") +apply (simp add: eucl_rel_poly_iff) +apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq) +apply (cases "r' = 0") +apply (simp add: eucl_rel_poly_iff degree_mult_eq) +apply (simp add: eucl_rel_poly_iff field_simps) +apply (simp add: degree_mult_eq degree_add_less) +done + +lemma poly_div_mult_right: + fixes x y z :: "'a::field poly" + shows "x div (y * z) = (x div y) div z" + by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) + +lemma poly_mod_mult_right: + fixes x y z :: "'a::field poly" + shows "x mod (y * z) = y * (x div y mod z) + x mod y" + by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) + +lemma mod_pCons: + fixes a and x + assumes y: "y \ 0" + defines b: "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" + shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" +unfolding b +apply (rule mod_poly_eq) +apply (rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) +done + +definition pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" +where + "pdivmod p q = (p div q, p mod q)" + +lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ pdivmod p q = (r, s)" + by (metis pdivmod_def eucl_rel_poly eucl_rel_poly_unique) + +lemma pdivmod_0: + "pdivmod 0 q = (0, 0)" + by (simp add: pdivmod_def) + +lemma pdivmod_pCons: + "pdivmod (pCons a p) q = + (if q = 0 then (0, pCons a p) else + (let (s, r) = pdivmod p q; + b = coeff (pCons a r) (degree q) / coeff q (degree q) + in (pCons b s, pCons a r - smult b q)))" + apply (simp add: pdivmod_def Let_def, safe) + apply (rule div_poly_eq) + apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl]) + apply (rule mod_poly_eq) + apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl]) + done + +lemma pdivmod_fold_coeffs: + "pdivmod p q = (if q = 0 then (0, p) + else fold_coeffs (\a (s, r). + let b = coeff (pCons a r) (degree q) / coeff q (degree q) + in (pCons b s, pCons a r - smult b q) + ) p (0, 0))" + apply (cases "q = 0") + apply (simp add: pdivmod_def) + apply (rule sym) + apply (induct p) + apply (simp_all add: pdivmod_0 pdivmod_pCons) + apply (case_tac "a = 0 \ p = 0") + apply (auto simp add: pdivmod_def) + done + + +subsubsection \List-based versions for fast implementation\ +(* Subsection by: + Sebastiaan Joosten + René Thiemann + Akihisa Yamada + *) +fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where + "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" +| "minus_poly_rev_list xs [] = xs" +| "minus_poly_rev_list [] (y # ys) = []" + +fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list + \ nat \ 'a list \ 'a list" where + "pseudo_divmod_main_list lc q r d (Suc n) = (let + rr = map (op * lc) r; + a = hd r; + qqq = cCons a (map (op * lc) q); + rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) + in pseudo_divmod_main_list lc qqq rrr d n)" +| "pseudo_divmod_main_list lc q r d 0 = (q,r)" + +fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list + \ nat \ 'a list" where + "pseudo_mod_main_list lc r d (Suc n) = (let + rr = map (op * lc) r; + a = hd r; + rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) + in pseudo_mod_main_list lc rrr d n)" +| "pseudo_mod_main_list lc r d 0 = r" + + +fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list + \ nat \ 'a list \ 'a list" where + "divmod_poly_one_main_list q r d (Suc n) = (let + a = hd r; + qqq = cCons a q; + rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) + in divmod_poly_one_main_list qqq rr d n)" +| "divmod_poly_one_main_list q r d 0 = (q,r)" + +fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list + \ nat \ 'a list" where + "mod_poly_one_main_list r d (Suc n) = (let + a = hd r; + rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) + in mod_poly_one_main_list rr d n)" +| "mod_poly_one_main_list r d 0 = r" + +definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where + "pseudo_divmod_list p q = + (if q = [] then ([],p) else + (let rq = rev q; + (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in + (qu,rev re)))" + +definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where + "pseudo_mod_list p q = + (if q = [] then p else + (let rq = rev q; + re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in + (rev re)))" + +lemma minus_zero_does_nothing: +"(minus_poly_rev_list x (map (op * 0) y)) = (x :: 'a :: ring list)" + by(induct x y rule: minus_poly_rev_list.induct, auto) + +lemma length_minus_poly_rev_list[simp]: + "length (minus_poly_rev_list xs ys) = length xs" + by(induct xs ys rule: minus_poly_rev_list.induct, auto) + +lemma if_0_minus_poly_rev_list: + "(if a = 0 then x else minus_poly_rev_list x (map (op * a) y)) + = minus_poly_rev_list x (map (op * (a :: 'a :: ring)) y)" + by(cases "a=0",simp_all add:minus_zero_does_nothing) + +lemma Poly_append: + "Poly ((a::'a::comm_semiring_1 list) @ b) = Poly a + monom 1 (length a) * Poly b" + by (induct a,auto simp: monom_0 monom_Suc) + +lemma minus_poly_rev_list: "length p \ length (q :: 'a :: comm_ring_1 list) \ + Poly (rev (minus_poly_rev_list (rev p) (rev q))) + = Poly p - monom 1 (length p - length q) * Poly q" +proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) + case (1 x xs y ys) + have "length (rev q) \ length (rev p)" using 1 by simp + from this[folded 1(2,3)] have ys_xs:"length ys \ length xs" by simp + hence a:"Poly (rev (minus_poly_rev_list xs ys)) = + Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" + by(subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev],auto) + have "Poly p - monom 1 (length p - length q) * Poly q + = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" + by simp + also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" + unfolding 1(2,3) by simp + also have "\ = Poly (rev xs) + monom x (length xs) - + (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" using ys_xs + by (simp add:Poly_append distrib_left mult_monom smult_monom) + also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" + unfolding a diff_monom[symmetric] by(simp) + finally show ?case + unfolding 1(2,3)[symmetric] by (simp add: smult_monom Poly_append) +qed auto + +lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" + using smult_monom [of a _ n] by (metis mult_smult_left) + +lemma head_minus_poly_rev_list: + "length d \ length r \ d\[] \ + hd (minus_poly_rev_list (map (op * (last d :: 'a :: comm_ring)) r) (map (op * (hd r)) (rev d))) = 0" +proof(induct r) + case (Cons a rs) + thus ?case by(cases "rev d", simp_all add:ac_simps) +qed simp + +lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)" +proof (induct p) + case(Cons x xs) thus ?case by (cases "Poly xs = 0",auto) +qed simp + +lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" + by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) + +lemma pseudo_divmod_main_list_invar : + assumes leading_nonzero:"last d \ 0" + and lc:"last d = lc" + and dNonempty:"d \ []" + and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q',rev r')" + and "n = (1 + length r - length d)" + shows + "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = + (Poly q', Poly r')" +using assms(4-) +proof(induct "n" arbitrary: r q) +case (Suc n r q) + have ifCond: "\ Suc (length r) \ length d" using Suc.prems by simp + have rNonempty:"r \ []" + using ifCond dNonempty using Suc_leI length_greater_0_conv list.size(3) by fastforce + let ?a = "(hd (rev r))" + let ?rr = "map (op * lc) (rev r)" + let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map (op * ?a) (rev d))))" + let ?qq = "cCons ?a (map (op * lc) q)" + have n: "n = (1 + length r - length d - 1)" + using ifCond Suc(3) by simp + have rr_val:"(length ?rrr) = (length r - 1)" using ifCond by auto + hence rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" + using rNonempty ifCond unfolding One_nat_def by auto + from ifCond have id: "Suc (length r) - length d = Suc (length r - length d)" by auto + have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" + using Suc.prems ifCond by (simp add:Let_def if_0_minus_poly_rev_list id) + hence v:"pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" + using n by auto + have sucrr:"Suc (length r) - length d = Suc (length r - length d)" + using Suc_diff_le ifCond not_less_eq_eq by blast + have n_ok : "n = 1 + (length ?rrr) - length d" using Suc(3) rNonempty by simp + have cong: "\ x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ + pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp + have hd_rev:"coeff (Poly r) (length r - Suc 0) = hd (rev r)" + using last_coeff_is_hd[OF rNonempty] by simp + show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def + proof (rule cong[OF _ _ refl], goal_cases) + case 1 + show ?case unfolding monom_Suc hd_rev[symmetric] + by (simp add: smult_monom Poly_map) + next + case 2 + show ?case + proof (subst Poly_on_rev_starting_with_0, goal_cases) + show "hd (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))) = 0" + by (fold lc, subst head_minus_poly_rev_list, insert ifCond dNonempty,auto) + from ifCond have "length d \ length r" by simp + then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = + Poly (rev (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))))" + by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] + minus_poly_rev_list) + qed + qed simp +qed simp + +lemma pseudo_divmod_impl[code]: "pseudo_divmod (f::'a::comm_ring_1 poly) g = + map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" +proof (cases "g=0") +case False + hence coeffs_g_nonempty:"(coeffs g) \ []" by simp + hence lastcoeffs:"last (coeffs g) = coeff g (degree g)" + by (simp add: hd_rev last_coeffs_eq_coeff_degree not_0_coeffs_not_Nil) + obtain q r where qr: "pseudo_divmod_main_list + (last (coeffs g)) (rev []) + (rev (coeffs f)) (rev (coeffs g)) + (1 + length (coeffs f) - + length (coeffs g)) = (q,rev (rev r))" by force + then have qr': "pseudo_divmod_main_list + (hd (rev (coeffs g))) [] + (rev (coeffs f)) (rev (coeffs g)) + (1 + length (coeffs f) - + length (coeffs g)) = (q,r)" using hd_rev[OF coeffs_g_nonempty] by(auto) + from False have cg: "(coeffs g = []) = False" by auto + have last_non0:"last (coeffs g) \ 0" using False by (simp add:last_coeffs_not_0) + show ?thesis + unfolding pseudo_divmod_def pseudo_divmod_list_def Let_def qr' map_prod_def split cg if_False + pseudo_divmod_main_list_invar[OF last_non0 _ _ qr,unfolded lastcoeffs,simplified,symmetric,OF False] + poly_of_list_def + using False by (auto simp: degree_eq_length_coeffs) +next + case True + show ?thesis unfolding True unfolding pseudo_divmod_def pseudo_divmod_list_def + by auto +qed + +lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q + xs ys n) = pseudo_mod_main_list l xs ys n" + by (induct n arbitrary: l q xs ys, auto simp: Let_def) + +lemma pseudo_mod_impl[code]: "pseudo_mod f g = + poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" +proof - + have snd_case: "\ f g p. snd ((\ (x,y). (f x, g y)) p) = g (snd p)" + by auto + show ?thesis + unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def + pseudo_mod_list_def Let_def + by (simp add: snd_case pseudo_mod_main_list) +qed + + +(* *************** *) +subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ + +lemma pdivmod_via_pseudo_divmod: "pdivmod f g = (if g = 0 then (0,f) + else let + ilc = inverse (coeff g (degree g)); + h = smult ilc g; + (q,r) = pseudo_divmod f h + in (smult ilc q, r))" (is "?l = ?r") +proof (cases "g = 0") + case False + define lc where "lc = inverse (coeff g (degree g))" + define h where "h = smult lc g" + from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" unfolding h_def lc_def by auto + hence h0: "h \ 0" by auto + obtain q r where p: "pseudo_divmod f h = (q,r)" by force + from False have id: "?r = (smult lc q, r)" + unfolding Let_def h_def[symmetric] lc_def[symmetric] p by auto + from pseudo_divmod[OF h0 p, unfolded h1] + have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto + have "eucl_rel_poly f h (q, r)" unfolding eucl_rel_poly_iff using f r h0 by auto + hence "pdivmod f h = (q,r)" by (simp add: pdivmod_pdivmodrel) + hence "pdivmod f g = (smult lc q, r)" + unfolding pdivmod_def h_def div_smult_right[OF lc] mod_smult_right[OF lc] + using lc by auto + with id show ?thesis by auto +qed (auto simp: pdivmod_def) + +lemma pdivmod_via_pseudo_divmod_list: "pdivmod f g = (let + cg = coeffs g + in if cg = [] then (0,f) + else let + cf = coeffs f; + ilc = inverse (last cg); + ch = map (op * ilc) cg; + (q,r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) + in (poly_of_list (map (op * ilc) q), poly_of_list (rev r)))" +proof - + note d = pdivmod_via_pseudo_divmod + pseudo_divmod_impl pseudo_divmod_list_def + show ?thesis + proof (cases "g = 0") + case True thus ?thesis unfolding d by auto + next + case False + define ilc where "ilc = inverse (coeff g (degree g))" + from False have ilc: "ilc \ 0" unfolding ilc_def by auto + with False have id: "(g = 0) = False" "(coeffs g = []) = False" + "last (coeffs g) = coeff g (degree g)" + "(coeffs (smult ilc g) = []) = False" + by (auto simp: last_coeffs_eq_coeff_degree) + have id2: "hd (rev (coeffs (smult ilc g))) = 1" + by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) + have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" + "rev (coeffs (smult ilc g)) = rev (map (op * ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto + obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map (op * ilc) (coeffs g))) + (1 + length (coeffs f) - length (coeffs g)) = (q,r)" by force + show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 + unfolding id3 pair map_prod_def split by (auto simp: Poly_map) + qed +qed + +lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" +proof (intro ext, goal_cases) + case (1 q r d n) + { + fix xs :: "'a list" + have "map (op * 1) xs = xs" by (induct xs, auto) + } note [simp] = this + show ?case by (induct n arbitrary: q r d, auto simp: Let_def) +qed + +fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list + \ nat \ 'a list" where + "divide_poly_main_list lc q r d (Suc n) = (let + cr = hd r + in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let + a = cr div lc; + qq = cCons a q; + rr = minus_poly_rev_list r (map (op * a) d) + in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" +| "divide_poly_main_list lc q r d 0 = q" + + +lemma divide_poly_main_list_simp[simp]: "divide_poly_main_list lc q r d (Suc n) = (let + cr = hd r; + a = cr div lc; + qq = cCons a q; + rr = minus_poly_rev_list r (map (op * a) d) + in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" + by (simp add: Let_def minus_zero_does_nothing) + +declare divide_poly_main_list.simps(1)[simp del] + +definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where + "divide_poly_list f g = + (let cg = coeffs g + in if cg = [] then g + else let cf = coeffs f; cgr = rev cg + in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" + +lemmas pdivmod_via_divmod_list[code] = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] + +lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" + by (induct n arbitrary: q r d, auto simp: Let_def) + +lemma mod_poly_code[code]: "f mod g = + (let cg = coeffs g + in if cg = [] then f + else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg; + r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) + in poly_of_list (rev r))" (is "?l = ?r") +proof - + have "?l = snd (pdivmod f g)" unfolding pdivmod_def by simp + also have "\ = ?r" unfolding pdivmod_via_divmod_list Let_def + mod_poly_one_main_list[symmetric, of _ _ _ Nil] by (auto split: prod.splits) + finally show ?thesis . +qed + +definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where + "div_field_poly_impl f g = ( + let cg = coeffs g + in if cg = [] then 0 + else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg; + q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) + in poly_of_list ((map (op * ilc) q)))" + +text \We do not declare the following lemma as code equation, since then polynomial division + on non-fields will no longer be executable. However, a code-unfold is possible, since + \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ +lemma div_field_poly_impl[code_unfold]: "op div = div_field_poly_impl" +proof (intro ext) + fix f g :: "'a poly" + have "f div g = fst (pdivmod f g)" unfolding pdivmod_def by simp + also have "\ = div_field_poly_impl f g" unfolding + div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) + finally show "f div g = div_field_poly_impl f g" . +qed + + +lemma divide_poly_main_list: + assumes lc0: "lc \ 0" + and lc:"last d = lc" + and d:"d \ []" + and "n = (1 + length r - length d)" + shows + "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = + divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" +using assms(4-) +proof(induct "n" arbitrary: r q) +case (Suc n r q) + have ifCond: "\ Suc (length r) \ length d" using Suc.prems by simp + have r: "r \ []" + using ifCond d using Suc_leI length_greater_0_conv list.size(3) by fastforce + then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases, auto) + from d lc obtain dd where d: "d = dd @ [lc]" + by (cases d rule: rev_cases, auto) + from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) + from ifCond have len: "length dd \ length rr" by (simp add: r d) + show ?case + proof (cases "lcr div lc * lc = lcr") + case False + thus ?thesis unfolding Suc(2)[symmetric] using r d + by (auto simp add: Let_def nth_default_append) + next + case True + hence id: + "?thesis = (Poly (divide_poly_main_list lc (cCons (lcr div lc) q) + (rev (rev (minus_poly_rev_list (rev rr) (rev (map (op * (lcr div lc)) dd))))) (rev d) n) = + divide_poly_main lc + (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) + (Poly r - monom (lcr div lc) n * Poly d) + (Poly d) (length rr - 1) n)" + using r d + by (cases r rule: rev_cases; cases "d" rule: rev_cases; + auto simp add: Let_def rev_map nth_default_append) + have cong: "\ x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ + divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp + show ?thesis unfolding id + proof (subst Suc(1), simp add: n, + subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) + case 2 + have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" + by (simp add: mult_monom len True) + thus ?case unfolding r d Poly_append n ring_distribs + by (auto simp: Poly_map smult_monom smult_monom_mult) + qed (auto simp: len monom_Suc smult_monom) + qed +qed simp + + +lemma divide_poly_list[code]: "f div g = divide_poly_list f g" +proof - + note d = divide_poly_def divide_poly_list_def + show ?thesis + proof (cases "g = 0") + case True + show ?thesis unfolding d True by auto + next + case False + then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases, auto) + with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto + from cg False have lcg: "coeff g (degree g) = lcg" + using last_coeffs_eq_coeff_degree last_snoc by force + with False have lcg0: "lcg \ 0" by auto + from cg have ltp: "Poly (cg @ [lcg]) = g" + using Poly_coeffs [of g] by auto + show ?thesis unfolding d cg Let_def id if_False poly_of_list_def + by (subst divide_poly_main_list, insert False cg lcg0, auto simp: lcg ltp, + simp add: degree_eq_length_coeffs) + qed qed no_notation cCons (infixr "##" 65) diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Library/Polynomial_Factorial.thy --- a/src/HOL/Library/Polynomial_Factorial.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Library/Polynomial_Factorial.thy Thu Jan 05 22:38:06 2017 +0100 @@ -19,40 +19,6 @@ lemma prod_mset_const_poly: "prod_mset (image_mset (\x. [:f x:]) A) = [:prod_mset (image_mset f A):]" by (induction A) (simp_all add: one_poly_def mult_ac) -lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" -proof - - have "smult c p = [:c:] * p" by simp - also have "\ dvd 1 \ c dvd 1 \ p dvd 1" - proof safe - assume A: "[:c:] * p dvd 1" - thus "p dvd 1" by (rule dvd_mult_right) - from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE) - have "c dvd c * (coeff p 0 * coeff q 0)" by simp - also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) - also note B [symmetric] - finally show "c dvd 1" by simp - next - assume "c dvd 1" "p dvd 1" - from \c dvd 1\ obtain d where "1 = c * d" by (erule dvdE) - hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac) - hence "[:c:] dvd 1" by (rule dvdI) - from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp - qed - finally show ?thesis . -qed - -lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" - using degree_mod_less[of b a] by auto - -lemma smult_eq_iff: - assumes "(b :: 'a :: field) \ 0" - shows "smult a p = smult b q \ smult (a / b) p = q" -proof - assume "smult a p = smult b q" - also from assms have "smult (inverse b) \ = q" by simp - finally show "smult (a / b) p = q" by (simp add: field_simps) -qed (insert assms, auto) - lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ irreducible c" @@ -160,145 +126,6 @@ by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) -subsection \Content and primitive part of a polynomial\ - -definition content :: "('a :: semiring_Gcd poly) \ 'a" where - "content p = Gcd (set (coeffs p))" - -lemma content_0 [simp]: "content 0 = 0" - by (simp add: content_def) - -lemma content_1 [simp]: "content 1 = 1" - by (simp add: content_def) - -lemma content_const [simp]: "content [:c:] = normalize c" - by (simp add: content_def cCons_def) - -lemma const_poly_dvd_iff_dvd_content: - fixes c :: "'a :: semiring_Gcd" - shows "[:c:] dvd p \ c dvd content p" -proof (cases "p = 0") - case [simp]: False - have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) - also have "\ \ (\a\set (coeffs p). c dvd a)" - proof safe - fix n :: nat assume "\a\set (coeffs p). c dvd a" - thus "c dvd coeff p n" - by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) - qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) - also have "\ \ c dvd content p" - by (simp add: content_def dvd_Gcd_iff mult.commute [of "unit_factor x" for x] - dvd_mult_unit_iff lead_coeff_nonzero) - finally show ?thesis . -qed simp_all - -lemma content_dvd [simp]: "[:content p:] dvd p" - by (subst const_poly_dvd_iff_dvd_content) simp_all - -lemma content_dvd_coeff [simp]: "content p dvd coeff p n" - by (cases "n \ degree p") - (auto simp: content_def coeffs_def not_le coeff_eq_0 simp del: upt_Suc intro: Gcd_dvd) - -lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" - by (simp add: content_def Gcd_dvd) - -lemma normalize_content [simp]: "normalize (content p) = content p" - by (simp add: content_def) - -lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" -proof - assume "is_unit (content p)" - hence "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) - thus "content p = 1" by simp -qed auto - -lemma content_smult [simp]: "content (smult c p) = normalize c * content p" - by (simp add: content_def coeffs_smult Gcd_mult) - -lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" - by (auto simp: content_def simp: poly_eq_iff coeffs_def) - -definition primitive_part :: "'a :: {semiring_Gcd,idom_divide} poly \ 'a poly" where - "primitive_part p = (if p = 0 then 0 else map_poly (\x. x div content p) p)" - -lemma primitive_part_0 [simp]: "primitive_part 0 = 0" - by (simp add: primitive_part_def) - -lemma content_times_primitive_part [simp]: - fixes p :: "'a :: {idom_divide, semiring_Gcd} poly" - shows "smult (content p) (primitive_part p) = p" -proof (cases "p = 0") - case False - thus ?thesis - unfolding primitive_part_def - by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs - intro: map_poly_idI) -qed simp_all - -lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" -proof (cases "p = 0") - case False - hence "primitive_part p = map_poly (\x. x div content p) p" - by (simp add: primitive_part_def) - also from False have "\ = 0 \ p = 0" - by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) - finally show ?thesis using False by simp -qed simp - -lemma content_primitive_part [simp]: - assumes "p \ 0" - shows "content (primitive_part p) = 1" -proof - - have "p = smult (content p) (primitive_part p)" by simp - also have "content \ = content p * content (primitive_part p)" - by (simp del: content_times_primitive_part) - finally show ?thesis using assms by simp -qed - -lemma content_decompose: - fixes p :: "'a :: semiring_Gcd poly" - obtains p' where "p = smult (content p) p'" "content p' = 1" -proof (cases "p = 0") - case True - thus ?thesis by (intro that[of 1]) simp_all -next - case False - from content_dvd[of p] obtain r where r: "p = [:content p:] * r" by (erule dvdE) - have "content p * 1 = content p * content r" by (subst r) simp - with False have "content r = 1" by (subst (asm) mult_left_cancel) simp_all - with r show ?thesis by (intro that[of r]) simp_all -qed - -lemma smult_content_normalize_primitive_part [simp]: - "smult (content p) (normalize (primitive_part p)) = normalize p" -proof - - have "smult (content p) (normalize (primitive_part p)) = - normalize ([:content p:] * primitive_part p)" - by (subst normalize_mult) (simp_all add: normalize_const_poly) - also have "[:content p:] * primitive_part p = p" by simp - finally show ?thesis . -qed - -lemma content_dvd_contentI [intro]: - "p dvd q \ content p dvd content q" - using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast - -lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" - by (simp add: primitive_part_def map_poly_pCons) - -lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" - by (auto simp: primitive_part_def) - -lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" -proof (cases "p = 0") - case False - have "p = smult (content p) (primitive_part p)" by simp - also from False have "degree \ = degree (primitive_part p)" - by (subst degree_smult_eq) simp_all - finally show ?thesis .. -qed simp_all - - subsection \Lifting polynomial coefficients to the field of fractions\ abbreviation (input) fract_poly @@ -690,7 +517,7 @@ fix p :: "'a poly" show "unit_factor_field_poly p * normalize_field_poly p = p" by (cases "p = 0") - (simp_all add: unit_factor_field_poly_def normalize_field_poly_def lead_coeff_nonzero) + (simp_all add: unit_factor_field_poly_def normalize_field_poly_def) next fix p :: "'a poly" assume "is_unit p" thus "normalize_field_poly p = 1" @@ -698,7 +525,7 @@ next fix p :: "'a poly" assume "p \ 0" thus "is_unit (unit_factor_field_poly p)" - by (simp add: unit_factor_field_poly_def lead_coeff_nonzero is_unit_pCons_iff) + by (simp add: unit_factor_field_poly_def is_unit_pCons_iff) next fix p q s :: "'a poly" assume "s \ 0" moreover assume "euclidean_size_field_poly p < euclidean_size_field_poly q" @@ -943,7 +770,7 @@ have "fract_poly p = unit_factor_field_poly (fract_poly p) * normalize_field_poly (fract_poly p)" by simp also have "unit_factor_field_poly (fract_poly p) = [:to_fract (lead_coeff p):]" - by (simp add: unit_factor_field_poly_def lead_coeff_def monom_0 degree_map_poly coeff_map_poly) + by (simp add: unit_factor_field_poly_def monom_0 degree_map_poly coeff_map_poly) also from assms have "normalize_field_poly (fract_poly p) = prod_mset ?P" by (subst field_poly_prod_mset_prime_factorization) simp_all also have "\ = prod_mset (image_mset id ?P)" by simp diff -r 81a5473e6d04 -r a0e1f64be67c src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Thu Jan 05 22:37:52 2017 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Thu Jan 05 22:38:06 2017 +0100 @@ -176,6 +176,21 @@ for x :: "'a::real_vector" using scaleR_minus_left [of 1 x] by simp +lemma scaleR_2: + fixes x :: "'a::real_vector" + shows "scaleR 2 x = x + x" + unfolding one_add_one [symmetric] scaleR_left_distrib by simp + +lemma scaleR_half_double [simp]: + fixes a :: "'a::real_vector" + shows "(1 / 2) *\<^sub>R (a + a) = a" +proof - + have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" + by (metis scaleR_2 scaleR_scaleR) + then show ?thesis + by simp +qed + class real_algebra = real_vector + ring + assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)" and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"