# HG changeset patch # User Andreas Lochbihler # Date 1448986714 -3600 # Node ID f58d75535f666a8e474fdc61468c3b9c994df4c7 # Parent 507b39df1a57fa11087d4b43490f1eba80d1a7a5# Parent 96d2c1b9a30a62f97f3e1acf43aedb03b3824620 merged diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Complex.thy Tue Dec 01 17:18:34 2015 +0100 @@ -748,9 +748,6 @@ subsubsection \Complex exponential\ -abbreviation Exp :: "complex \ complex" - where "Exp \ exp" - lemma cis_conv_exp: "cis b = exp (\ * b)" proof - { fix n :: nat @@ -766,29 +763,29 @@ intro!: sums_unique sums_add sums_mult sums_of_real) qed -lemma Exp_eq_polar: "Exp z = exp (Re z) * cis (Im z)" +lemma exp_eq_polar: "exp z = exp (Re z) * cis (Im z)" unfolding cis_conv_exp exp_of_real [symmetric] mult_exp_exp by (cases z) simp lemma Re_exp: "Re (exp z) = exp (Re z) * cos (Im z)" - unfolding Exp_eq_polar by simp + unfolding exp_eq_polar by simp lemma Im_exp: "Im (exp z) = exp (Re z) * sin (Im z)" - unfolding Exp_eq_polar by simp + unfolding exp_eq_polar by simp lemma norm_cos_sin [simp]: "norm (Complex (cos t) (sin t)) = 1" by (simp add: norm_complex_def) lemma norm_exp_eq_Re [simp]: "norm (exp z) = exp (Re z)" - by (simp add: cis.code cmod_complex_polar Exp_eq_polar) + by (simp add: cis.code cmod_complex_polar exp_eq_polar) -lemma complex_Exp_Ex: "\a r. z = complex_of_real r * Exp a" +lemma complex_exp_exists: "\a r. z = complex_of_real r * exp a" apply (insert rcis_Ex [of z]) - apply (auto simp add: Exp_eq_polar rcis_def mult.assoc [symmetric]) + apply (auto simp add: exp_eq_polar rcis_def mult.assoc [symmetric]) apply (rule_tac x = "ii * complex_of_real a" in exI, auto) done -lemma Exp_two_pi_i [simp]: "Exp((2::complex) * complex_of_real pi * ii) = 1" - by (simp add: Exp_eq_polar complex_eq_iff) +lemma exp_two_pi_i [simp]: "exp(2 * complex_of_real pi * ii) = 1" + by (simp add: exp_eq_polar complex_eq_iff) subsubsection \Complex argument\ diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Tue Dec 01 17:18:34 2015 +0100 @@ -1644,9 +1644,8 @@ "(real_of_int (a::int) + b > 0) = (real_of_int a + real_of_int (floor b) > 0 \ (real_of_int a + real_of_int (floor b) = 0 \ real_of_int (floor b) - b < 0))" proof- have th: "(real_of_int a + b >0) = (real_of_int (-a) + (-b)< 0)" by arith - show ?thesis using myless[of _ "real_of_int (floor b)"] - by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) - (simp add: algebra_simps,arith) + show ?thesis + by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) (auto simp add: algebra_simps) qed lemma split_int_le_real: @@ -3765,8 +3764,7 @@ proof- let ?ss = "s - real_of_int (floor s)" from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] - of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" - by (auto simp add: myle[of _ "s", symmetric] myless[of "?ss"]) + of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" by (auto simp: floor_less_cancel) from np have n0: "real_of_int n \ 0" by simp from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] have nu0:"real_of_int n * u - s \ -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto @@ -4807,7 +4805,7 @@ shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real_of_int i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") proof- have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real_of_int i)#bs) (exsplit p))" - by (simp add: myless[of _ "1"] myless[of _ "0"] ac_simps) + by auto also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real_of_int i + x) #bs) p)" by (simp only: exsplit[OF qf] ac_simps) also have "\ = (\ x. Ifm (x#bs) p)" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Finite_Set.thy Tue Dec 01 17:18:34 2015 +0100 @@ -319,6 +319,16 @@ apply (simp add: finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2) done +lemma finite_finite_vimage_IntI: + assumes "finite F" and "\y. y \ F \ finite ((h -` {y}) \ A)" + shows "finite (h -` F \ A)" +proof - + have *: "h -` F \ A = (\ y\F. (h -` {y}) \ A)" + by blast + show ?thesis + by (simp only: * assms finite_UN_I) +qed + lemma finite_vimageI: "finite F \ inj h \ finite (h -` F)" using finite_vimage_IntI[of F h UNIV] by auto diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Groups.thy --- a/src/HOL/Groups.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Groups.thy Tue Dec 01 17:18:34 2015 +0100 @@ -999,6 +999,9 @@ apply (simp add: algebra_simps) done +lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \ a > b" +by (simp add: less_diff_eq) + lemma diff_le_eq[algebra_simps, field_simps]: "a - b \ c \ a \ c + b" by (auto simp add: le_less diff_less_eq ) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Inequalities.thy --- a/src/HOL/Inequalities.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Inequalities.thy Tue Dec 01 17:18:34 2015 +0100 @@ -66,7 +66,6 @@ using assms by (cases "i \ j") (auto simp: algebra_simps) } hence "?S \ 0" by (auto intro!: setsum_nonpos simp: mult_le_0_iff) - (auto simp: field_simps) finally show ?thesis by (simp add: algebra_simps) qed diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Library/BigO.thy Tue Dec 01 17:18:34 2015 +0100 @@ -200,8 +200,6 @@ apply (auto simp add: fun_Compl_def func_plus) apply (drule_tac x = x in spec)+ apply force - apply (drule_tac x = x in spec)+ - apply force done lemma bigo_abs: "(\x. abs (f x)) =o O(f)" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Library/Float.thy Tue Dec 01 17:18:34 2015 +0100 @@ -1116,10 +1116,11 @@ proof - have "0 \ log 2 x - real_of_int \log 2 x\" by (simp add: algebra_simps) - from this assms + with assms show ?thesis - by (auto simp: truncate_down_def round_down_def mult_powr_eq + apply (auto simp: truncate_down_def round_down_def mult_powr_eq intro!: ge_one_powr_ge_zero mult_pos_pos) + by linarith qed lemma truncate_down_nonneg: "0 \ y \ 0 \ truncate_down prec y" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Library/Infinite_Set.thy Tue Dec 01 17:18:34 2015 +0100 @@ -72,10 +72,7 @@ by(induction rule: finite_psubset_induct)(meson Diff_subset card_Diff1_less card_psubset finite_Diff step) qed -text \ - As a concrete example, we prove that the set of natural numbers is - infinite. -\ +text \As a concrete example, we prove that the set of natural numbers is infinite.\ lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) \ (\m. \n\m. n \ S)" using frequently_cofinite[of "\x. x \ S"] @@ -94,6 +91,7 @@ lemma finite_nat_bounded: "finite (S::nat set) \ \k. S \ {.. For a set of natural numbers to be infinite, it is enough to know that for any number larger than some \k\, there is some larger @@ -150,6 +148,32 @@ obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) +proposition finite_image_absD: + fixes S :: "'a::linordered_ring set" + shows "finite (abs ` S) \ finite S" + by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) + +text \The set of integers is also infinite.\ + +lemma infinite_int_iff_infinite_nat_abs: "infinite (S::int set) \ infinite ((nat o abs) ` S)" + by (auto simp: transfer_nat_int_set_relations o_def image_comp dest: finite_image_absD) + +proposition infinite_int_iff_unbounded_le: "infinite (S::int set) \ (\m. \n. abs n \ m \ n \ S)" + apply (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) + apply (metis abs_ge_zero nat_le_eq_zle le_nat_iff) + done + +proposition infinite_int_iff_unbounded: "infinite (S::int set) \ (\m. \n. abs n > m \ n \ S)" + apply (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) + apply (metis (full_types) nat_le_iff nat_mono not_le) + done + +proposition finite_int_iff_bounded: "finite (S::int set) \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" + using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) + subsection "Infinitely Many and Almost All" text \ @@ -385,24 +409,5 @@ unfolding bij_betw_def by (auto intro: enumerate_in_set) qed -subsection "Miscellaneous" - -text \ - A few trivial lemmas about sets that contain at most one element. - These simplify the reasoning about deterministic automata. -\ - -definition atmost_one :: "'a set \ bool" - where "atmost_one S \ (\x y. x\S \ y\S \ x = y)" - -lemma atmost_one_empty: "S = {} \ atmost_one S" - by (simp add: atmost_one_def) - -lemma atmost_one_singleton: "S = {x} \ atmost_one S" - by (simp add: atmost_one_def) - -lemma atmost_one_unique [elim]: "atmost_one S \ x \ S \ y \ S \ y = x" - by (simp add: atmost_one_def) - end diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Tue Dec 01 17:18:34 2015 +0100 @@ -3012,7 +3012,7 @@ \ ball (p t) (ee (p t))" apply (intro subset_path_image_join pi_hgn pi_ghn') using \N>0\ Suc.prems - apply (auto simp: dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) + apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) done have pi0: "(f has_contour_integral 0) (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ @@ -3492,7 +3492,7 @@ by (simp add: winding_number_valid_path) lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" - by (simp add: winding_number_valid_path) + by (simp add: path_image_subpath winding_number_valid_path) lemma winding_number_join: assumes g1: "path g1" "z \ path_image g1" @@ -3742,7 +3742,7 @@ by (rule continuous_at_imp_continuous_within) have gdx: "\ differentiable at x" using x by (simp add: g_diff_at) - have "((\c. Exp (- integral {a..c} (\x. vector_derivative \ (at x) / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) + have "((\c. exp (- integral {a..c} (\x. vector_derivative \ (at x) / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) (at x within {a..b})" using x gdx t apply (clarsimp simp add: differentiable_iff_scaleR) @@ -3781,7 +3781,7 @@ "pathstart p = pathstart \" "pathfinish p = pathfinish \" "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF assms, of 1] by auto - have [simp]: "(winding_number \ z \ \) = (Exp (contour_integral p (\w. 1 / (w - z))) = 1)" + have [simp]: "(winding_number \ z \ \) = (exp (contour_integral p (\w. 1 / (w - z))) = 1)" using p by (simp add: exp_eq_1 complex_is_Int_iff) have "winding_number p z \ \ \ pathfinish p = pathstart p" using p z @@ -3840,7 +3840,7 @@ using eqArg by (simp add: i_def) have gpdt: "\ piecewise_C1_differentiable_on {0..t}" by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) - have "Exp (- i) * (\ t - z) = \ 0 - z" + have "exp (- i) * (\ t - z) = \ 0 - z" unfolding i_def apply (rule winding_number_exp_integral [OF gpdt]) using t z unfolding path_image_def @@ -3855,7 +3855,7 @@ apply (subst Complex_Transcendental.Arg_eq [of r]) apply (simp add: iArg) using * - apply (simp add: Exp_eq_polar field_simps) + apply (simp add: exp_eq_polar field_simps) done with t show ?thesis by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) @@ -4225,8 +4225,8 @@ also have "... = winding_number (subpath 0 x \) z" apply (subst winding_number_valid_path) using assms x - apply (simp_all add: valid_path_subpath) - by (force simp: closed_segment_eq_real_ivl path_image_def) + apply (simp_all add: path_image_subpath valid_path_subpath) + by (force simp: path_image_def) finally show ?thesis . qed show ?thesis @@ -4277,7 +4277,7 @@ have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" using winding_number_exp_2pi [of "subpath 0 t \" z] apply (simp add: t \ valid_path_imp_path) - using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp add: Euler sub12) + using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) have "b < a \ \ 0" proof - have "\ 0 \ {c. b < a \ c}" @@ -4321,7 +4321,7 @@ have "isCont (winding_number \) z" by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < abs(Re(winding_number \ z)) - 1/2" - using continuous_at_eps_delta wnz_12 diff_less_iff(1) by blast + using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast def z' \ "z - (d / (2 * cmod a)) *\<^sub>R a" have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Complex_Transcendental.thy --- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Tue Dec 01 17:18:34 2015 +0100 @@ -654,7 +654,6 @@ done qed - corollary shows Arg_ge_0: "0 \ Arg z" and Arg_lt_2pi: "Arg z < 2*pi" @@ -772,7 +771,7 @@ lemma Arg_inverse: "Arg(inverse z) = (if z \ \ \ 0 \ Re z then Arg z else 2*pi - Arg z)" apply (cases "z=0", simp) apply (rule Arg_unique [of "inverse (norm z)"]) - using Arg_ge_0 [of z] Arg_lt_2pi [of z] Arg_eq [of z] Arg_eq_0 [of z] Exp_two_pi_i + using Arg_ge_0 [of z] Arg_lt_2pi [of z] Arg_eq [of z] Arg_eq_0 [of z] exp_two_pi_i apply (auto simp: of_real_numeral algebra_simps exp_diff divide_simps) done @@ -849,8 +848,11 @@ by auto lemma Arg_exp: "0 \ Im z \ Im z < 2*pi \ Arg(exp z) = Im z" - by (rule Arg_unique [of "exp(Re z)"]) (auto simp: Exp_eq_polar) - + by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) + +lemma complex_split_polar: + obtains r a::real where "z = complex_of_real r * (cos a + \ * sin a)" "0 \ r" "0 \ a" "a < 2*pi" + using Arg cis.ctr cis_conv_exp by fastforce subsection\Analytic properties of tangent function\ @@ -898,7 +900,7 @@ have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \ pi" unfolding ln_complex_def apply (rule theI [where a = "Complex (ln(norm z)) \"]) using z assms \ - apply (auto simp: field_simps exp_complex_eqI Exp_eq_polar cis.code) + apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) done then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \ pi" by auto @@ -1516,14 +1518,14 @@ shows "((\z. z powr r :: complex) has_field_derivative r * z powr (r - 1)) (at z)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "eventually (\z. z \ 0) (nhds z)" by (intro t1_space_nhds) auto - thus "eventually (\z. z powr r = Exp (r * Ln z)) (nhds z)" + thus "eventually (\z. z powr r = exp (r * Ln z)) (nhds z)" unfolding powr_def by eventually_elim simp - have "((\z. Exp (r * Ln z)) has_field_derivative Exp (r * Ln z) * (inverse z * r)) (at z)" + have "((\z. exp (r * Ln z)) has_field_derivative exp (r * Ln z) * (inverse z * r)) (at z)" using assms by (auto intro!: derivative_eq_intros has_field_derivative_powr) - also have "Exp (r * Ln z) * (inverse z * r) = r * z powr (r - 1)" + also have "exp (r * Ln z) * (inverse z * r) = r * z powr (r - 1)" unfolding powr_def by (simp add: assms exp_diff field_simps) - finally show "((\z. Exp (r * Ln z)) has_field_derivative r * z powr (r - 1)) (at z)" + finally show "((\z. exp (r * Ln z)) has_field_derivative r * z powr (r - 1)) (at z)" by simp qed @@ -2405,7 +2407,7 @@ lemma Re_Arcsin_bound: "abs(Re(Arcsin z)) \ pi" by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) - + subsection\Interrelations between Arcsin and Arccos\ @@ -2481,7 +2483,6 @@ apply (simp add: cos_squared_eq) using assms apply (auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff) - apply (auto simp: algebra_simps) done lemma sin_cos_csqrt: @@ -2491,7 +2492,6 @@ apply (simp add: sin_squared_eq) using assms apply (auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff) - apply (auto simp: algebra_simps) done lemma Arcsin_Arccos_csqrt_pos: @@ -2661,7 +2661,7 @@ by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" - apply (simp add: abs_square_le_1 diff_le_iff arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) + apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arcsin_Arccos_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done @@ -2671,7 +2671,7 @@ by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" - apply (simp add: abs_square_le_1 diff_le_iff arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) + apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arccos_Arcsin_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Dec 01 17:18:34 2015 +0100 @@ -4397,7 +4397,7 @@ using \y \ s\ proof - show "inner (y - z) z < inner (y - z) y" - apply (subst diff_less_iff(1)[symmetric]) + apply (subst diff_gt_0_iff_gt [symmetric]) unfolding inner_diff_right[symmetric] and inner_gt_zero_iff using \y\s\ \z\s\ apply auto diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Tue Dec 01 17:18:34 2015 +0100 @@ -688,7 +688,7 @@ "x \ {a <..< b}" "(\xa. f' x xa - (f b - f a) / (b - a) * xa) = (\v. 0)" .. then show ?thesis - by (metis (hide_lams) assms(1) diff_less_iff(1) eq_iff_diff_eq_0 + by (metis (hide_lams) assms(1) diff_gt_0_iff_gt eq_iff_diff_eq_0 zero_less_mult_iff nonzero_mult_divide_cancel_right not_real_square_gt_zero times_divide_eq_left) qed diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Dec 01 17:18:34 2015 +0100 @@ -599,7 +599,7 @@ then have cd_ne: "\i\Basis. c \ i \ d \ i" using assms unfolding box_ne_empty by auto have "\i. i \ Basis \ 0 \ b \ i - a \ i" - using ab_ne by (metis diff_le_iff(1)) + using ab_ne by auto moreover have "\i. i \ Basis \ b \ i - a \ i \ d \ i - c \ i" using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)] diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Tue Dec 01 17:18:34 2015 +0100 @@ -375,7 +375,7 @@ lemma path_compose_reversepath: "f o reversepath p = reversepath(f o p)" by (rule ext) (simp add: reversepath_def) -lemma join_paths_eq: +lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" @@ -453,8 +453,6 @@ lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by (simp add: path_join) -lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join - lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" @@ -563,18 +561,18 @@ definition subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" -lemma path_image_subpath_gen [simp]: - fixes g :: "real \ 'a::real_normed_vector" +lemma path_image_subpath_gen: + fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" apply (simp add: closed_segment_real_eq path_image_def subpath_def) apply (subst o_def [of g, symmetric]) apply (simp add: image_comp [symmetric]) done -lemma path_image_subpath [simp]: +lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" - by (simp add: closed_segment_eq_real_ivl) + by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" @@ -614,7 +612,7 @@ lemma affine_ineq: fixes x :: "'a::linordered_idom" - assumes "x \ 1" "v < u" + assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" @@ -726,7 +724,7 @@ lemma path_image_subpath_subset: "\path g; u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" - apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost) + apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath) apply (auto simp: path_image_def) done @@ -805,7 +803,7 @@ apply (rule that [OF `0 \ u` `u \ 1`]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) using `0 \ u` g0 disj - apply (simp add:) + apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) apply (drule_tac x="y/u" in spec) @@ -825,7 +823,7 @@ show ?thesis apply (rule that [of "subpath 0 u g"]) using assms u - apply simp_all + apply (simp_all add: path_image_subpath) apply (simp add: pathstart_def) apply (force simp: closed_segment_eq_real_ivl path_image_def) done @@ -966,7 +964,7 @@ unfolding linepath_def by (intro continuous_intros) -lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)" +lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) @@ -982,6 +980,9 @@ unfolding reversepath_def linepath_def by auto +lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" + by (simp add: linepath_def) + lemma arc_linepath: assumes "a \ b" shows "arc (linepath a b)" @@ -1566,7 +1567,7 @@ have CC: "1 \ 1 + (C - 1) * u" using `x \ a` `0 \ u` apply (simp add: C_def divide_simps norm_minus_commute) - by (metis Bx diff_le_iff(1) less_eq_real_def mult_nonneg_nonneg) + using Bx by auto have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = @@ -1601,7 +1602,7 @@ have DD: "1 \ 1 + (D - 1) * u" using `y \ a` `0 \ u` apply (simp add: D_def divide_simps norm_minus_commute) - by (metis By diff_le_iff(1) less_eq_real_def mult_nonneg_nonneg) + using By by auto have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = @@ -2793,7 +2794,7 @@ proposition homotopic_paths_sym_eq: "homotopic_paths s p q \ homotopic_paths s q p" by (metis homotopic_paths_sym) -proposition homotopic_paths_trans: +proposition homotopic_paths_trans [trans]: "\homotopic_paths s p q; homotopic_paths s q r\ \ homotopic_paths s p r" apply (simp add: homotopic_paths_def) apply (rule homotopic_with_trans, assumption) @@ -3262,4 +3263,83 @@ by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) qed +subsection\ Homotopy and subpaths\ + +lemma homotopic_join_subpaths1: + assumes "path g" and pag: "path_image g \ s" + and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" + shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" +proof - + have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t + using affine_ineq \u \ v\ by fastforce + have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t + by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) + have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto + show ?thesis + apply (rule homotopic_paths_subset [OF _ pag]) + using assms + apply (cases "w = u") + using homotopic_paths_rinv [of "subpath u v g" "path_image g"] + apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl) + apply (rule homotopic_paths_sym) + apply (rule homotopic_paths_reparametrize + [where f = "\t. if t \ 1 / 2 + then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t + else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"]) + using \path g\ path_subpath u w apply blast + using \path g\ path_image_subpath_subset u w(1) apply blast + apply simp_all + apply (subst split_01) + apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ + apply (simp_all add: field_simps not_le) + apply (force dest!: t2) + apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2) + apply (simp add: joinpaths_def subpath_def) + apply (force simp: algebra_simps) + done +qed + +lemma homotopic_join_subpaths2: + assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" + shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)" +by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) + +lemma homotopic_join_subpaths3: + assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" + and "path g" and pag: "path_image g \ s" + and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" + shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)" +proof - + have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" + apply (rule homotopic_paths_join) + using hom homotopic_paths_sym_eq apply blast + apply (metis \path g\ homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w) + apply (simp add:) + done + also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" + apply (rule homotopic_paths_sym [OF homotopic_paths_assoc]) + using assms by (simp_all add: path_image_subpath_subset [THEN order_trans]) + also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g) + (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" + apply (rule homotopic_paths_join) + apply (metis \path g\ homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v) + apply (metis (no_types, lifting) \path g\ homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w) + apply (simp add:) + done + also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" + apply (rule homotopic_paths_rid) + using \path g\ path_subpath u v apply blast + apply (meson \path g\ order.trans pag path_image_subpath_subset u v) + done + finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" . + then show ?thesis + using homotopic_join_subpaths2 by blast +qed + +proposition homotopic_join_subpaths: + "\path g; path_image g \ s; u \ {0..1}; v \ {0..1}; w \ {0..1}\ + \ homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" +apply (rule le_cases3 [of u v w]) +using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+ + end diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Dec 01 17:18:34 2015 +0100 @@ -817,6 +817,9 @@ definition cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" +definition sphere :: "'a::metric_space \ real \ 'a set" + where "sphere x e = {y. dist x y = e}" + lemma mem_ball [simp]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) @@ -863,19 +866,6 @@ lemma cball_diff_eq_sphere: "cball a r - ball a r = {x. dist x a = r}" by (auto simp: cball_def ball_def dist_commute) -lemma diff_less_iff: - "(a::real) - b > 0 \ a > b" - "(a::real) - b < 0 \ a < b" - "a - b < c \ a < c + b" "a - b > c \ a > c + b" - by arith+ - -lemma diff_le_iff: - "(a::real) - b \ 0 \ a \ b" - "(a::real) - b \ 0 \ a \ b" - "a - b \ c \ a \ c + b" - "a - b \ c \ a \ c + b" - by arith+ - lemma open_ball [intro, simp]: "open (ball x e)" proof - have "open (dist x -` {.. (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult.commute diff_le_iff inner_distrib inner_diff_left) + apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) done } moreover @@ -7357,7 +7347,7 @@ then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult.commute diff_le_iff inner_distrib inner_diff_left) + apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) done } ultimately show ?thesis using False by (auto simp: cbox_def) @@ -8187,8 +8177,8 @@ shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" (is "?lhs = ?rhs") proof - assume ?lhs - then show ?rhs + assume ?lhs + then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next @@ -8209,13 +8199,13 @@ using \a \ a'\ by (simp add: abs_mult_pos field_simps) finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" by linarith show ?thesis - using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ + using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ by (simp add: dist_norm scaleR_add_left) qed then show ?rhs by (simp add: dist_norm) qed next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs apply (auto simp: ball_def dist_norm ) apply (metis add.commute add_le_cancel_right dist_norm dist_triangle_alt order_trans) using le_less_trans apply fastforce @@ -8227,8 +8217,8 @@ shows "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" (is "?lhs = ?rhs") proof - assume ?lhs - then show ?rhs + assume ?lhs + then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next @@ -8256,7 +8246,7 @@ then show ?rhs by (simp add: dist_norm) qed next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs apply (auto simp: ball_def dist_norm ) apply (metis add.commute add_le_cancel_right dist_norm dist_triangle_alt le_less_trans) using le_less_trans apply fastforce @@ -8268,10 +8258,10 @@ shows "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") - case True then show ?thesis + case True then show ?thesis using dist_not_less_zero less_le_trans by force next - case False show ?thesis + case False show ?thesis proof assume ?lhs then have "(cball a r \ cball a' r')" @@ -8280,7 +8270,7 @@ using False cball_subset_cball_iff by fastforce next assume ?rhs with False show ?lhs - using ball_subset_cball cball_subset_cball_iff by blast + using ball_subset_cball cball_subset_cball_iff by blast qed qed @@ -8289,10 +8279,10 @@ shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") - case True then show ?thesis + case True then show ?thesis using dist_not_less_zero less_le_trans by force next - case False show ?thesis + case False show ?thesis proof assume ?lhs then have "0 < r'" @@ -8316,22 +8306,22 @@ shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof - assume ?lhs - then show ?rhs + assume ?lhs + then show ?rhs proof (cases "d \ 0 \ e \ 0") - case True + case True with \?lhs\ show ?rhs by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) next case False - with \?lhs\ show ?rhs + with \?lhs\ show ?rhs apply (auto simp add: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs by (auto simp add: set_eq_subset ball_subset_ball_iff) qed @@ -8340,22 +8330,22 @@ shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof - assume ?lhs - then show ?rhs + assume ?lhs + then show ?rhs proof (cases "d < 0 \ e < 0") - case True + case True with \?lhs\ show ?rhs by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) next case False - with \?lhs\ show ?rhs + with \?lhs\ show ?rhs apply (auto simp add: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs by (auto simp add: set_eq_subset cball_subset_cball_iff) qed @@ -8363,7 +8353,7 @@ fixes x :: "'a :: euclidean_space" shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") proof - assume ?lhs + assume ?lhs then show ?rhs apply (auto simp add: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) @@ -8371,7 +8361,7 @@ using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ done next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs by (auto simp add: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff) qed @@ -8379,7 +8369,7 @@ fixes x :: "'a :: euclidean_space" shows "cball x d = ball y e \ d < 0 \ e \ 0" (is "?lhs = ?rhs") proof - assume ?lhs + assume ?lhs then show ?rhs apply (auto simp add: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) @@ -8387,7 +8377,7 @@ using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ done next - assume ?rhs then show ?lhs + assume ?rhs then show ?lhs by (auto simp add: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff) qed diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Multivariate_Analysis/Weierstrass.thy --- a/src/HOL/Multivariate_Analysis/Weierstrass.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Weierstrass.thy Tue Dec 01 17:18:34 2015 +0100 @@ -380,9 +380,7 @@ by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric]) also have "... \ (1/(k * (p t))^n) * 1" apply (rule mult_left_mono [OF power_le_one]) - apply (metis diff_le_iff(1) less_eq_real_def mult.commute power_le_one power_mult ptn_pos ptn_le) - using pt_pos [OF t] \k>0\ - apply auto + using pt_pos \k>0\ p01 power_le_one t apply auto done also have "... \ (1 / (k*\))^n" using \k>0\ \01 power_mono pt_\ t diff -r 507b39df1a57 -r f58d75535f66 src/HOL/NSA/Examples/NSPrimes.thy --- a/src/HOL/NSA/Examples/NSPrimes.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/NSA/Examples/NSPrimes.thy Tue Dec 01 17:18:34 2015 +0100 @@ -61,7 +61,7 @@ (* Goldblatt: Exercise 5.11(3a) - p 57 *) lemma starprime: "starprime = {p. 1 < p & (\m. m dvd p --> m = 1 | m = p)}" -by (transfer, auto simp add: prime_nat_def) +by (transfer, auto simp add: prime_def) (* Goldblatt Exercise 5.11(3b) - p 57 *) lemma hyperprime_factor_exists [rule_format]: @@ -262,17 +262,14 @@ text{*Already proved as @{text primes_infinite}, but now using non-standard naturals.*} theorem not_finite_prime: "~ finite {p::nat. prime p}" apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2]) -apply (cut_tac hypnat_dvd_all_hypnat_of_nat) -apply (erule exE) -apply (erule conjE) -apply (subgoal_tac "1 < N + 1") -prefer 2 apply (blast intro: hypnat_add_one_gt_one) +using hypnat_dvd_all_hypnat_of_nat +apply clarify +apply (drule hypnat_add_one_gt_one) apply (drule hyperprime_factor_exists) -apply auto +apply clarify apply (subgoal_tac "k \ hypnat_of_nat ` {p. prime p}") -apply (force simp add: starprime_def, safe) -apply (drule_tac x = x in bspec, auto) -apply (metis add.commute hdvd_diff hdvd_one_eq_one hypnat_diff_add_inverse2 hypnat_one_not_prime) +apply (force simp add: starprime_def) +apply (metis Compl_iff add.commute dvd_add_left_iff empty_iff hdvd_one_eq_one hypnat_one_not_prime imageE insert_iff mem_Collect_eq zero_not_prime_nat) done end diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Number_Theory/Eratosthenes.thy --- a/src/HOL/Number_Theory/Eratosthenes.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Number_Theory/Eratosthenes.thy Tue Dec 01 17:18:34 2015 +0100 @@ -294,8 +294,8 @@ from 2 show ?thesis apply (auto simp add: numbers_of_marks_sieve numeral_2_eq_2 set_primes_upto dest: prime_gt_Suc_0_nat) - apply (metis One_nat_def Suc_le_eq less_not_refl prime_nat_def) - apply (metis One_nat_def Suc_le_eq aux prime_nat_def) + apply (metis One_nat_def Suc_le_eq less_not_refl prime_def) + apply (metis One_nat_def Suc_le_eq aux prime_def) done qed qed diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Number_Theory/Pocklington.thy --- a/src/HOL/Number_Theory/Pocklington.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Number_Theory/Pocklington.thy Tue Dec 01 17:18:34 2015 +0100 @@ -457,11 +457,11 @@ proof assume "prime n" then show ?rhs - by (metis one_not_prime_nat prime_nat_def) + by (metis one_not_prime_nat prime_def) next assume ?rhs with False show "prime n" - by (auto simp: prime_def) (metis One_nat_def prime_factor_nat prime_nat_def) + by (auto simp: prime_def) (metis One_nat_def prime_factor_nat prime_def) qed qed @@ -538,7 +538,7 @@ and pp: "prime p" and pn: "p dvd n" shows "[p = 1] (mod q)" proof - - have p01: "p \ 0" "p \ 1" using pp one_not_prime_nat zero_not_prime_nat by auto + have p01: "p \ 0" "p \ 1" using pp one_not_prime_nat zero_not_prime_nat by (auto intro: prime_gt_0_nat) obtain k where k: "a ^ (q * r) - 1 = n*k" by (metis an cong_to_1_nat dvd_def nqr) from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast @@ -689,7 +689,7 @@ from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast from n m have m0: "m > 0" "m\0" by auto have "1 < p" - by (metis p(1) prime_nat_def) + by (metis p(1) prime_def) with m0 m have mn: "m < n" by auto from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" .. from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Number_Theory/Primes.thy --- a/src/HOL/Number_Theory/Primes.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Number_Theory/Primes.thy Tue Dec 01 17:18:34 2015 +0100 @@ -37,38 +37,33 @@ definition prime :: "nat \ bool" where "prime p = (1 < p \ (\m. m dvd p \ m = 1 \ m = p))" -lemmas prime_nat_def = prime_def - - subsection \Primes\ lemma prime_odd_nat: "prime p \ p > 2 \ odd p" - apply (auto simp add: prime_nat_def even_iff_mod_2_eq_zero dvd_eq_mod_eq_0) + apply (auto simp add: prime_def even_iff_mod_2_eq_zero dvd_eq_mod_eq_0) apply (metis dvd_eq_mod_eq_0 even_Suc mod_by_1 nat_dvd_not_less not_mod_2_eq_0_eq_1 zero_less_numeral) done -(* FIXME Is there a better way to handle these, rather than making them elim rules? *) +lemma prime_gt_0_nat: "prime p \ p > 0" + unfolding prime_def by auto -lemma prime_gt_0_nat [elim]: "prime p \ p > 0" - unfolding prime_nat_def by auto +lemma prime_ge_1_nat: "prime p \ p >= 1" + unfolding prime_def by auto -lemma prime_ge_1_nat [elim]: "prime p \ p >= 1" - unfolding prime_nat_def by auto +lemma prime_gt_1_nat: "prime p \ p > 1" + unfolding prime_def by auto -lemma prime_gt_1_nat [elim]: "prime p \ p > 1" - unfolding prime_nat_def by auto - -lemma prime_ge_Suc_0_nat [elim]: "prime p \ p >= Suc 0" - unfolding prime_nat_def by auto +lemma prime_ge_Suc_0_nat: "prime p \ p >= Suc 0" + unfolding prime_def by auto -lemma prime_gt_Suc_0_nat [elim]: "prime p \ p > Suc 0" - unfolding prime_nat_def by auto +lemma prime_gt_Suc_0_nat: "prime p \ p > Suc 0" + unfolding prime_def by auto -lemma prime_ge_2_nat [elim]: "prime p \ p >= 2" - unfolding prime_nat_def by auto +lemma prime_ge_2_nat: "prime p \ p >= 2" + unfolding prime_def by auto lemma prime_imp_coprime_nat: "prime p \ \ p dvd n \ coprime p n" - apply (unfold prime_nat_def) + apply (unfold prime_def) apply (metis gcd_dvd1_nat gcd_dvd2_nat) done @@ -105,7 +100,7 @@ lemma not_prime_eq_prod_nat: "(n::nat) > 1 \ ~ prime n \ EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n" - unfolding prime_nat_def dvd_def apply auto + unfolding prime_def dvd_def apply auto by (metis mult.commute linorder_neq_iff linorder_not_le mult_1 n_less_n_mult_m one_le_mult_iff less_imp_le_nat) @@ -129,18 +124,18 @@ subsubsection \Make prime naively executable\ lemma zero_not_prime_nat [simp]: "~prime (0::nat)" - by (simp add: prime_nat_def) + by (simp add: prime_def) lemma one_not_prime_nat [simp]: "~prime (1::nat)" - by (simp add: prime_nat_def) + by (simp add: prime_def) lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)" - by (simp add: prime_nat_def) + by (simp add: prime_def) lemma prime_nat_code [code]: "prime p \ p > 1 \ (\n \ {1<..One property of coprimality is easier to prove via prime factors.\ @@ -239,7 +234,8 @@ by (rule dvd_diff_nat) then have "p dvd 1" by simp then have "p <= 1" by auto - moreover from \prime p\ have "p > 1" by auto + moreover from \prime p\ have "p > 1" + using prime_def by blast ultimately have False by auto} then have "n < p" by presburger with \prime p\ and \p <= fact n + 1\ show ?thesis by auto @@ -270,7 +266,7 @@ proof - from assms have "1 < p * q" and P: "\m. m dvd p * q \ m = 1 \ m = p * q" - unfolding prime_nat_def by auto + unfolding prime_def by auto from \1 < p * q\ have "p \ 0" by (cases p) auto then have Q: "p = p * q \ q = 1" by auto have "p dvd p * q" by simp diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Number_Theory/UniqueFactorization.thy --- a/src/HOL/Number_Theory/UniqueFactorization.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy Tue Dec 01 17:18:34 2015 +0100 @@ -107,9 +107,7 @@ ultimately have "a ^ count M a dvd a ^ count N a" by (elim coprime_dvd_mult_nat) with a show ?thesis - apply (intro power_dvd_imp_le) - apply auto - done + using power_dvd_imp_le prime_def by blast next case False then show ?thesis @@ -247,14 +245,14 @@ using assms apply auto done -lemma prime_factors_gt_0_nat [elim]: +lemma prime_factors_gt_0_nat: fixes p :: nat shows "p \ prime_factors x \ p > 0" - by (auto dest!: prime_factors_prime_nat) + using prime_factors_prime_nat by force -lemma prime_factors_gt_0_int [elim]: +lemma prime_factors_gt_0_int: shows "x \ 0 \ p \ prime_factors x \ int p > (0::int)" - by auto + by (simp add: prime_factors_gt_0_nat) lemma prime_factors_finite_nat [iff]: fixes n :: nat @@ -303,7 +301,8 @@ proof - from assms have "f \ multiset" by (auto simp add: multiset_def) - moreover from assms have "n > 0" by force + moreover from assms have "n > 0" + by (auto intro: prime_gt_0_nat) ultimately have "multiset_prime_factorization n = Abs_multiset f" apply (unfold multiset_prime_factorization_def) apply (subst if_P, assumption) @@ -723,16 +722,16 @@ (\p \ prime_factors x \ prime_factors y. p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?z") proof - - have [arith]: "?z > 0" - by auto + have [arith]: "?z > 0" + using prime_factors_gt_0_nat by auto have aux: "\p. prime p \ multiplicity p ?z = min (multiplicity p x) (multiplicity p y)" apply (subst multiplicity_prod_prime_powers_nat) apply auto done have "?z dvd x" - by (intro multiplicity_dvd'_nat) (auto simp add: aux) + by (intro multiplicity_dvd'_nat) (auto simp add: aux intro: prime_gt_0_nat) moreover have "?z dvd y" - by (intro multiplicity_dvd'_nat) (auto simp add: aux) + by (intro multiplicity_dvd'_nat) (auto simp add: aux intro: prime_gt_0_nat) moreover have "w dvd x \ w dvd y \ w dvd ?z" for w proof (cases "w = 0") case True @@ -758,7 +757,7 @@ (is "_ = ?z") proof - have [arith]: "?z > 0" - by auto + by (auto intro: prime_gt_0_nat) have aux: "\p. prime p \ multiplicity p ?z = max (multiplicity p x) (multiplicity p y)" apply (subst multiplicity_prod_prime_powers_nat) apply auto @@ -776,7 +775,7 @@ then show ?thesis apply auto apply (rule multiplicity_dvd'_nat) - apply (auto intro: dvd_multiplicity_nat simp add: aux) + apply (auto intro: prime_gt_0_nat dvd_multiplicity_nat simp add: aux) done qed ultimately have "?z = lcm x y" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Old_Number_Theory/Legacy_GCD.thy --- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy Tue Dec 01 17:18:34 2015 +0100 @@ -665,7 +665,8 @@ apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib) apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if) apply (frule_tac a = m in pos_mod_bound) - apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle) + apply (simp del: pos_mod_bound add: algebra_simps nat_diff_distrib gcd_diff2 nat_le_eq_zle) + apply (metis dual_order.strict_implies_order gcd.simps gcd_0_left gcd_diff2 mod_by_0 nat_mono) done lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Orderings.thy Tue Dec 01 17:18:34 2015 +0100 @@ -310,6 +310,11 @@ "(x \ y \ P) \ (y \ x \ P) \ P" using linear by blast +lemma (in linorder) le_cases3: + "\\x \ y; y \ z\ \ P; \y \ x; x \ z\ \ P; \x \ z; z \ y\ \ P; + \z \ y; y \ x\ \ P; \y \ z; z \ x\ \ P; \z \ x; x \ y\ \ P\ \ P" +by (blast intro: le_cases) + lemma linorder_cases [case_names less equal greater]: "(x < y \ P) \ (x = y \ P) \ (y < x \ P) \ P" using less_linear by blast diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Dec 01 17:18:34 2015 +0100 @@ -46,8 +46,8 @@ assume l_r[simp]: "\n. l n \ r n" and "a \ b" and disj: "disjoint_family (\n. {l n<..r n})" assume lr_eq_ab: "(\i. {l i<..r i}) = {a<..b}" - have [intro, simp]: "\a b. a \ b \ 0 \ F b - F a" - by (auto intro!: l_r mono_F simp: diff_le_iff) + have [intro, simp]: "\a b. a \ b \ F a \ F b" + by (auto intro!: l_r mono_F) { fix S :: "nat set" assume "finite S" moreover note `a \ b` @@ -92,7 +92,7 @@ by (auto simp add: Ioc_subset_iff intro!: mono_F) finally show ?case by (auto intro: add_mono) - qed (simp add: `a \ b` less_le) + qed (auto simp add: `a \ b` less_le) qed } note claim1 = this @@ -280,7 +280,7 @@ by (auto simp add: claim1 intro!: suminf_bound) ultimately show "(\n. ereal (F (r n) - F (l n))) = ereal (F b - F a)" by simp -qed (auto simp: Ioc_inj diff_le_iff mono_F) +qed (auto simp: Ioc_inj mono_F) lemma measure_interval_measure_Ioc: assumes "a \ b" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Proofs/Extraction/Euclid.thy --- a/src/HOL/Proofs/Extraction/Euclid.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Proofs/Extraction/Euclid.thy Tue Dec 01 17:18:34 2015 +0100 @@ -29,7 +29,7 @@ by (induct m) auto lemma prime_eq: "prime (p::nat) = (1 < p \ (\m. m dvd p \ 1 < m \ m = p))" - apply (simp add: prime_nat_def) + apply (simp add: prime_def) apply (rule iffI) apply blast apply (erule conjE) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Tue Dec 01 17:18:34 2015 +0100 @@ -9,6 +9,10 @@ imports Real Topological_Spaces begin + +lemma (in ordered_ab_group_add) diff_ge_0_iff_ge [simp]: "a - b \ 0 \ a \ b" + by (simp add: le_diff_eq) + subsection \Locale for additive functions\ locale additive = @@ -777,6 +781,11 @@ thus ?thesis by simp qed +lemma norm_add_leD: + fixes a b :: "'a::real_normed_vector" + shows "norm (a + b) \ c \ norm b \ norm a + c" + by (metis add.commute diff_le_eq norm_diff_ineq order.trans) + lemma norm_diff_triangle_ineq: fixes a b c d :: "'a::real_normed_vector" shows "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Rings.thy --- a/src/HOL/Rings.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Rings.thy Tue Dec 01 17:18:34 2015 +0100 @@ -1559,6 +1559,9 @@ lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) +proposition abs_eq_iff: "abs x = abs y \ x = y \ x = -y" + by (auto simp add: abs_if split: split_if_asm) + end class linordered_ring_strict = ring + linordered_semiring_strict diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_comp.ML --- a/src/HOL/Tools/BNF/bnf_comp.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_comp.ML Tue Dec 01 17:18:34 2015 +0100 @@ -181,7 +181,8 @@ val CCA = mk_T_of_bnf oDs CAs outer; val CBs = @{map 3} mk_T_of_bnf Dss Bss_repl inners; val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer; - val inner_setss = @{map 3} mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners; + val inner_setss = + @{map 3} mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners; val inner_bds = @{map 3} mk_bd_of_bnf Dss Ass_repl inners; val outer_bd = mk_bd_of_bnf oDs CAs outer; @@ -692,7 +693,8 @@ val ((kill_poss, As), (inners', ((cache', unfold_set'), lthy'))) = normalize_bnfs qualify Ass Ds flatten_tyargs inners accum; - val Ds = oDs @ flat (@{map 3} (uncurry append oo curry swap oo map o nth) tfreess kill_poss Dss); + val Ds = + oDs @ flat (@{map 3} (uncurry append oo curry swap oo map o nth) tfreess kill_poss Dss); val As = map TFree As; in apfst (rpair (Ds, As)) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_comp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_comp_tactics.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML Tue Dec 01 17:18:34 2015 +0100 @@ -74,14 +74,16 @@ map (fn thm => rtac ctxt (thm RS fun_cong)) set_map0s @ [rtac ctxt (Gset_map0 RS comp_eq_dest_lhs), rtac ctxt sym, rtac ctxt trans_o_apply, rtac ctxt trans_image_cong_o_apply, rtac ctxt trans_image_cong_o_apply, - rtac ctxt (@{thm image_cong} OF [Gset_map0 RS comp_eq_dest_lhs RS arg_cong_Union, refl] RS trans), + rtac ctxt (@{thm image_cong} OF [Gset_map0 RS comp_eq_dest_lhs RS arg_cong_Union, refl] + RS trans), rtac ctxt @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac ctxt arg_cong_Union, rtac ctxt @{thm trans[OF comp_eq_dest_lhs[OF image_o_collect[symmetric]]]}, rtac ctxt @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @ [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac ctxt @{thm trans[OF image_insert]}, - rtac ctxt @{thm arg_cong2[of _ _ _ _ insert]}, rtac ctxt @{thm ext}, rtac ctxt trans_o_apply, - rtac ctxt trans_image_cong_o_apply, rtac ctxt @{thm trans[OF image_image]}, - rtac ctxt @{thm sym[OF trans[OF o_apply]]}, rtac ctxt @{thm image_cong[OF refl o_apply]}], + rtac ctxt @{thm arg_cong2[of _ _ _ _ insert]}, rtac ctxt @{thm ext}, + rtac ctxt trans_o_apply, rtac ctxt trans_image_cong_o_apply, + rtac ctxt @{thm trans[OF image_image]}, rtac ctxt @{thm sym[OF trans[OF o_apply]]}, + rtac ctxt @{thm image_cong[OF refl o_apply]}], rtac ctxt @{thm image_empty}]) 1; fun mk_comp_map_cong0_tac ctxt set'_eq_sets comp_set_alts map_cong0 map_cong0s = @@ -96,9 +98,9 @@ EVERY' [select_prem_tac ctxt n (dtac ctxt @{thm meta_spec}) (k + 1), etac ctxt meta_mp, rtac ctxt (equalityD2 RS set_mp), rtac ctxt (set_alt RS fun_cong RS trans), rtac ctxt trans_o_apply, rtac ctxt (@{thm collect_def} RS arg_cong_Union), - rtac ctxt @{thm UnionI}, rtac ctxt @{thm UN_I}, REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, - rtac ctxt @{thm insertI1}, rtac ctxt (o_apply RS equalityD2 RS set_mp), - etac ctxt @{thm imageI}, assume_tac ctxt]) + rtac ctxt @{thm UnionI}, rtac ctxt @{thm UN_I}, + REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, rtac ctxt @{thm insertI1}, + rtac ctxt (o_apply RS equalityD2 RS set_mp), etac ctxt @{thm imageI}, assume_tac ctxt]) comp_set_alts)) map_cong0s) 1) end; @@ -220,14 +222,15 @@ fun mk_permute_in_alt_tac ctxt src dest = (rtac ctxt @{thm Collect_cong} THEN' - mk_rotate_eq_tac ctxt (rtac ctxt refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong} - dest src) 1; + mk_rotate_eq_tac ctxt (rtac ctxt refl) trans @{thm conj_assoc} @{thm conj_commute} + @{thm conj_cong} dest src) 1; (* Miscellaneous *) fun mk_le_rel_OO_tac ctxt outer_le_rel_OO outer_rel_mono inner_le_rel_OOs = - EVERY' (map (rtac ctxt) (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: inner_le_rel_OOs)) 1; + EVERY' (map (rtac ctxt) (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: + inner_le_rel_OOs)) 1; fun mk_simple_rel_OO_Grp_tac ctxt rel_OO_Grp in_alt_thm = rtac ctxt (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_def.ML Tue Dec 01 17:18:34 2015 +0100 @@ -1300,10 +1300,12 @@ val funTs = map (fn T => bdT --> T) ranTs; val ran_bnfT = mk_bnf_T ranTs Calpha; val (revTs, Ts) = `rev (bd_bnfT :: funTs); - val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; - val tinst = fold (fn T => fn t => HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs) - (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, - map Bound (live - 1 downto 0)) $ Bound live)); + val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, + Library.foldr1 HOLogic.mk_prodT Ts]; + val tinst = fold (fn T => fn t => + HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs) + (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, + map Bound (live - 1 downto 0)) $ Bound live)); val cts = [NONE, SOME (Thm.cterm_of lthy tinst)]; in Thm.instantiate' cTs cts @{thm surj_imp_ordLeq} @@ -1420,7 +1422,8 @@ val in_rel = Lazy.lazy mk_in_rel; fun mk_rel_flip () = - unfold_thms lthy @{thms conversep_iff} (Lazy.force rel_conversep RS @{thm predicate2_eqD}); + unfold_thms lthy @{thms conversep_iff} + (Lazy.force rel_conversep RS @{thm predicate2_eqD}); val rel_flip = Lazy.lazy mk_rel_flip; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_def_tactics.ML --- a/src/HOL/Tools/BNF/bnf_def_tactics.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML Tue Dec 01 17:18:34 2015 +0100 @@ -97,18 +97,20 @@ fun mk_rel_Grp_tac ctxt rel_OO_Grps map_id0 map_cong0 map_id map_comp set_maps = let val n = length set_maps; - val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac else rtac ctxt (hd rel_OO_Grps RS trans); + val rel_OO_Grps_tac = + if null rel_OO_Grps then K all_tac else rtac ctxt (hd rel_OO_Grps RS trans); in if null set_maps then unfold_thms_tac ctxt ((map_id0 RS @{thm Grp_UNIV_id}) :: rel_OO_Grps) THEN rtac ctxt @{thm Grp_UNIV_idI[OF refl]} 1 else EVERY' [rel_OO_Grps_tac, rtac ctxt @{thm antisym}, rtac ctxt @{thm predicate2I}, - REPEAT_DETERM o - eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}], - hyp_subst_tac ctxt, rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0, - REPEAT_DETERM_N n o - EVERY' [rtac ctxt @{thm Collect_case_prod_Grp_eqD}, etac ctxt @{thm set_mp}, assume_tac ctxt], + REPEAT_DETERM o eresolve_tac ctxt + [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}], + hyp_subst_tac ctxt, rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, + rtac ctxt map_cong0, + REPEAT_DETERM_N n o EVERY' [rtac ctxt @{thm Collect_case_prod_Grp_eqD}, + etac ctxt @{thm set_mp}, assume_tac ctxt], rtac ctxt CollectI, CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), rtac ctxt @{thm image_subsetI}, rtac ctxt @{thm Collect_case_prod_Grp_in}, @@ -151,8 +153,9 @@ unfold_thms_tac ctxt [rel_compp, rel_conversep, rel_Grp, @{thm vimage2p_Grp}] THEN TRYALL (EVERY' [rtac ctxt iffI, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac ctxt CollectI, - CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, rtac ctxt @{thm relcomppI}, assume_tac ctxt, - rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac ctxt CollectI, + CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, rtac ctxt @{thm relcomppI}, + assume_tac ctxt, rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, + resolve_tac ctxt [map_id, refl], rtac ctxt CollectI, CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE}, dtac ctxt (trans OF [sym, map_id]), hyp_subst_tac ctxt, assume_tac ctxt]) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Tue Dec 01 17:18:34 2015 +0100 @@ -75,8 +75,8 @@ val fp_sugar_of_global: theory -> string -> fp_sugar option val fp_sugars_of: Proof.context -> fp_sugar list val fp_sugars_of_global: theory -> fp_sugar list - val fp_sugars_interpretation: string -> - (fp_sugar list -> local_theory -> local_theory)-> theory -> theory + val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) -> + theory -> theory val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory @@ -334,7 +334,7 @@ ); fun fp_sugar_of_generic context = - Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context) + Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); fun fp_sugars_of_generic context = Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; @@ -1206,7 +1206,8 @@ let val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; val g_absTs = map range_type fun_Ts; - val g_Tsss = map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); + val g_Tsss = + map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) Cs ctr_Tsss' g_Tsss; val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; @@ -1312,7 +1313,8 @@ ctor_rec_absTs reps fss xssss))) end; -fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = +fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss + dtor_corec = let val nn = length fpTs; val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); @@ -1371,8 +1373,8 @@ val rel_induct0_thm = Goal.prove_sorry lthy vars prems goal (fn {context = ctxt, prems} => - mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss - ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) + mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts + ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation; in (postproc_co_induct lthy (length fpA_Ts) @{thm predicate2D} @{thm predicate2D_conj} @@ -1713,8 +1715,8 @@ val thm = Goal.prove_sorry lthy vars prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => - mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts - set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) + mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts + exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> Thm.close_derivation; val case_names_attr = @@ -1811,7 +1813,8 @@ [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], - Library.foldr1 HOLogic.mk_conj (@{map 3} (build_rel_app rs') usels vsels ctrXs_Ts))]); + Library.foldr1 HOLogic.mk_conj + (@{map 3} (build_rel_app rs') usels vsels ctrXs_Ts))]); fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss = Library.foldr1 HOLogic.mk_conj (flat (@{map 6} (mk_prem_ctr_concls rs' n) @@ -2323,8 +2326,9 @@ in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => - mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss - rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) + mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) + (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs + live_nesting_rel_eqs) |> Thm.close_derivation |> Conjunction.elim_balanced nn end; @@ -2431,7 +2435,8 @@ val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs; val simp_thmss = - @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; + @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss + set_thmss; val common_notes = (if nn > 1 then diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Tue Dec 01 17:18:34 2015 +0100 @@ -305,13 +305,15 @@ end; fun solve_prem_prem_tac ctxt = - REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE' - hyp_subst_tac ctxt ORELSE' resolve_tac ctxt @{thms disjI1 disjI2}) THEN' + REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' + rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE' hyp_subst_tac ctxt ORELSE' + resolve_tac ctxt @{thms disjI1 disjI2}) THEN' (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' rtac ctxt @{thm singletonI}); fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps pre_set_defs = - HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, etac ctxt meta_mp, + HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, + etac ctxt meta_mp, SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @ sumprod_thms_set)), solve_prem_prem_tac ctxt]) (rev kks))); @@ -366,9 +368,10 @@ fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse dtor_ctor exhaust ctr_defs discss selss = let val ks = 1 upto n in - EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, - dtac ctxt meta_spec, dtac ctxt meta_mp, assume_tac ctxt, rtac ctxt exhaust, - K (co_induct_inst_as_projs_tac ctxt 0), hyp_subst_tac ctxt] @ + EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, dtac ctxt meta_spec, dtac ctxt meta_mp, + assume_tac ctxt, rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 0), + hyp_subst_tac ctxt] @ @{map 4} (fn k => fn ctr_def => fn discs => fn sels => EVERY' ([rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @ map2 (fn k' => fn discs' => @@ -435,8 +438,8 @@ abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @ @{thms id_bnf_def rel_sum_simps rel_prod_apply vimage2p_def Inl_Inr_False iffD2[OF eq_False Inr_not_Inl] sum.inject prod.inject}) THEN - REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN' (REPEAT_DETERM o rtac ctxt conjI) THEN' - (rtac ctxt refl ORELSE' assume_tac ctxt)))) + REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN' + (REPEAT_DETERM o rtac ctxt conjI) THEN' (rtac ctxt refl ORELSE' assume_tac ctxt)))) cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses); @@ -445,7 +448,8 @@ rtac ctxt ctor_rel_induct 1 THEN EVERY (@{map 6} (fn cterm => fn exhaust => fn ctor_defs => fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse => HEADGOAL (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW - (rtac ctxt (infer_instantiate' ctxt (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2) + (rtac ctxt (infer_instantiate' ctxt (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} + RS iffD2) THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' TRY o resolve_tac ctxt assms))) THEN unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @ @{thms id_bnf_def vimage2p_def}) THEN @@ -485,12 +489,14 @@ (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt)); fun mk_set_cases_tac ctxt ct assms exhaust sets = - HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN + HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) + THEN_ALL_NEW hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt sets THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE' hyp_subst_tac ctxt ORELSE' - SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o assume_tac ctxt))))); + SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o + assume_tac ctxt))))); fun mk_set_intros_tac ctxt sets = TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt sets THEN @@ -505,7 +511,8 @@ val assms_tac = let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in fold (curry (op ORELSE')) (map (fn thm => - funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' assume_tac ctxt) (rtac ctxt thm)) assms') + funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' assume_tac ctxt) + (rtac ctxt thm)) assms') (etac ctxt FalseE) end; val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts @@ -519,8 +526,8 @@ unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @ @{thms id_bnf_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert Un_empty_left Un_empty_right empty_iff singleton_iff}) THEN - REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE' eresolve_tac ctxt @{thms UN_E UnE singletonE} ORELSE' - assms_tac)) + REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE' + eresolve_tac ctxt @{thms UN_E UnE singletonE} ORELSE' assms_tac)) end; end; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Tue Dec 01 17:18:34 2015 +0100 @@ -1009,8 +1009,8 @@ |> map2 abs_curried_balanced arg_Tss |> (fn ts => Syntax.check_terms ctxt ts handle ERROR _ => nonprimitive_corec ctxt []) - |> @{map 3} (fn b => fn mx => fn t => ((b, mx), ((Binding.concealed (Thm.def_binding b), []), t))) - bs mxs + |> @{map 3} (fn b => fn mx => fn t => + ((b, mx), ((Binding.concealed (Thm.def_binding b), []), t))) bs mxs |> rpair excludess' end; @@ -1037,7 +1037,8 @@ val prems = maps (s_not_conj o #prems) disc_eqns; val ctr_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE; val code_rhs_opt = Option.map #code_rhs_opt sel_eqn_opt |> the_default NONE; - val eqn_pos = Option.map (curry (op +) 1 o #eqn_pos) sel_eqn_opt |> the_default 100000; (* FIXME *) + val eqn_pos = Option.map (curry (op +) 1 o #eqn_pos) sel_eqn_opt + |> the_default 100000; (* FIXME *) val extra_disc_eqn = {fun_name = fun_name, fun_T = fun_T, fun_args = fun_args, ctr = ctr, ctr_no = ctr_no, @@ -1307,7 +1308,8 @@ Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_primcorec_sel_tac ctxt def_thms distincts split_sels split_sel_asms - fp_nesting_maps fp_nesting_map_ident0s fp_nesting_map_comps corec_sel k m excludesss) + fp_nesting_maps fp_nesting_map_ident0s fp_nesting_map_comps corec_sel k m + excludesss) |> Thm.close_derivation |> `(is_some code_rhs_opt ? fold_thms lthy sel_defs) (*mildly too aggressive*) |> pair sel @@ -1444,7 +1446,8 @@ Goal.prove_sorry lthy [] [] raw_goal (fn {context = ctxt, prems = _} => mk_primcorec_raw_code_tac ctxt distincts discIs split_sels split_sel_asms - ms ctr_thms (if exhaustive_code then try the_single nchotomys else NONE)) + ms ctr_thms + (if exhaustive_code then try the_single nchotomys else NONE)) |> Thm.close_derivation; in Goal.prove_sorry lthy [] [] goal diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Tue Dec 01 17:18:34 2015 +0100 @@ -72,7 +72,8 @@ fun mk_primcorec_assumption_tac ctxt discIs = SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN - SOLVE (HEADGOAL (REPEAT o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' etac ctxt conjE ORELSE' + SOLVE (HEADGOAL (REPEAT o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' + etac ctxt conjE ORELSE' eresolve_tac ctxt falseEs ORELSE' resolve_tac ctxt @{thms TrueI conjI disjI1 disjI2} ORELSE' dresolve_tac ctxt discIs THEN' assume_tac ctxt ORELSE' @@ -137,7 +138,8 @@ resolve_tac ctxt split_connectI ORELSE' Splitter.split_asm_tac ctxt (split_if_asm :: split_asms) ORELSE' Splitter.split_tac ctxt (split_if :: splits) ORELSE' - eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' assume_tac ctxt ORELSE' + eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' + assume_tac ctxt ORELSE' etac ctxt notE THEN' assume_tac ctxt ORELSE' (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def sum.case sum.sel sum.distinct[THEN eq_False[THEN iffD2]]} @ @@ -148,7 +150,8 @@ fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs = HEADGOAL (rtac ctxt ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN' - (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' REPEAT_DETERM_N m o assume_tac ctxt) THEN + (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' + REPEAT_DETERM_N m o assume_tac ctxt) THEN unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac ctxt refl); fun inst_split_eq ctxt split = diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Tue Dec 01 17:18:34 2015 +0100 @@ -458,7 +458,8 @@ (recs, ctr_poss) |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos) |> Syntax.check_terms ctxt - |> @{map 3} (fn b => fn mx => fn t => ((b, mx), ((Binding.concealed (Thm.def_binding b), []), t))) + |> @{map 3} (fn b => fn mx => fn t => + ((b, mx), ((Binding.concealed (Thm.def_binding b), []), t))) bs mxs end; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML --- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML Tue Dec 01 17:18:34 2015 +0100 @@ -40,7 +40,8 @@ | basic_lfp_sugars_of bs arg_Ts callers callssss0 lthy0 = let val ((missing_arg_Ts, perm0_kks, - fp_sugars as {fp_nesting_bnfs, fp_co_induct_sugar = {common_co_inducts = [common_induct], ...}, ...} :: _, + fp_sugars as {fp_nesting_bnfs, + fp_co_induct_sugar = {common_co_inducts = [common_induct], ...}, ...} :: _, (lfp_sugar_thms, _)), lthy) = nested_to_mutual_fps (K true) Least_FP bs arg_Ts callers callssss0 lthy0; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML Tue Dec 01 17:18:34 2015 +0100 @@ -42,7 +42,8 @@ Context.proof_map (Data.map (Symtab.update (T_name, (size_name, (size_simps, size_gen_o_maps))))); fun register_size_global T_name size_name size_simps size_gen_o_maps = - Context.theory_map (Data.map (Symtab.update (T_name, (size_name, (size_simps, size_gen_o_maps))))); + Context.theory_map + (Data.map (Symtab.update (T_name, (size_name, (size_simps, size_gen_o_maps))))); val size_of = Symtab.lookup o Data.get o Context.Proof; val size_of_global = Symtab.lookup o Data.get o Context.Theory; @@ -70,8 +71,9 @@ fun mk_size_neq ctxt cts exhaust sizes = HEADGOAL (rtac ctxt (infer_instantiate' ctxt (map SOME cts) exhaust)) THEN ALLGOALS (hyp_subst_tac ctxt) THEN - Ctr_Sugar_Tactics.unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN - ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' rtac ctxt @{thm trans_less_add2})); + unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN + ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' + rtac ctxt @{thm trans_less_add2})); fun generate_datatype_size (fp_sugars as ({T = Type (_, As), BT = Type (_, Bs), fp = Least_FP, fp_res = {bnfs = fp_bnfs, ...}, fp_nesting_bnfs, live_nesting_bnfs, ...} : fp_sugar) :: _) @@ -236,7 +238,8 @@ (Spec_Rules.retrieve lthy0 @{const size ('a)} |> map_filter (try (fn (Spec_Rules.Equational, (_, [thm])) => thm))); - val nested_size_maps = map (mk_pointful lthy2) nested_size_gen_o_maps @ nested_size_gen_o_maps; + val nested_size_maps = + map (mk_pointful lthy2) nested_size_gen_o_maps @ nested_size_gen_o_maps; val all_inj_maps = @{thm prod.inj_map} :: map inj_map_of_bnf (fp_bnfs @ fp_nesting_bnfs @ live_nesting_bnfs) |> distinct Thm.eq_thm_prop; diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Tue Dec 01 17:18:34 2015 +0100 @@ -50,8 +50,8 @@ val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option - val ctr_sugar_interpretation: string -> - (ctr_sugar -> local_theory -> local_theory) -> theory -> theory + val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> + theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Tue Dec 01 17:18:34 2015 +0100 @@ -54,7 +54,8 @@ fun mk_nchotomy_tac ctxt n exhaust = HEADGOAL (rtac ctxt allI THEN' rtac ctxt exhaust THEN' - EVERY' (maps (fn k => [rtac ctxt (mk_disjIN n k), REPEAT_DETERM o rtac ctxt exI, assume_tac ctxt]) + EVERY' (maps (fn k => + [rtac ctxt (mk_disjIN n k), REPEAT_DETERM o rtac ctxt exI, assume_tac ctxt]) (1 upto n))); fun mk_unique_disc_def_tac ctxt m uexhaust = @@ -114,7 +115,8 @@ else let val ks = 1 upto n in HEADGOAL (rtac ctxt uexhaust_disc THEN' - EVERY' (@{map 5} (fn k => fn m => fn distinct_discss => fn distinct_discss' => fn uuncollapse => + EVERY' (@{map 5} (fn k => fn m => fn distinct_discss => fn distinct_discss' => + fn uuncollapse => EVERY' [rtac ctxt (uuncollapse RS trans) THEN' TRY o assume_tac ctxt, rtac ctxt sym, rtac ctxt vexhaust_disc, EVERY' (@{map 4} (fn k' => fn distinct_discs => fn distinct_discs' => fn vuncollapse => @@ -124,13 +126,17 @@ (if m = 0 then [rtac ctxt refl] else - [if n = 1 then K all_tac - else EVERY' [dtac ctxt meta_mp, assume_tac ctxt, dtac ctxt meta_mp, assume_tac ctxt], - REPEAT_DETERM_N (Int.max (0, m - 1)) o etac ctxt conjE, - asm_simp_tac (ss_only [] ctxt)]) + [if n = 1 then + K all_tac + else + EVERY' [dtac ctxt meta_mp, assume_tac ctxt, dtac ctxt meta_mp, + assume_tac ctxt], + REPEAT_DETERM_N (Int.max (0, m - 1)) o etac ctxt conjE, + asm_simp_tac (ss_only [] ctxt)]) else [dtac ctxt (the_single (if k = n then distinct_discs else distinct_discs')), - etac ctxt (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}), + etac ctxt (if k = n then @{thm iff_contradict(1)} + else @{thm iff_contradict(2)}), assume_tac ctxt, assume_tac ctxt])) ks distinct_discss distinct_discss' uncollapses)]) ks ms distinct_discsss distinct_discsss' uncollapses)) @@ -152,8 +158,8 @@ val ks = 1 upto n; in HEADGOAL (rtac ctxt (case_def' RS trans) THEN' rtac ctxt @{thm the_equality} THEN' - rtac ctxt (mk_disjIN n k) THEN' REPEAT_DETERM o rtac ctxt exI THEN' rtac ctxt conjI THEN' rtac ctxt refl THEN' - rtac ctxt refl THEN' + rtac ctxt (mk_disjIN n k) THEN' REPEAT_DETERM o rtac ctxt exI THEN' rtac ctxt conjI THEN' + rtac ctxt refl THEN' rtac ctxt refl THEN' EVERY' (map2 (fn k' => fn distincts => (if k' < n then etac ctxt disjE else K all_tac) THEN' (if k' = k then mk_case_same_ctr_tac ctxt injects @@ -182,7 +188,8 @@ simp_tac (ss_only (@{thms simp_thms} @ cases @ nth selss (k - 1) @ nth injectss (k - 1) @ flat (nth distinctsss (k - 1))) ctxt)) k) THEN ALLGOALS (etac ctxt thin_rl THEN' rtac ctxt iffI THEN' - REPEAT_DETERM o rtac ctxt allI THEN' rtac ctxt impI THEN' REPEAT_DETERM o etac ctxt conjE THEN' + REPEAT_DETERM o rtac ctxt allI THEN' rtac ctxt impI THEN' + REPEAT_DETERM o etac ctxt conjE THEN' hyp_subst_tac ctxt THEN' assume_tac ctxt THEN' REPEAT_DETERM o etac ctxt allE THEN' etac ctxt impE THEN' REPEAT_DETERM o (rtac ctxt conjI THEN' rtac ctxt refl) THEN' diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Tools/Transfer/transfer_bnf.ML --- a/src/HOL/Tools/Transfer/transfer_bnf.ML Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Tue Dec 01 17:18:34 2015 +0100 @@ -66,11 +66,6 @@ fun bnf_of_fp_sugar (fp_sugar:fp_sugar) = nth (#bnfs (#fp_res fp_sugar)) (#fp_res_index fp_sugar) -fun fp_sugar_only_type_ctr f fp_sugars = - (case filter (is_Type o T_of_bnf o bnf_of_fp_sugar) fp_sugars of - [] => I - | fp_sugars' => f fp_sugars') - (* relation constraints - bi_total & co. *) fun mk_relation_constraint name arg = @@ -410,7 +405,7 @@ fun transfer_fp_sugars_interpretation fp_sugar lthy = let - val lthy = pred_injects fp_sugar lthy + val lthy = lthy |> (is_Type o T_of_bnf o bnf_of_fp_sugar) fp_sugar ? pred_injects fp_sugar val transfer_rules_notes = fp_sugar_transfer_rules fp_sugar in lthy @@ -419,7 +414,6 @@ end val _ = - Theory.setup (fp_sugars_interpretation transfer_plugin - (fp_sugar_only_type_ctr (fold transfer_fp_sugars_interpretation))) + Theory.setup (fp_sugars_interpretation transfer_plugin (fold transfer_fp_sugars_interpretation)) end diff -r 507b39df1a57 -r f58d75535f66 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/Transcendental.thy Tue Dec 01 17:18:34 2015 +0100 @@ -10,6 +10,16 @@ imports Binomial Series Deriv NthRoot begin +lemma of_int_leD: + fixes x :: "'a :: linordered_idom" + shows "\of_int n\ \ x \ n=0 \ x\1" + by (metis (no_types) le_less_trans not_less of_int_abs of_int_less_1_iff zabs_less_one_iff) + +lemma of_int_lessD: + fixes x :: "'a :: linordered_idom" + shows "\of_int n\ < x \ n=0 \ x>1" + by (metis less_le_trans not_less of_int_abs of_int_less_1_iff zabs_less_one_iff) + lemma fact_in_Reals: "fact n \ \" by (induction n) auto lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" @@ -1979,8 +1989,7 @@ assume "x \ y" "y \ a" with \0 < x\ \a < 1\ have "0 < 1 / y - 1" "0 < y" by (auto simp: field_simps) - with D show "\z. DERIV ?l y :> z \ 0 < z" - by auto + with D show "\z. DERIV ?l y :> z \ 0 < z" by blast qed also have "\ \ 0" using ln_le_minus_one \0 < x\ \x < a\ by (auto simp: field_simps) @@ -3090,6 +3099,11 @@ using cos_add [where x=x and y=x] by (simp add: power2_eq_square) +lemma sin_cos_le1: + fixes x::real shows "abs (sin x * sin y + cos x * cos y) \ 1" + using cos_diff [of x y] + by (metis abs_cos_le_one add.commute) + lemma DERIV_fun_pow: "DERIV g x :> m ==> DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" by (auto intro!: derivative_eq_intros simp:) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/ex/Dedekind_Real.thy --- a/src/HOL/ex/Dedekind_Real.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/ex/Dedekind_Real.thy Tue Dec 01 17:18:34 2015 +0100 @@ -784,8 +784,7 @@ qed hence "y < r" by simp with ypos have dless: "?d < (r * ?d)/y" - by (simp add: pos_less_divide_eq mult.commute [of ?d] - mult_strict_right_mono dpos) + using dpos less_divide_eq_1 by fastforce have "r + ?d < r*x" proof - have "r + ?d < r + (r * ?d)/y" by (simp add: dless) diff -r 507b39df1a57 -r f58d75535f66 src/HOL/ex/Sqrt.thy --- a/src/HOL/ex/Sqrt.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/ex/Sqrt.thy Tue Dec 01 17:18:34 2015 +0100 @@ -14,7 +14,7 @@ assumes "prime (p::nat)" shows "sqrt p \ \" proof - from \prime p\ have p: "1 < p" by (simp add: prime_nat_def) + from \prime p\ have p: "1 < p" by (simp add: prime_def) assume "sqrt p \ \" then obtain m n :: nat where n: "n \ 0" and sqrt_rat: "\sqrt p\ = m / n" @@ -59,7 +59,7 @@ assumes "prime (p::nat)" shows "sqrt p \ \" proof - from \prime p\ have p: "1 < p" by (simp add: prime_nat_def) + from \prime p\ have p: "1 < p" by (simp add: prime_def) assume "sqrt p \ \" then obtain m n :: nat where n: "n \ 0" and sqrt_rat: "\sqrt p\ = m / n" diff -r 507b39df1a57 -r f58d75535f66 src/HOL/ex/Sqrt_Script.thy --- a/src/HOL/ex/Sqrt_Script.thy Tue Dec 01 12:35:11 2015 +0100 +++ b/src/HOL/ex/Sqrt_Script.thy Tue Dec 01 17:18:34 2015 +0100 @@ -17,7 +17,7 @@ subsection \Preliminaries\ lemma prime_nonzero: "prime (p::nat) \ p \ 0" - by (force simp add: prime_nat_def) + by (force simp add: prime_def) lemma prime_dvd_other_side: "(n::nat) * n = p * (k * k) \ prime p \ p dvd n" @@ -32,7 +32,7 @@ apply (erule disjE) apply (frule mult_le_mono, assumption) apply auto - apply (force simp add: prime_nat_def) + apply (force simp add: prime_def) done lemma rearrange: "(j::nat) * (p * j) = k * k \ k * k = p * (j * j)"